src/HOL/Nominal/Examples/Class.thy
author urbanc
Thu, 24 May 2007 12:09:38 +0200
changeset 23092 f3615235dc4d
parent 22542 8279a25ad0ae
child 23093 e3735771e7ba
permissions -rw-r--r--
formalisation of my PhD (the result was correct, but the proof needed several corrections)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
19500
188d4e44c1a6 Removed broken proof.
berghofe
parents: 19477
diff changeset
     1
(* $Id$ *)
18425
bcf13dbaa339 I think the earlier version was completely broken
urbanc
parents: 18395
diff changeset
     2
23092
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
     3
theory ClassWork 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
     4
imports "Nominal" 
18425
bcf13dbaa339 I think the earlier version was completely broken
urbanc
parents: 18395
diff changeset
     5
begin
18395
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
     6
18881
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
     7
section {* Term-Calculus from Urban's PhD *}
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
     8
18395
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
     9
atom_decl name coname
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    10
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
    11
text {* types *}
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
    12
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
    13
nominal_datatype ty =
23092
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
    14
    PR "string"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
    15
  | NOT  "ty"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
    16
  | AND  "ty" "ty"   ("_ AND _" [100,100] 100)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
    17
  | OR   "ty" "ty"   ("_ OR _" [100,100] 100)
23092
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
    18
  | IMP  "ty" "ty"   ("_ IMP _" [100,100] 100)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
    19
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
    20
instance ty :: size ..
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
    21
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
    22
nominal_primrec
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
    23
  "size (PR s)     = (1::nat)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
    24
  "size (NOT T)     = 1 + size T"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
    25
  "size (T1 AND T2) = 1 + size T1 + size T2"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
    26
  "size (T1 OR T2)  = 1 + size T1 + size T2"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
    27
  "size (T1 IMP T2) = 1 + size T1 + size T2"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
    28
by (rule TrueI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
    29
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
    30
lemma ty_cases:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
    31
  fixes T::ty
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
    32
  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)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
    33
by (induct T rule:ty.weak_induct) (auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
    34
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
    35
lemma fresh_ty:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
    36
  fixes a::"coname"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
    37
  and   x::"name"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
    38
  and   T::"ty"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
    39
  shows "a\<sharp>T" and "x\<sharp>T"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
    40
by (nominal_induct T rule: ty.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
    41
   (auto simp add: fresh_string)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
    42
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
    43
text {* terms *}
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
    44
18881
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    45
nominal_datatype trm = 
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    46
    Ax   "name" "coname"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
    47
  | Cut  "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm"            ("Cut <_>._ (_)._" [100,100,100,100] 100)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
    48
  | NotR "\<guillemotleft>name\<guillemotright>trm" "coname"                 ("NotR (_)._ _" [100,100,100] 100)
18881
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    49
  | NotL "\<guillemotleft>coname\<guillemotright>trm" "name"                 ("NotL <_>._ _" [100,100,100] 100)
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    50
  | AndR "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>coname\<guillemotright>trm" "coname" ("AndR <_>._ <_>._ _" [100,100,100,100,100] 100)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
    51
  | AndL1 "\<guillemotleft>name\<guillemotright>trm" "name"                  ("AndL1 (_)._ _" [100,100,100] 100)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
    52
  | AndL2 "\<guillemotleft>name\<guillemotright>trm" "name"                  ("AndL2 (_)._ _" [100,100,100] 100)
18881
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    53
  | OrR1 "\<guillemotleft>coname\<guillemotright>trm" "coname"               ("OrR1 <_>._ _" [100,100,100] 100)
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    54
  | OrR2 "\<guillemotleft>coname\<guillemotright>trm" "coname"               ("OrR2 <_>._ _" [100,100,100] 100)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
    55
  | OrL "\<guillemotleft>name\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" "name"        ("OrL (_)._ (_)._ _" [100,100,100,100,100] 100)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
    56
  | ImpR "\<guillemotleft>name\<guillemotright>(\<guillemotleft>coname\<guillemotright>trm)" "coname"       ("ImpR (_).<_>._ _" [100,100,100,100] 100)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
    57
  | ImpL "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" "name"     ("ImpL <_>._ (_)._ _" [100,100,100,100,100] 100)
19319
7e1f85ceb1a2 added the first two simple proofs of the recursion
urbanc
parents: 18881
diff changeset
    58
18881
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    59
text {* named terms *}
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    60
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
    61
nominal_datatype ntrm = Na "\<guillemotleft>name\<guillemotright>trm" ("((_):_)" [100,100] 100)
18881
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    62
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    63
text {* conamed terms *}
18395
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    64
22232
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    65
nominal_datatype ctrm = Co "\<guillemotleft>coname\<guillemotright>trm" ("(<_>:_)" [100,100] 100)
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    66
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    67
text {* renaming functions *}
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    68
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    69
consts
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    70
  nrename :: "trm \<Rightarrow> name \<Rightarrow> name \<Rightarrow> trm"      ("_[_\<turnstile>n>_]" [100,100,100] 100) 
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    71
  crename :: "trm \<Rightarrow> coname \<Rightarrow> coname \<Rightarrow> trm"  ("_[_\<turnstile>c>_]" [100,100,100] 100) 
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    72
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    73
nominal_primrec (freshness_context: "(d::coname,e::coname)") 
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    74
  "(Ax x a)[d\<turnstile>c>e] = (if a=d then Ax x e else Ax x a)" 
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
    75
  "\<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])" 
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
    76
  "(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)" 
22232
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    77
  "a\<sharp>(d,e) \<Longrightarrow> (NotL <a>.M x)[d\<turnstile>c>e] = (NotL <a>.(M[d\<turnstile>c>e]) x)" 
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    78
  "\<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] = 
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    79
          (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)" 
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
    80
  "x\<sharp>y \<Longrightarrow> (AndL1 (x).M y)[d\<turnstile>c>e] = AndL1 (x).(M[d\<turnstile>c>e]) y"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
    81
  "x\<sharp>y \<Longrightarrow> (AndL2 (x).M y)[d\<turnstile>c>e] = AndL2 (x).(M[d\<turnstile>c>e]) y"
22232
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    82
  "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)"
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    83
  "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)"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
    84
  "\<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"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
    85
  "a\<sharp>(d,e,b) \<Longrightarrow> (ImpR (x).<a>.M b)[d\<turnstile>c>e] = 
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
    86
       (if b=d then ImpR (x).<a>.(M[d\<turnstile>c>e]) e else ImpR (x).<a>.(M[d\<turnstile>c>e]) b)"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
    87
  "\<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"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
    88
apply(finite_guess)+
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
    89
apply(rule TrueI)+
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
    90
apply(simp add: abs_fresh abs_supp fin_supp)+
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
    91
apply(fresh_guess)+
22232
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    92
done
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    93
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    94
nominal_primrec (freshness_context: "(u::name,v::name)") 
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    95
  "(Ax x a)[u\<turnstile>n>v] = (if x=u then Ax v a else Ax x a)" 
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
    96
  "\<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])" 
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
    97
  "x\<sharp>(u,v) \<Longrightarrow> (NotR (x).M a)[u\<turnstile>n>v] = NotR (x).(M[u\<turnstile>n>v]) a" 
22232
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    98
  "(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)" 
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
    99
  "\<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" 
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   100
  "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)"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   101
  "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)"
22232
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
   102
  "a\<sharp>b \<Longrightarrow> (OrR1 <a>.M b)[u\<turnstile>n>v] = OrR1 <a>.(M[u\<turnstile>n>v]) b"
340cb955008e added the two renaming functions from my PhD
urbanc
parents: 19500
diff changeset
   103
  "a\<sharp>b \<Longrightarrow> (OrR2 <a>.M b)[u\<turnstile>n>v] = OrR2 <a>.(M[u\<turnstile>n>v]) b"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   104
  "\<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] = 
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   105
        (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)"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   106
  "\<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"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   107
  "\<lbrakk>a\<sharp>N;x\<sharp>(u,v,M,y)\<rbrakk> \<Longrightarrow> (ImpL <a>.M (x).N y)[u\<turnstile>n>v] = 
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   108
        (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)"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   109
apply(finite_guess)+
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   110
apply(rule TrueI)+
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   111
apply(simp add: abs_fresh abs_supp fs_name1 fs_coname1)+
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   112
apply(fresh_guess)+
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   113
done
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   114
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   115
lemmas eq_bij = pt_bij[OF pt_name_inst, OF at_name_inst] pt_bij[OF pt_coname_inst, OF at_coname_inst]
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   116
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   117
lemma crename_name_eqvt[eqvt]:
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   118
  fixes pi::"name prm"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   119
  shows "pi\<bullet>(M[d\<turnstile>c>e]) = (pi\<bullet>M)[(pi\<bullet>d)\<turnstile>c>(pi\<bullet>e)]"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   120
apply(nominal_induct M avoiding: d e rule: trm.induct)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   121
apply(auto simp add: fresh_bij eq_bij)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   122
done
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   123
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   124
lemma crename_coname_eqvt[eqvt]:
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   125
  fixes pi::"coname prm"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   126
  shows "pi\<bullet>(M[d\<turnstile>c>e]) = (pi\<bullet>M)[(pi\<bullet>d)\<turnstile>c>(pi\<bullet>e)]"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   127
apply(nominal_induct M avoiding: d e rule: trm.induct)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   128
apply(auto simp add: fresh_bij eq_bij)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   129
done
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   130
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   131
lemma nrename_name_eqvt[eqvt]:
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   132
  fixes pi::"name prm"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   133
  shows "pi\<bullet>(M[x\<turnstile>n>y]) = (pi\<bullet>M)[(pi\<bullet>x)\<turnstile>n>(pi\<bullet>y)]"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   134
apply(nominal_induct M avoiding: x y rule: trm.induct)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   135
apply(auto simp add: fresh_bij eq_bij)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   136
done
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   137
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   138
lemma nrename_coname_eqvt[eqvt]:
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   139
  fixes pi::"coname prm"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   140
  shows "pi\<bullet>(M[x\<turnstile>n>y]) = (pi\<bullet>M)[(pi\<bullet>x)\<turnstile>n>(pi\<bullet>y)]"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   141
apply(nominal_induct M avoiding: x y rule: trm.induct)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   142
apply(auto simp add: fresh_bij eq_bij)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   143
done
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   144
23092
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   145
lemmas rename_eqvts = crename_name_eqvt crename_coname_eqvt
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   146
                      nrename_name_eqvt nrename_coname_eqvt
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   147
lemma nrename_fresh:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   148
  assumes a: "x\<sharp>M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   149
  shows "M[x\<turnstile>n>y] = M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   150
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   151
by (nominal_induct M avoiding: x y rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   152
   (auto simp add: trm.inject fresh_atm abs_fresh fin_supp abs_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   153
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   154
lemma crename_fresh:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   155
  assumes a: "a\<sharp>M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   156
  shows "M[a\<turnstile>c>b] = M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   157
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   158
by (nominal_induct M avoiding: a b rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   159
   (auto simp add: trm.inject fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   160
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   161
lemma nrename_nfresh:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   162
  fixes x::"name"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   163
  shows "x\<sharp>y\<Longrightarrow>x\<sharp>M[x\<turnstile>n>y]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   164
by (nominal_induct M avoiding: x y rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   165
   (auto simp add: fresh_atm abs_fresh abs_supp fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   166
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   167
 lemma crename_nfresh:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   168
  fixes x::"name"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   169
  shows "x\<sharp>M\<Longrightarrow>x\<sharp>M[a\<turnstile>c>b]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   170
by (nominal_induct M avoiding: a b rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   171
   (auto simp add: fresh_atm abs_fresh abs_supp fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   172
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   173
lemma crename_cfresh:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   174
  fixes a::"coname"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   175
  shows "a\<sharp>b\<Longrightarrow>a\<sharp>M[a\<turnstile>c>b]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   176
by (nominal_induct M avoiding: a b rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   177
   (auto simp add: fresh_atm abs_fresh abs_supp fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   178
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   179
lemma nrename_cfresh:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   180
  fixes c::"coname"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   181
  shows "c\<sharp>M\<Longrightarrow>c\<sharp>M[x\<turnstile>n>y]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   182
by (nominal_induct M avoiding: x y rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   183
   (auto simp add: fresh_atm abs_fresh abs_supp fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   184
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   185
lemma nrename_nfresh':
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   186
  fixes x::"name"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   187
  shows "x\<sharp>(M,z,y)\<Longrightarrow>x\<sharp>M[z\<turnstile>n>y]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   188
by (nominal_induct M avoiding: x z y rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   189
   (auto simp add: fresh_prod fresh_atm abs_fresh abs_supp fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   190
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   191
lemma crename_cfresh':
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   192
  fixes a::"coname"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   193
  shows "a\<sharp>(M,b,c)\<Longrightarrow>a\<sharp>M[b\<turnstile>c>c]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   194
by (nominal_induct M avoiding: a b c rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   195
   (auto simp add: fresh_prod fresh_atm abs_fresh abs_supp fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   196
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   197
lemma nrename_rename:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   198
  assumes a: "x'\<sharp>M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   199
  shows "([(x',x)]\<bullet>M)[x'\<turnstile>n>y]= M[x\<turnstile>n>y]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   200
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   201
apply(nominal_induct M avoiding: x x' y rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   202
apply(auto simp add: abs_fresh fresh_bij fresh_atm fresh_prod fresh_right calc_atm abs_supp fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   203
apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   204
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   205
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   206
lemma crename_rename:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   207
  assumes a: "a'\<sharp>M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   208
  shows "([(a',a)]\<bullet>M)[a'\<turnstile>c>b]= M[a\<turnstile>c>b]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   209
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   210
apply(nominal_induct M avoiding: a a' b rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   211
apply(auto simp add: abs_fresh fresh_bij fresh_atm fresh_prod fresh_right calc_atm abs_supp fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   212
apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   213
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   214
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   215
lemmas rename_fresh = nrename_fresh crename_fresh 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   216
                      nrename_nfresh crename_nfresh crename_cfresh nrename_cfresh
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   217
                      nrename_nfresh' crename_cfresh'
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   218
                      nrename_rename crename_rename
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   219
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   220
lemma better_nrename_Cut:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   221
  assumes a: "x\<sharp>(u,v)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   222
  shows "(Cut <a>.M (x).N)[u\<turnstile>n>v] = Cut <a>.(M[u\<turnstile>n>v]) (x).(N[u\<turnstile>n>v])"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   223
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   224
  obtain x'::"name"   where fs1: "x'\<sharp>(M,N,a,x,u,v)" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   225
  obtain a'::"coname" where fs2: "a'\<sharp>(M,N,a,x,u,v)" by (rule exists_fresh(2), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   226
  have eq1: "(Cut <a>.M (x).N) = (Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   227
    using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   228
  have "(Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N))[u\<turnstile>n>v] = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   229
                        Cut <a'>.(([(a',a)]\<bullet>M)[u\<turnstile>n>v]) (x').(([(x',x)]\<bullet>N)[u\<turnstile>n>v])"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   230
    using fs1 fs2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   231
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   232
    apply(rule nrename.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   233
    apply(simp add: fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   234
    apply(simp add: fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   235
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   236
  also have "\<dots> = Cut <a>.(M[u\<turnstile>n>v]) (x).(N[u\<turnstile>n>v])" using fs1 fs2 a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   237
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   238
    apply(simp add: trm.inject alpha fresh_atm fresh_prod rename_eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   239
    apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   240
    apply(simp add: rename_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   241
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   242
  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
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   243
    by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   244
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   245
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   246
lemma better_crename_Cut:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   247
  assumes a: "a\<sharp>(b,c)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   248
  shows "(Cut <a>.M (x).N)[b\<turnstile>c>c] = Cut <a>.(M[b\<turnstile>c>c]) (x).(N[b\<turnstile>c>c])"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   249
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   250
  obtain x'::"name"   where fs1: "x'\<sharp>(M,N,a,x,b,c)" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   251
  obtain a'::"coname" where fs2: "a'\<sharp>(M,N,a,x,b,c)" by (rule exists_fresh(2), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   252
  have eq1: "(Cut <a>.M (x).N) = (Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   253
    using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   254
  have "(Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N))[b\<turnstile>c>c] = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   255
                        Cut <a'>.(([(a',a)]\<bullet>M)[b\<turnstile>c>c]) (x').(([(x',x)]\<bullet>N)[b\<turnstile>c>c])"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   256
    using fs1 fs2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   257
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   258
    apply(rule crename.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   259
    apply(simp add: fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   260
    apply(simp add: fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   261
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   262
  also have "\<dots> = Cut <a>.(M[b\<turnstile>c>c]) (x).(N[b\<turnstile>c>c])" using fs1 fs2 a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   263
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   264
    apply(simp add: trm.inject alpha fresh_atm fresh_prod rename_eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   265
    apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   266
    apply(simp add: rename_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   267
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   268
  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
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   269
    by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   270
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   271
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   272
lemma crename_id:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   273
  shows "M[a\<turnstile>c>a] = M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   274
by (nominal_induct M avoiding: a rule: trm.induct) (auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   275
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   276
lemma nrename_id:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   277
  shows "M[x\<turnstile>n>x] = M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   278
by (nominal_induct M avoiding: x rule: trm.induct) (auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   279
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   280
lemma nrename_swap:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   281
  shows "x\<sharp>M \<Longrightarrow> [(x,y)]\<bullet>M = M[y\<turnstile>n>x]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   282
by (nominal_induct M avoiding: x y rule: trm.induct) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   283
   (simp_all add: calc_atm fresh_atm trm.inject alpha abs_fresh abs_supp fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   284
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   285
lemma crename_swap:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   286
  shows "a\<sharp>M \<Longrightarrow> [(a,b)]\<bullet>M = M[b\<turnstile>c>a]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   287
by (nominal_induct M avoiding: a b rule: trm.induct) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   288
   (simp_all add: calc_atm fresh_atm trm.inject alpha abs_fresh abs_supp fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   289
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   290
lemma crename_ax:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   291
  assumes a: "M[a\<turnstile>c>b] = Ax x c" "c\<noteq>a" "c\<noteq>b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   292
  shows "M = Ax x c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   293
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   294
apply(nominal_induct M avoiding: a b x c rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   295
apply(simp_all add: trm.inject split: if_splits)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   296
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   297
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   298
lemma nrename_ax:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   299
  assumes a: "M[x\<turnstile>n>y] = Ax z a" "z\<noteq>x" "z\<noteq>y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   300
  shows "M = Ax z a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   301
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   302
apply(nominal_induct M avoiding: x y z a rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   303
apply(simp_all add: trm.inject split: if_splits)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   304
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   305
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   306
text {* substitution functions *}
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   307
23092
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   308
lemma fresh_perm_coname:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   309
  fixes c::"coname"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   310
  and   pi::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   311
  and   M::"trm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   312
  assumes a: "c\<sharp>pi" "c\<sharp>M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   313
  shows "c\<sharp>(pi\<bullet>M)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   314
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   315
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   316
apply(simp add: fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   317
apply(simp add: at_prm_fresh[OF at_coname_inst] fresh_list_rev)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   318
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   319
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   320
lemma fresh_perm_name:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   321
  fixes x::"name"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   322
  and   pi::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   323
  and   M::"trm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   324
  assumes a: "x\<sharp>pi" "x\<sharp>M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   325
  shows "x\<sharp>(pi\<bullet>M)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   326
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   327
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   328
apply(simp add: fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   329
apply(simp add: at_prm_fresh[OF at_name_inst] fresh_list_rev)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   330
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   331
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   332
lemma fresh_fun_simp_NotL:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   333
  assumes a: "x'\<sharp>P" "x'\<sharp>M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   334
  shows "fresh_fun (\<lambda>x'. Cut <c>.P (x').NotL <a>.M x') = Cut <c>.P (x').NotL <a>.M x'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   335
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   336
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   337
apply(rule fresh_fun_app)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   338
apply(rule pt_name_inst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   339
apply(rule at_name_inst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   340
apply(finite_guess)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   341
apply(subgoal_tac "\<exists>n::name. n\<sharp>(c,P,a,M)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   342
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   343
apply(rule_tac x="n" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   344
apply(simp add: fresh_prod abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   345
apply(fresh_guess)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   346
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   347
apply(simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   348
apply(fresh_guess)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   349
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   350
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   351
lemma fresh_fun_NotL[eqvt_force]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   352
  fixes pi1::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   353
  and   pi2::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   354
  shows "pi1\<bullet>fresh_fun (\<lambda>x'. Cut <c>.P (x').NotL <a>.M x')=
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   355
             fresh_fun (pi1\<bullet>(\<lambda>x'. Cut <c>.P (x').NotL <a>.M x'))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   356
  and   "pi2\<bullet>fresh_fun (\<lambda>x'. Cut <c>.P (x').NotL <a>.M x')=
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   357
             fresh_fun (pi2\<bullet>(\<lambda>x'. Cut <c>.P (x').NotL <a>.M x'))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   358
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   359
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   360
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   361
apply(auto simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   362
apply(simp add: fresh_fun_simp_NotL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   363
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   364
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   365
apply(rule fresh_fun_simp_NotL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   366
apply(rule fresh_perm_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   367
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   368
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   369
apply(rule fresh_perm_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   370
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   371
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   372
apply(simp add: at_prm_fresh[OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   373
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   374
apply(subgoal_tac "\<exists>n::name. n\<sharp>(P,M,pi2\<bullet>P,pi2\<bullet>M,pi2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   375
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   376
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   377
apply(simp add: fresh_fun_simp_NotL calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   378
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   379
apply(simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   380
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   381
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   382
lemma fresh_fun_simp_AndL1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   383
  assumes a: "z'\<sharp>P" "z'\<sharp>M" "z'\<sharp>x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   384
  shows "fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL1 (x).M z') = Cut <c>.P (z').AndL1 (x).M z'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   385
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   386
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   387
apply(rule fresh_fun_app)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   388
apply(rule pt_name_inst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   389
apply(rule at_name_inst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   390
apply(finite_guess)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   391
apply(subgoal_tac "\<exists>n::name. n\<sharp>(c,P,x,M)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   392
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   393
apply(rule_tac x="n" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   394
apply(simp add: fresh_prod abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   395
apply(fresh_guess)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   396
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   397
apply(simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   398
apply(fresh_guess)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   399
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   400
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   401
lemma fresh_fun_AndL1[eqvt_force]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   402
  fixes pi1::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   403
  and   pi2::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   404
  shows "pi1\<bullet>fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL1 (x).M z')=
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   405
             fresh_fun (pi1\<bullet>(\<lambda>z'. Cut <c>.P (z').AndL1 (x).M z'))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   406
  and   "pi2\<bullet>fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL1 (x).M z')=
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   407
             fresh_fun (pi2\<bullet>(\<lambda>z'. Cut <c>.P (z').AndL1 (x).M z'))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   408
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   409
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   410
apply(subgoal_tac "\<exists>n::name. n\<sharp>(P,M,x,pi1\<bullet>P,pi1\<bullet>M,pi1\<bullet>x,pi1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   411
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   412
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   413
apply(simp add: fresh_fun_simp_AndL1 at_prm_fresh[OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   414
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   415
apply(simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   416
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   417
apply(subgoal_tac "\<exists>n::name. n\<sharp>(P,M,x,pi2\<bullet>P,pi2\<bullet>M,pi2\<bullet>x,pi2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   418
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   419
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   420
apply(simp add: fresh_fun_simp_AndL1 calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   421
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   422
apply(simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   423
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   424
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   425
lemma fresh_fun_simp_AndL2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   426
  assumes a: "z'\<sharp>P" "z'\<sharp>M" "z'\<sharp>x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   427
  shows "fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL2 (x).M z') = Cut <c>.P (z').AndL2 (x).M z'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   428
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   429
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   430
apply(rule fresh_fun_app)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   431
apply(rule pt_name_inst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   432
apply(rule at_name_inst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   433
apply(finite_guess)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   434
apply(subgoal_tac "\<exists>n::name. n\<sharp>(c,P,x,M)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   435
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   436
apply(rule_tac x="n" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   437
apply(simp add: fresh_prod abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   438
apply(fresh_guess)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   439
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   440
apply(simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   441
apply(fresh_guess)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   442
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   443
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   444
lemma fresh_fun_AndL2[eqvt_force]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   445
  fixes pi1::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   446
  and   pi2::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   447
  shows "pi1\<bullet>fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL2 (x).M z')=
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   448
             fresh_fun (pi1\<bullet>(\<lambda>z'. Cut <c>.P (z').AndL2 (x).M z'))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   449
  and   "pi2\<bullet>fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL2 (x).M z')=
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   450
             fresh_fun (pi2\<bullet>(\<lambda>z'. Cut <c>.P (z').AndL2 (x).M z'))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   451
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   452
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   453
apply(subgoal_tac "\<exists>n::name. n\<sharp>(P,M,x,pi1\<bullet>P,pi1\<bullet>M,pi1\<bullet>x,pi1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   454
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   455
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   456
apply(simp add: fresh_fun_simp_AndL2 at_prm_fresh[OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   457
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   458
apply(simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   459
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   460
apply(subgoal_tac "\<exists>n::name. n\<sharp>(P,M,x,pi2\<bullet>P,pi2\<bullet>M,pi2\<bullet>x,pi2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   461
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   462
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   463
apply(simp add: fresh_fun_simp_AndL2 calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   464
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   465
apply(simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   466
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   467
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   468
lemma fresh_fun_simp_OrL:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   469
  assumes a: "z'\<sharp>P" "z'\<sharp>M" "z'\<sharp>N" "z'\<sharp>u" "z'\<sharp>x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   470
  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'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   471
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   472
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   473
apply(rule fresh_fun_app)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   474
apply(rule pt_name_inst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   475
apply(rule at_name_inst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   476
apply(finite_guess)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   477
apply(subgoal_tac "\<exists>n::name. n\<sharp>(c,P,x,M,u,N)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   478
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   479
apply(rule_tac x="n" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   480
apply(simp add: fresh_prod abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   481
apply(fresh_guess)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   482
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   483
apply(simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   484
apply(fresh_guess)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   485
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   486
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   487
lemma fresh_fun_OrL[eqvt_force]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   488
  fixes pi1::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   489
  and   pi2::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   490
  shows "pi1\<bullet>fresh_fun (\<lambda>z'. Cut <c>.P (z').OrL (x).M (u).N z')=
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   491
             fresh_fun (pi1\<bullet>(\<lambda>z'. Cut <c>.P (z').OrL (x).M (u).N z'))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   492
  and   "pi2\<bullet>fresh_fun (\<lambda>z'. Cut <c>.P (z').OrL (x).M (u).N z')=
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   493
             fresh_fun (pi2\<bullet>(\<lambda>z'. Cut <c>.P (z').OrL (x).M (u).N z'))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   494
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   495
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   496
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)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   497
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   498
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   499
apply(simp add: fresh_fun_simp_OrL at_prm_fresh[OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   500
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   501
apply(simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   502
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   503
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)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   504
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   505
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   506
apply(simp add: fresh_fun_simp_OrL calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   507
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   508
apply(simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   509
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   510
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   511
lemma fresh_fun_simp_ImpL:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   512
  assumes a: "z'\<sharp>P" "z'\<sharp>M" "z'\<sharp>N" "z'\<sharp>x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   513
  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'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   514
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   515
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   516
apply(rule fresh_fun_app)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   517
apply(rule pt_name_inst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   518
apply(rule at_name_inst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   519
apply(finite_guess)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   520
apply(subgoal_tac "\<exists>n::name. n\<sharp>(c,P,x,M,N)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   521
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   522
apply(rule_tac x="n" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   523
apply(simp add: fresh_prod abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   524
apply(fresh_guess)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   525
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   526
apply(simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   527
apply(fresh_guess)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   528
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   529
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   530
lemma fresh_fun_ImpL[eqvt_force]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   531
  fixes pi1::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   532
  and   pi2::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   533
  shows "pi1\<bullet>fresh_fun (\<lambda>z'. Cut <c>.P (z').ImpL <a>.M (x).N z')=
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   534
             fresh_fun (pi1\<bullet>(\<lambda>z'. Cut <c>.P (z').ImpL <a>.M (x).N z'))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   535
  and   "pi2\<bullet>fresh_fun (\<lambda>z'. Cut <c>.P (z').ImpL <a>.M (x).N z')=
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   536
             fresh_fun (pi2\<bullet>(\<lambda>z'. Cut <c>.P (z').ImpL <a>.M (x).N z'))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   537
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   538
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   539
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)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   540
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   541
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   542
apply(simp add: fresh_fun_simp_ImpL at_prm_fresh[OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   543
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   544
apply(simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   545
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   546
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)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   547
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   548
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   549
apply(simp add: fresh_fun_simp_ImpL calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   550
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   551
apply(simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   552
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   553
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   554
lemma fresh_fun_simp_NotR:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   555
  assumes a: "a'\<sharp>P" "a'\<sharp>M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   556
  shows "fresh_fun (\<lambda>a'. Cut <a'>.(NotR (y).M a') (x).P) = Cut <a'>.(NotR (y).M a') (x).P"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   557
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   558
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   559
apply(rule fresh_fun_app)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   560
apply(rule pt_coname_inst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   561
apply(rule at_coname_inst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   562
apply(finite_guess)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   563
apply(subgoal_tac "\<exists>n::coname. n\<sharp>(x,P,y,M)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   564
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   565
apply(rule_tac x="n" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   566
apply(simp add: fresh_prod abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   567
apply(fresh_guess)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   568
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   569
apply(simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   570
apply(fresh_guess)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   571
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   572
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   573
lemma fresh_fun_NotR[eqvt_force]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   574
  fixes pi1::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   575
  and   pi2::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   576
  shows "pi1\<bullet>fresh_fun (\<lambda>a'. Cut <a'>.(NotR (y).M a') (x).P)=
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   577
             fresh_fun (pi1\<bullet>(\<lambda>a'. Cut <a'>.(NotR (y).M a') (x).P))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   578
  and   "pi2\<bullet>fresh_fun (\<lambda>a'. Cut <a'>.(NotR (y).M a') (x).P)=
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   579
             fresh_fun (pi2\<bullet>(\<lambda>a'. Cut <a'>.(NotR (y).M a') (x).P))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   580
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   581
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   582
apply(subgoal_tac "\<exists>n::coname. n\<sharp>(P,M,pi1\<bullet>P,pi1\<bullet>M,pi1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   583
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   584
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   585
apply(simp add: fresh_fun_simp_NotR calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   586
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   587
apply(simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   588
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   589
apply(subgoal_tac "\<exists>n::coname. n\<sharp>(P,M,pi2\<bullet>P,pi2\<bullet>M,pi2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   590
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   591
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   592
apply(simp add: fresh_fun_simp_NotR at_prm_fresh[OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   593
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   594
apply(simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   595
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   596
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   597
lemma fresh_fun_simp_AndR:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   598
  assumes a: "a'\<sharp>P" "a'\<sharp>M" "a'\<sharp>N" "a'\<sharp>b" "a'\<sharp>c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   599
  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"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   600
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   601
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   602
apply(rule fresh_fun_app)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   603
apply(rule pt_coname_inst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   604
apply(rule at_coname_inst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   605
apply(finite_guess)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   606
apply(subgoal_tac "\<exists>n::coname. n\<sharp>(x,P,b,M,c,N)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   607
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   608
apply(rule_tac x="n" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   609
apply(simp add: fresh_prod abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   610
apply(fresh_guess)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   611
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   612
apply(simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   613
apply(fresh_guess)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   614
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   615
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   616
lemma fresh_fun_AndR[eqvt_force]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   617
  fixes pi1::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   618
  and   pi2::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   619
  shows "pi1\<bullet>fresh_fun (\<lambda>a'. Cut <a'>.(AndR <b>.M <c>.N a') (x).P)=
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   620
             fresh_fun (pi1\<bullet>(\<lambda>a'. Cut <a'>.(AndR <b>.M <c>.N a') (x).P))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   621
  and   "pi2\<bullet>fresh_fun (\<lambda>a'. Cut <a'>.(AndR <b>.M <c>.N a') (x).P)=
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   622
             fresh_fun (pi2\<bullet>(\<lambda>a'. Cut <a'>.(AndR <b>.M <c>.N a') (x).P))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   623
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   624
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   625
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)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   626
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   627
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   628
apply(simp add: fresh_fun_simp_AndR calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   629
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   630
apply(simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   631
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   632
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)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   633
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   634
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   635
apply(simp add: fresh_fun_simp_AndR at_prm_fresh[OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   636
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   637
apply(simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   638
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   639
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   640
lemma fresh_fun_simp_OrR1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   641
  assumes a: "a'\<sharp>P" "a'\<sharp>M" "a'\<sharp>b" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   642
  shows "fresh_fun (\<lambda>a'. Cut <a'>.(OrR1 <b>.M a') (x).P) = Cut <a'>.(OrR1 <b>.M a') (x).P"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   643
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   644
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   645
apply(rule fresh_fun_app)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   646
apply(rule pt_coname_inst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   647
apply(rule at_coname_inst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   648
apply(finite_guess)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   649
apply(subgoal_tac "\<exists>n::coname. n\<sharp>(x,P,b,M)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   650
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   651
apply(rule_tac x="n" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   652
apply(simp add: fresh_prod abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   653
apply(fresh_guess)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   654
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   655
apply(simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   656
apply(fresh_guess)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   657
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   658
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   659
lemma fresh_fun_OrR1[eqvt_force]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   660
  fixes pi1::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   661
  and   pi2::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   662
  shows "pi1\<bullet>fresh_fun (\<lambda>a'. Cut <a'>.(OrR1 <b>.M a') (x).P)=
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   663
             fresh_fun (pi1\<bullet>(\<lambda>a'. Cut <a'>.(OrR1 <b>.M  a') (x).P))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   664
  and   "pi2\<bullet>fresh_fun (\<lambda>a'. Cut <a'>.(OrR1 <b>.M a') (x).P)=
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   665
             fresh_fun (pi2\<bullet>(\<lambda>a'. Cut <a'>.(OrR1 <b>.M a') (x).P))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   666
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   667
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   668
apply(subgoal_tac "\<exists>n::coname. n\<sharp>(P,M,b,pi1\<bullet>P,pi1\<bullet>M,pi1\<bullet>b,pi1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   669
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   670
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   671
apply(simp add: fresh_fun_simp_OrR1 calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   672
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   673
apply(simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   674
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   675
apply(subgoal_tac "\<exists>n::coname. n\<sharp>(P,M,b,pi2\<bullet>P,pi2\<bullet>M,pi2\<bullet>b,pi2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   676
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   677
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   678
apply(simp add: fresh_fun_simp_OrR1 at_prm_fresh[OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   679
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   680
apply(simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   681
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   682
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   683
lemma fresh_fun_simp_OrR2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   684
  assumes a: "a'\<sharp>P" "a'\<sharp>M" "a'\<sharp>b" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   685
  shows "fresh_fun (\<lambda>a'. Cut <a'>.(OrR2 <b>.M a') (x).P) = Cut <a'>.(OrR2 <b>.M a') (x).P"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   686
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   687
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   688
apply(rule fresh_fun_app)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   689
apply(rule pt_coname_inst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   690
apply(rule at_coname_inst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   691
apply(finite_guess)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   692
apply(subgoal_tac "\<exists>n::coname. n\<sharp>(x,P,b,M)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   693
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   694
apply(rule_tac x="n" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   695
apply(simp add: fresh_prod abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   696
apply(fresh_guess)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   697
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   698
apply(simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   699
apply(fresh_guess)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   700
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   701
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   702
lemma fresh_fun_OrR2[eqvt_force]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   703
  fixes pi1::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   704
  and   pi2::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   705
  shows "pi1\<bullet>fresh_fun (\<lambda>a'. Cut <a'>.(OrR2 <b>.M a') (x).P)=
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   706
             fresh_fun (pi1\<bullet>(\<lambda>a'. Cut <a'>.(OrR2 <b>.M  a') (x).P))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   707
  and   "pi2\<bullet>fresh_fun (\<lambda>a'. Cut <a'>.(OrR2 <b>.M a') (x).P)=
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   708
             fresh_fun (pi2\<bullet>(\<lambda>a'. Cut <a'>.(OrR2 <b>.M a') (x).P))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   709
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   710
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   711
apply(subgoal_tac "\<exists>n::coname. n\<sharp>(P,M,b,pi1\<bullet>P,pi1\<bullet>M,pi1\<bullet>b,pi1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   712
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   713
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   714
apply(simp add: fresh_fun_simp_OrR2 calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   715
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   716
apply(simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   717
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   718
apply(subgoal_tac "\<exists>n::coname. n\<sharp>(P,M,b,pi2\<bullet>P,pi2\<bullet>M,pi2\<bullet>b,pi2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   719
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   720
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   721
apply(simp add: fresh_fun_simp_OrR2 at_prm_fresh[OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   722
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   723
apply(simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   724
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   725
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   726
lemma fresh_fun_simp_ImpR:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   727
  assumes a: "a'\<sharp>P" "a'\<sharp>M" "a'\<sharp>b" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   728
  shows "fresh_fun (\<lambda>a'. Cut <a'>.(ImpR (y).<b>.M a') (x).P) = Cut <a'>.(ImpR (y).<b>.M a') (x).P"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   729
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   730
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   731
apply(rule fresh_fun_app)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   732
apply(rule pt_coname_inst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   733
apply(rule at_coname_inst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   734
apply(finite_guess)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   735
apply(subgoal_tac "\<exists>n::coname. n\<sharp>(x,P,y,b,M)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   736
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   737
apply(rule_tac x="n" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   738
apply(simp add: fresh_prod abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   739
apply(fresh_guess)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   740
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   741
apply(simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   742
apply(fresh_guess)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   743
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   744
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   745
lemma fresh_fun_ImpR[eqvt_force]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   746
  fixes pi1::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   747
  and   pi2::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   748
  shows "pi1\<bullet>fresh_fun (\<lambda>a'. Cut <a'>.(ImpR (y).<b>.M a') (x).P)=
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   749
             fresh_fun (pi1\<bullet>(\<lambda>a'. Cut <a'>.(ImpR (y).<b>.M  a') (x).P))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   750
  and   "pi2\<bullet>fresh_fun (\<lambda>a'. Cut <a'>.(ImpR (y).<b>.M a') (x).P)=
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   751
             fresh_fun (pi2\<bullet>(\<lambda>a'. Cut <a'>.(ImpR (y).<b>.M a') (x).P))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   752
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   753
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   754
apply(subgoal_tac "\<exists>n::coname. n\<sharp>(P,M,b,pi1\<bullet>P,pi1\<bullet>M,pi1\<bullet>b,pi1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   755
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   756
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   757
apply(simp add: fresh_fun_simp_ImpR calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   758
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   759
apply(simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   760
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   761
apply(subgoal_tac "\<exists>n::coname. n\<sharp>(P,M,b,pi2\<bullet>P,pi2\<bullet>M,pi2\<bullet>b,pi2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   762
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   763
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   764
apply(simp add: fresh_fun_simp_ImpR at_prm_fresh[OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   765
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   766
apply(simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   767
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   768
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   769
consts
23092
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   770
  substn :: "trm \<Rightarrow> name   \<Rightarrow> coname \<Rightarrow> trm \<Rightarrow> trm" ("_{_:=<_>._}" [100,100,100,100] 100) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   771
  substc :: "trm \<Rightarrow> coname \<Rightarrow> name   \<Rightarrow> trm \<Rightarrow> trm" ("_{_:=(_)._}" [100,100,100,100] 100)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   772
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   773
nominal_primrec (freshness_context: "(y::name,c::coname,P::trm)")
23092
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   774
  "(Ax x a){y:=<c>.P} = (if x=y then Cut <c>.P (y).Ax y a else Ax x a)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   775
  "\<lbrakk>a\<sharp>(c,P,N);x\<sharp>(y,P,M)\<rbrakk> \<Longrightarrow> (Cut <a>.M (x).N){y:=<c>.P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   776
  (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}))" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   777
  "x\<sharp>(y,P) \<Longrightarrow> (NotR (x).M a){y:=<c>.P} = NotR (x).(M{y:=<c>.P}) a" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   778
  "a\<sharp>(c,P) \<Longrightarrow> (NotL <a>.M x){y:=<c>.P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   779
  (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)"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   780
  "\<lbrakk>a\<sharp>(c,P,N,d);b\<sharp>(c,P,M,d);b\<noteq>a\<rbrakk> \<Longrightarrow> 
23092
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   781
  (AndR <a>.M <b>.N d){y:=<c>.P} = AndR <a>.(M{y:=<c>.P}) <b>.(N{y:=<c>.P}) d" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   782
  "x\<sharp>(y,P,z) \<Longrightarrow> (AndL1 (x).M z){y:=<c>.P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   783
  (if z=y then fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL1 (x).(M{y:=<c>.P}) z') 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   784
   else AndL1 (x).(M{y:=<c>.P}) z)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   785
  "x\<sharp>(y,P,z) \<Longrightarrow> (AndL2 (x).M z){y:=<c>.P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   786
  (if z=y then fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL2 (x).(M{y:=<c>.P}) z') 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   787
   else AndL2 (x).(M{y:=<c>.P}) z)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   788
  "a\<sharp>(c,P,b) \<Longrightarrow> (OrR1 <a>.M b){y:=<c>.P} = OrR1 <a>.(M{y:=<c>.P}) b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   789
  "a\<sharp>(c,P,b) \<Longrightarrow> (OrR2 <a>.M b){y:=<c>.P} = OrR2 <a>.(M{y:=<c>.P}) b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   790
  "\<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} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   791
  (if z=y then fresh_fun (\<lambda>z'. Cut <c>.P (z').OrL (x).(M{y:=<c>.P}) (u).(N{y:=<c>.P}) z') 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   792
   else OrL (x).(M{y:=<c>.P}) (u).(N{y:=<c>.P}) z)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   793
  "\<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"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   794
  "\<lbrakk>a\<sharp>(N,c,P);x\<sharp>(y,P,M,z)\<rbrakk> \<Longrightarrow> (ImpL <a>.M (x).N z){y:=<c>.P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   795
  (if y=z then fresh_fun (\<lambda>z'. Cut <c>.P (z').ImpL <a>.(M{y:=<c>.P}) (x).(N{y:=<c>.P}) z') 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   796
   else ImpL <a>.(M{y:=<c>.P}) (x).(N{y:=<c>.P}) z)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   797
apply(finite_guess)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   798
apply(rule TrueI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   799
apply(simp add: abs_fresh abs_supp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   800
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   801
apply(subgoal_tac "\<exists>x::name. x\<sharp>(x1,P,y1)", erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   802
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   803
apply(simp add: fresh_fun_simp_NotL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   804
apply(rule exists_fresh', simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   805
apply(simp add: abs_fresh abs_supp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   806
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   807
apply(subgoal_tac "\<exists>x::name. x\<sharp>(x1,P,y1)", erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   808
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   809
apply(simp add: fresh_fun_simp_AndL1 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   810
apply(rule exists_fresh', simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   811
apply(simp add: abs_fresh abs_supp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   812
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   813
apply(subgoal_tac "\<exists>x::name. x\<sharp>(x1,P,y1)", erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   814
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   815
apply(simp add: fresh_fun_simp_AndL2 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   816
apply(rule exists_fresh', simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   817
apply(simp add: abs_fresh abs_supp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   818
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   819
apply(subgoal_tac "\<exists>x::name. x\<sharp>(x1,P,y1,x3,y2)", erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   820
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   821
apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   822
apply(rule exists_fresh', simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   823
apply(simp add: abs_fresh abs_supp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   824
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   825
apply(subgoal_tac "\<exists>x::name. x\<sharp>(x1,P,y1,x3,y2)", erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   826
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   827
apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   828
apply(rule exists_fresh', simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   829
apply(simp add: abs_fresh abs_supp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   830
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   831
apply(subgoal_tac "\<exists>x::name. x\<sharp>(x3,P,y1,y2)", erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   832
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   833
apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   834
apply(rule exists_fresh', simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   835
apply(simp add: abs_fresh abs_supp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   836
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   837
apply(subgoal_tac "\<exists>x::name. x\<sharp>(x3,P,y1,y2)", erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   838
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   839
apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   840
apply(rule exists_fresh', simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   841
apply(fresh_guess)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   842
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   843
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   844
nominal_primrec (freshness_context: "(d::name,z::coname,P::trm)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   845
  "(Ax x a){d:=(z).P} = (if d=a then Cut <a>.(Ax x a) (z).P else Ax x a)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   846
  "\<lbrakk>a\<sharp>(d,P,N);x\<sharp>(z,P,M)\<rbrakk> \<Longrightarrow> (Cut <a>.M (x).N){d:=(z).P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   847
  (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}))" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   848
  "x\<sharp>(z,P) \<Longrightarrow> (NotR (x).M a){d:=(z).P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   849
  (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)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   850
  "a\<sharp>(d,P) \<Longrightarrow> (NotL <a>.M x){d:=(z).P} = NotL <a>.(M{d:=(z).P}) x" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   851
  "\<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} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   852
  (if d=c then fresh_fun (\<lambda>a'. Cut <a'>.(AndR <a>.(M{d:=(z).P}) <b>.(N{d:=(z).P}) a') (z).P)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   853
   else AndR <a>.(M{d:=(z).P}) <b>.(N{d:=(z).P}) c)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   854
  "x\<sharp>(y,z,P) \<Longrightarrow> (AndL1 (x).M y){d:=(z).P} = AndL1 (x).(M{d:=(z).P}) y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   855
  "x\<sharp>(y,P,z) \<Longrightarrow> (AndL2 (x).M y){d:=(z).P} = AndL2 (x).(M{d:=(z).P}) y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   856
  "a\<sharp>(d,P,b) \<Longrightarrow> (OrR1 <a>.M b){d:=(z).P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   857
  (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)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   858
  "a\<sharp>(d,P,b) \<Longrightarrow> (OrR2 <a>.M b){d:=(z).P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   859
  (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)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   860
  "\<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} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   861
  OrL (x).(M{d:=(z).P}) (y).(N{d:=(z).P}) u" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   862
  "\<lbrakk>a\<sharp>(b,d,P); x\<sharp>(z,P)\<rbrakk> \<Longrightarrow> (ImpR (x).<a>.M b){d:=(z).P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   863
  (if d=b then fresh_fun (\<lambda>a'. Cut <a'>.ImpR (x).<a>.(M{d:=(z).P}) a' (z).P) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   864
   else ImpR (x).<a>.(M{d:=(z).P}) b)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   865
  "\<lbrakk>a\<sharp>(N,d,P);x\<sharp>(y,z,P,M)\<rbrakk> \<Longrightarrow> (ImpL <a>.M (x).N y){d:=(z).P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   866
  ImpL <a>.(M{d:=(z).P}) (x).(N{d:=(z).P}) y"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   867
apply(finite_guess)+
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   868
apply(rule TrueI)+
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
   869
apply(simp add: abs_fresh abs_supp fs_name1 fs_coname1)+
23092
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   870
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   871
apply(subgoal_tac "\<exists>x::coname. x\<sharp>(x1,P,y1)", erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   872
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   873
apply(simp add: fresh_fun_simp_NotR abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   874
apply(rule exists_fresh', simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   875
apply(simp add: abs_fresh abs_supp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   876
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   877
apply(subgoal_tac "\<exists>x::coname. x\<sharp>(x1,P,y1,x3,y2)", erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   878
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   879
apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   880
apply(rule exists_fresh', simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   881
apply(simp add: abs_fresh abs_supp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   882
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   883
apply(subgoal_tac "\<exists>x::coname. x\<sharp>(x1,P,y1,x3,y2)", erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   884
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   885
apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   886
apply(rule exists_fresh', simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   887
apply(simp add: abs_fresh abs_supp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   888
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   889
apply(subgoal_tac "\<exists>x::coname. x\<sharp>(x1,P,y1)", erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   890
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   891
apply(simp add: fresh_fun_simp_OrR1 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   892
apply(rule exists_fresh', simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   893
apply(simp add: abs_fresh abs_supp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   894
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   895
apply(subgoal_tac "\<exists>x::coname. x\<sharp>(x1,P,y1)", erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   896
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   897
apply(simp add: fresh_fun_simp_OrR2 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   898
apply(rule exists_fresh', simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   899
apply(simp add: abs_fresh abs_supp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   900
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   901
apply(subgoal_tac "\<exists>x::coname. x\<sharp>(x1,P,x2,y1)", erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   902
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   903
apply(simp add: fresh_fun_simp_ImpR abs_fresh fresh_atm abs_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   904
apply(rule exists_fresh', simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   905
apply(simp add: abs_fresh abs_supp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   906
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   907
apply(subgoal_tac "\<exists>x::coname. x\<sharp>(x1,P,x2,y1)", erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   908
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   909
apply(simp add: fresh_fun_simp_ImpR abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   910
apply(rule exists_fresh', simp add: fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   911
apply(simp add: abs_fresh abs_supp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   912
apply(fresh_guess add: abs_fresh fresh_prod)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   913
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   914
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   915
lemma csubst_eqvt[eqvt]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   916
  fixes pi1::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   917
  and   pi2::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   918
  shows "pi1\<bullet>(M{c:=(x).N}) = (pi1\<bullet>M){(pi1\<bullet>c):=(pi1\<bullet>x).(pi1\<bullet>N)}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   919
  and   "pi2\<bullet>(M{c:=(x).N}) = (pi2\<bullet>M){(pi2\<bullet>c):=(pi2\<bullet>x).(pi2\<bullet>N)}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   920
apply(nominal_induct M avoiding: c x N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   921
apply(auto simp add: eq_bij fresh_bij eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   922
apply(perm_simp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   923
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   924
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   925
lemma nsubst_eqvt[eqvt]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   926
  fixes pi1::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   927
  and   pi2::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   928
  shows "pi1\<bullet>(M{x:=<c>.N}) = (pi1\<bullet>M){(pi1\<bullet>x):=<(pi1\<bullet>c)>.(pi1\<bullet>N)}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   929
  and   "pi2\<bullet>(M{x:=<c>.N}) = (pi2\<bullet>M){(pi2\<bullet>x):=<(pi2\<bullet>c)>.(pi2\<bullet>N)}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   930
apply(nominal_induct M avoiding: c x N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   931
apply(auto simp add: eq_bij fresh_bij eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   932
apply(perm_simp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   933
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   934
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   935
lemma supp_subst1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   936
  shows "supp (M{y:=<c>.P}) \<subseteq> ((supp M) - {y}) \<union> (supp P)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   937
apply(nominal_induct M avoiding: y P c rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   938
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   939
apply(auto simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   940
apply(blast)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   941
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<c>.P},P)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   942
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   943
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   944
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   945
apply(simp add: fresh_fun_simp_NotL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   946
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   947
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   948
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   949
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<c>.P},P)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   950
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   951
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   952
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   953
apply(simp add: fresh_fun_simp_NotL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   954
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   955
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   956
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   957
apply(blast)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   958
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<c>.P},P,name1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   959
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   960
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   961
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   962
apply(simp add: fresh_fun_simp_AndL1 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   963
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   964
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   965
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   966
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<c>.P},P,name1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   967
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   968
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   969
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   970
apply(simp add: fresh_fun_simp_AndL1 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   971
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   972
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   973
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<c>.P},P,name1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   974
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   975
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   976
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   977
apply(simp add: fresh_fun_simp_AndL1 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   978
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   979
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   980
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   981
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<c>.P},P,name1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   982
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   983
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   984
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   985
apply(simp add: fresh_fun_simp_AndL2 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   986
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   987
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   988
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   989
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<c>.P},P,name1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   990
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   991
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   992
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   993
apply(simp add: fresh_fun_simp_AndL2 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   994
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   995
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   996
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<c>.P},P,name1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   997
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   998
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
   999
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1000
apply(simp add: fresh_fun_simp_AndL2 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1001
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1002
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1003
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1004
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{y:=<c>.P},P,name1,trm2{y:=<c>.P},name2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1005
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1006
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1007
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1008
apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1009
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1010
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1011
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1012
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{y:=<c>.P},P,name1,trm2{y:=<c>.P},name2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1013
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1014
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1015
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1016
apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1017
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1018
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1019
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1020
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{y:=<c>.P},P,name1,trm2{y:=<c>.P},name2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1021
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1022
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1023
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1024
apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1025
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1026
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1027
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1028
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{y:=<c>.P},P,name1,trm2{y:=<c>.P},name2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1029
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1030
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1031
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1032
apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1033
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1034
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1035
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1036
apply(blast)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1037
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{name2:=<c>.P},P,name1,trm2{name2:=<c>.P})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1038
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1039
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1040
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1041
apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1042
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1043
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1044
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1045
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{name2:=<c>.P},P,name1,trm2{name2:=<c>.P})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1046
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1047
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1048
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1049
apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1050
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1051
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1052
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1053
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{name2:=<c>.P},P,name1,trm2{name2:=<c>.P})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1054
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1055
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1056
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1057
apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1058
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1059
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1060
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1061
apply(blast)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1062
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1063
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1064
lemma supp_subst2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1065
  shows "supp (M{y:=<c>.P}) \<subseteq> supp (M) \<union> ((supp P) - {c})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1066
apply(nominal_induct M avoiding: y P c rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1067
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1068
apply(auto simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1069
apply(blast)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1070
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<c>.P},P)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1071
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1072
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1073
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1074
apply(simp add: fresh_fun_simp_NotL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1075
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1076
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1077
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1078
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<c>.P},P)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1079
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1080
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1081
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1082
apply(simp add: fresh_fun_simp_NotL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1083
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1084
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1085
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<c>.P},P)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1086
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1087
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1088
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1089
apply(simp add: fresh_fun_simp_NotL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1090
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1091
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1092
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1093
apply(blast)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1094
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<c>.P},P,name1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1095
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1096
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1097
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1098
apply(simp add: fresh_fun_simp_AndL1 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1099
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1100
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1101
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1102
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<c>.P},P,name1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1103
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1104
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1105
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1106
apply(simp add: fresh_fun_simp_AndL1 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1107
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1108
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1109
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1110
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<c>.P},P,name1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1111
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1112
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1113
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1114
apply(simp add: fresh_fun_simp_AndL2 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1115
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1116
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1117
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1118
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<c>.P},P,name1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1119
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1120
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1121
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1122
apply(simp add: fresh_fun_simp_AndL2 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1123
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1124
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1125
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1126
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{y:=<c>.P},P,name1,trm2{y:=<c>.P},name2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1127
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1128
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1129
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1130
apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1131
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1132
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1133
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1134
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{y:=<c>.P},P,name1,trm2{y:=<c>.P},name2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1135
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1136
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1137
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1138
apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1139
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1140
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1141
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1142
apply(blast)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1143
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{name2:=<c>.P},P,name1,trm2{name2:=<c>.P})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1144
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1145
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1146
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1147
apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1148
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1149
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1150
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1151
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{name2:=<c>.P},P,name1,trm2{name2:=<c>.P})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1152
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1153
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1154
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1155
apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1156
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1157
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1158
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1159
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{name2:=<c>.P},P,name1,trm2{name2:=<c>.P})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1160
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1161
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1162
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1163
apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1164
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1165
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1166
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{name2:=<c>.P},P,name1,trm2{name2:=<c>.P})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1167
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1168
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1169
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1170
apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1171
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1172
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1173
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1174
apply(blast)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1175
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1176
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1177
lemma supp_subst3:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1178
  shows "supp (M{c:=(x).P}) \<subseteq> ((supp M) - {c}) \<union> (supp P)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1179
apply(nominal_induct M avoiding: x P c rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1180
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1181
apply(auto simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1182
apply(blast)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1183
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname:=(x).P},P)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1184
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1185
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1186
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1187
apply(simp add: fresh_fun_simp_NotR abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1188
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1189
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1190
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1191
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname:=(x).P},P)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1192
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1193
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1194
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1195
apply(simp add: fresh_fun_simp_NotR abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1196
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1197
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1198
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1199
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1200
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1201
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1202
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1203
apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1204
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1205
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1206
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1207
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1208
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1209
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1210
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1211
apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1212
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1213
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1214
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1215
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1216
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1217
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1218
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1219
apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1220
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1221
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1222
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1223
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1224
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1225
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1226
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1227
apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1228
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1229
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1230
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1231
apply(blast)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1232
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(x).P},P,coname1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1233
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1234
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1235
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1236
apply(simp add: fresh_fun_simp_OrR1 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1237
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1238
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1239
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1240
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(x).P},P,coname1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1241
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1242
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1243
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1244
apply(simp add: fresh_fun_simp_OrR1 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1245
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1246
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1247
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(x).P},P,coname1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1248
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1249
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1250
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1251
apply(simp add: fresh_fun_simp_OrR1 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1252
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1253
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1254
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1255
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(x).P},P,coname1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1256
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1257
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1258
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1259
apply(simp add: fresh_fun_simp_OrR2 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1260
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1261
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1262
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1263
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(x).P},P,coname1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1264
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1265
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1266
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1267
apply(simp add: fresh_fun_simp_OrR2 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1268
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1269
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1270
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(x).P},P,coname1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1271
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1272
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1273
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1274
apply(simp add: fresh_fun_simp_OrR2 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1275
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1276
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1277
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1278
apply(blast)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1279
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(x).P},P,coname1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1280
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1281
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1282
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1283
apply(simp add: fresh_fun_simp_ImpR abs_fresh abs_supp fin_supp fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1284
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1285
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1286
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1287
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(x).P},P,coname1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1288
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1289
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1290
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1291
apply(simp add: fresh_fun_simp_ImpR abs_fresh abs_supp fin_supp fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1292
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1293
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1294
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(x).P},P,coname1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1295
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1296
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1297
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1298
apply(simp add: fresh_fun_simp_ImpR abs_fresh abs_supp fin_supp fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1299
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1300
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1301
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1302
apply(blast)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1303
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1304
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1305
lemma supp_subst4:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1306
  shows "supp (M{c:=(x).P}) \<subseteq> (supp M) \<union> ((supp P) - {x})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1307
apply(nominal_induct M avoiding: x P c rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1308
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1309
apply(auto simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1310
apply(blast)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1311
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname:=(x).P},P)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1312
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1313
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1314
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1315
apply(simp add: fresh_fun_simp_NotR abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1316
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1317
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1318
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1319
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname:=(x).P},P)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1320
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1321
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1322
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1323
apply(simp add: fresh_fun_simp_NotR abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1324
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1325
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1326
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname:=(x).P},P)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1327
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1328
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1329
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1330
apply(simp add: fresh_fun_simp_NotR abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1331
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1332
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1333
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1334
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1335
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1336
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1337
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1338
apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1339
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1340
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1341
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1342
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1343
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1344
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1345
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1346
apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1347
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1348
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1349
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1350
apply(blast)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1351
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(x).P},P,coname1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1352
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1353
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1354
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1355
apply(simp add: fresh_fun_simp_OrR1 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1356
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1357
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1358
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1359
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(x).P},P,coname1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1360
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1361
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1362
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1363
apply(simp add: fresh_fun_simp_OrR1 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1364
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1365
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1366
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1367
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(x).P},P,coname1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1368
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1369
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1370
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1371
apply(simp add: fresh_fun_simp_OrR2 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1372
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1373
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1374
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1375
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(x).P},P,coname1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1376
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1377
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1378
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1379
apply(simp add: fresh_fun_simp_OrR2 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1380
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1381
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1382
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1383
apply(blast)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1384
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(x).P},P,coname1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1385
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1386
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1387
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1388
apply(simp add: fresh_fun_simp_ImpR abs_fresh abs_supp fin_supp fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1389
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1390
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1391
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1392
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(x).P},P,coname1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1393
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1394
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1395
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1396
apply(simp add: fresh_fun_simp_ImpR abs_fresh abs_supp fin_supp fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1397
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1398
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1399
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(x).P},P,coname1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1400
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1401
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1402
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1403
apply(simp add: fresh_fun_simp_ImpR abs_fresh abs_supp fin_supp fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1404
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1405
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1406
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1407
apply(blast)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1408
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1409
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1410
lemma supp_subst5:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1411
  shows "(supp M - {y}) \<subseteq> supp (M{y:=<c>.P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1412
apply(nominal_induct M avoiding: y P c rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1413
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1414
apply(auto simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1415
apply(blast)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1416
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<c>.P},P)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1417
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1418
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1419
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1420
apply(simp add: fresh_fun_simp_NotL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1421
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1422
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1423
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1424
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1425
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<c>.P},P,name1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1426
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1427
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1428
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1429
apply(simp add: fresh_fun_simp_AndL1 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1430
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1431
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1432
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1433
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<c>.P},P,name1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1434
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1435
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1436
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1437
apply(simp add: fresh_fun_simp_AndL2 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1438
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1439
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1440
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1441
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{y:=<c>.P},P,name1,trm2{y:=<c>.P},name2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1442
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1443
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1444
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1445
apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1446
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1447
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1448
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1449
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{y:=<c>.P},P,name1,trm2{y:=<c>.P},name2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1450
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1451
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1452
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1453
apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1454
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1455
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1456
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1457
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1458
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{name2:=<c>.P},P,name1,trm2{name2:=<c>.P})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1459
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1460
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1461
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1462
apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1463
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1464
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1465
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1466
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{name2:=<c>.P},P,name1,trm2{name2:=<c>.P})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1467
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1468
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1469
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1470
apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1471
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1472
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1473
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1474
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1475
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1476
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1477
lemma supp_subst6:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1478
  shows "(supp M) \<subseteq> ((supp (M{y:=<c>.P}))::coname set)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1479
apply(nominal_induct M avoiding: y P c rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1480
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1481
apply(auto simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1482
apply(blast)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1483
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<c>.P},P)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1484
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1485
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1486
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1487
apply(simp add: fresh_fun_simp_NotL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1488
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1489
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1490
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1491
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1492
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<c>.P},P,name1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1493
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1494
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1495
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1496
apply(simp add: fresh_fun_simp_AndL1 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1497
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1498
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1499
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1500
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<c>.P},P,name1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1501
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1502
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1503
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1504
apply(simp add: fresh_fun_simp_AndL2 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1505
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1506
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1507
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1508
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{y:=<c>.P},P,name1,trm2{y:=<c>.P},name2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1509
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1510
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1511
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1512
apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1513
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1514
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1515
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1516
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{y:=<c>.P},P,name1,trm2{y:=<c>.P},name2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1517
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1518
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1519
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1520
apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1521
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1522
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1523
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1524
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1525
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{name2:=<c>.P},P,name1,trm2{name2:=<c>.P})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1526
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1527
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1528
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1529
apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1530
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1531
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1532
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1533
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{name2:=<c>.P},P,name1,trm2{name2:=<c>.P})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1534
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1535
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1536
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1537
apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1538
apply(simp add: fresh_def abs_supp trm.supp supp_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1539
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1540
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1541
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1542
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1543
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1544
lemma supp_subst7:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1545
  shows "(supp M - {c}) \<subseteq>  supp (M{c:=(x).P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1546
apply(nominal_induct M avoiding: x P c rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1547
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1548
apply(auto simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1549
apply(blast)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1550
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname:=(x).P},P)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1551
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1552
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1553
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1554
apply(simp add: fresh_fun_simp_NotR abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1555
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1556
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1557
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1558
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1559
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1560
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1561
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1562
apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1563
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1564
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1565
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1566
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1567
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1568
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1569
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1570
apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1571
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1572
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1573
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1574
apply(blast)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1575
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(x).P},P,coname1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1576
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1577
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1578
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1579
apply(simp add: fresh_fun_simp_OrR1 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1580
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1581
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1582
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1583
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(x).P},P,coname1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1584
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1585
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1586
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1587
apply(simp add: fresh_fun_simp_OrR2 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1588
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1589
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1590
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1591
apply(blast)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1592
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(x).P},P,coname1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1593
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1594
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1595
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1596
apply(simp add: fresh_fun_simp_ImpR abs_fresh abs_supp fin_supp fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1597
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1598
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1599
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1600
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1601
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1602
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1603
lemma supp_subst8:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1604
  shows "(supp M) \<subseteq> ((supp (M{c:=(x).P}))::name set)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1605
apply(nominal_induct M avoiding: x P c rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1606
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1607
apply(auto simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1608
apply(blast)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1609
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname:=(x).P},P)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1610
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1611
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1612
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1613
apply(simp add: fresh_fun_simp_NotR abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1614
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1615
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1616
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1617
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1618
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1619
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1620
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1621
apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1622
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1623
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1624
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1625
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1626
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1627
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1628
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1629
apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1630
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1631
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1632
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1633
apply(blast)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1634
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(x).P},P,coname1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1635
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1636
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1637
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1638
apply(simp add: fresh_fun_simp_OrR1 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1639
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1640
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1641
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1642
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(x).P},P,coname1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1643
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1644
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1645
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1646
apply(simp add: fresh_fun_simp_OrR2 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1647
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1648
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1649
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1650
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1651
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(x).P},P,coname1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1652
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1653
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1654
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1655
apply(simp add: fresh_fun_simp_ImpR abs_fresh abs_supp fin_supp fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1656
apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1657
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1658
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1659
apply(blast)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1660
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1661
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1662
lemmas subst_supp = supp_subst1 supp_subst2 supp_subst3 supp_subst4
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1663
                    supp_subst5 supp_subst6 supp_subst7 supp_subst8
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1664
lemma subst_fresh:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1665
  fixes x::"name"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1666
  and   c::"coname"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1667
  shows "x\<sharp>P \<Longrightarrow> x\<sharp>M{x:=<c>.P}"   
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1668
  and   "b\<sharp>P \<Longrightarrow> b\<sharp>M{b:=(y).P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1669
  and   "x\<sharp>(M,P) \<Longrightarrow> x\<sharp>M{y:=<c>.P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1670
  and   "x\<sharp>M \<Longrightarrow> x\<sharp>M{c:=(x).P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1671
  and   "x\<sharp>(M,P) \<Longrightarrow> x\<sharp>M{c:=(y).P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1672
  and   "b\<sharp>(M,P) \<Longrightarrow> b\<sharp>M{c:=(y).P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1673
  and   "b\<sharp>M \<Longrightarrow> b\<sharp>M{y:=<b>.P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1674
  and   "b\<sharp>(M,P) \<Longrightarrow> b\<sharp>M{y:=<c>.P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1675
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1676
apply(insert subst_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1677
apply(simp_all add: fresh_def supp_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1678
apply(blast)+ 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1679
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1680
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1681
lemma forget:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1682
  shows "x\<sharp>M \<Longrightarrow> M{x:=<c>.P} = M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1683
  and   "c\<sharp>M \<Longrightarrow> M{c:=(x).P} = M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1684
apply(nominal_induct M avoiding: x c P rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1685
apply(auto simp add: fresh_atm abs_fresh abs_supp fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1686
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1687
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1688
lemma substc_rename1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1689
  assumes a: "c\<sharp>(M,a)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1690
  shows "M{a:=(x).N} = ([(c,a)]\<bullet>M){c:=(x).N}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1691
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1692
proof(nominal_induct M avoiding: c a x N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1693
  case (Ax z d)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1694
  then show ?case by (auto simp add: fresh_prod fresh_atm calc_atm trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1695
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1696
  case NotL
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1697
  then show ?case by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1698
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1699
  case (NotR y M d)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1700
  then show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1701
    apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1702
    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(N,M{d:=(x).N},([(c,d)]\<bullet>M){c:=(x).N})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1703
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1704
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1705
    apply(simp add: fresh_fun_simp_NotR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1706
    apply(simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1707
    apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1708
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1709
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1710
  case (AndR c1 M c2 M' c3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1711
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1712
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1713
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1714
    apply(simp add: fresh_prod calc_atm fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1715
    apply(simp add: fresh_prod calc_atm fresh_atm abs_fresh fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1716
    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(N,M{c3:=(x).N},
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1717
                  M'{c3:=(x).N},c1,c2,c3,([(c,c3)]\<bullet>M){c:=(x).N},([(c,c3)]\<bullet>M'){c:=(x).N})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1718
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1719
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1720
    apply(simp add: fresh_fun_simp_AndR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1721
    apply (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1722
    apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1723
    apply(simp add: fresh_prod calc_atm fresh_atm abs_fresh fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1724
    apply(simp add: fresh_prod calc_atm fresh_atm abs_fresh fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1725
    apply(auto simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1726
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1727
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1728
  case AndL1
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1729
  then show ?case by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1730
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1731
  case AndL2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1732
  then show ?case by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1733
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1734
  case (OrR1 d M e)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1735
  then show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1736
    apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1737
    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(N,M{e:=(x).N},([(c,e)]\<bullet>M){c:=(x).N},d,e)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1738
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1739
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1740
    apply(simp add: fresh_fun_simp_OrR1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1741
    apply(simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1742
    apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1743
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1744
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1745
  case (OrR2 d M e)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1746
  then show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1747
    apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1748
    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(N,M{e:=(x).N},([(c,e)]\<bullet>M){c:=(x).N},d,e)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1749
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1750
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1751
    apply(simp add: fresh_fun_simp_OrR2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1752
    apply(simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1753
    apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1754
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1755
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1756
  case (OrL x1 M x2 M' x3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1757
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1758
    by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1759
next 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1760
  case ImpL
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1761
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1762
    by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1763
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1764
  case (ImpR y d M e)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1765
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1766
    apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1767
    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(N,M{e:=(x).N},([(c,e)]\<bullet>M){c:=(x).N},d,e)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1768
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1769
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1770
    apply(simp add: fresh_fun_simp_ImpR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1771
    apply(simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1772
    apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1773
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1774
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1775
  case (Cut d M y M')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1776
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1777
    apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1778
    apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1779
    apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1780
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1781
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1782
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1783
lemma substc_rename2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1784
  assumes a: "y\<sharp>(N,x)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1785
  shows "M{a:=(x).N} = M{a:=(y).([(y,x)]\<bullet>N)}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1786
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1787
proof(nominal_induct M avoiding: a x y N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1788
  case (Ax z d)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1789
  then show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1790
    by (auto simp add: fresh_prod fresh_atm calc_atm trm.inject alpha perm_swap fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1791
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1792
  case NotL
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1793
  then show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1794
    by (auto simp add: fresh_prod fresh_atm calc_atm trm.inject alpha perm_swap fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1795
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1796
  case (NotR y M d)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1797
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1798
    apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1799
    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(N,M{d:=(y).([(y,x)]\<bullet>N)},[(y,x)]\<bullet>N)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1800
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1801
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1802
    apply(simp add: fresh_fun_simp_NotR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1803
    apply(simp add: trm.inject alpha perm_swap fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1804
    apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1805
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1806
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1807
  case (AndR c1 M c2 M' c3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1808
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1809
    apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1810
    apply(subgoal_tac 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1811
       "\<exists>a'::coname. a'\<sharp>(N,M{c3:=(y).([(y,x)]\<bullet>N)},M'{c3:=(y).([(y,x)]\<bullet>N)},[(y,x)]\<bullet>N,c1,c2,c3)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1812
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1813
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1814
    apply(simp add: fresh_fun_simp_AndR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1815
    apply (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh subst_fresh perm_swap fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1816
    apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1817
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1818
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1819
  case AndL1
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1820
  then show ?case by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1821
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1822
  case AndL2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1823
  then show ?case by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1824
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1825
  case (OrR1 d M e)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1826
  then show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1827
    apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1828
    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(N,M{e:=(y).([(y,x)]\<bullet>N)},[(y,x)]\<bullet>N,d,e)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1829
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1830
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1831
    apply(simp add: fresh_fun_simp_OrR1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1832
    apply(simp add: trm.inject alpha perm_swap fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1833
    apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1834
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1835
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1836
  case (OrR2 d M e)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1837
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1838
    apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1839
    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(N,M{e:=(y).([(y,x)]\<bullet>N)},[(y,x)]\<bullet>N,d,e)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1840
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1841
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1842
    apply(simp add: fresh_fun_simp_OrR2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1843
    apply(simp add: trm.inject alpha perm_swap fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1844
    apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1845
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1846
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1847
  case (OrL x1 M x2 M' x3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1848
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1849
    by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1850
next 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1851
  case ImpL
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1852
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1853
    by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1854
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1855
  case (ImpR y d M e)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1856
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1857
    apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1858
    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(N,M{e:=(y).([(y,x)]\<bullet>N)},[(y,x)]\<bullet>N,d,e)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1859
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1860
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1861
    apply(simp add: fresh_fun_simp_ImpR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1862
    apply(simp add: trm.inject alpha perm_swap fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1863
    apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1864
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1865
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1866
  case (Cut d M y M')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1867
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1868
    by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left perm_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1869
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1870
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1871
lemma substn_rename3:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1872
  assumes a: "y\<sharp>(M,x)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1873
  shows "M{x:=<a>.N} = ([(y,x)]\<bullet>M){y:=<a>.N}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1874
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1875
proof(nominal_induct M avoiding: a x y N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1876
  case (Ax z d)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1877
  then show ?case by (auto simp add: fresh_prod fresh_atm calc_atm trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1878
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1879
  case NotR
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1880
  then show ?case by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1881
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1882
  case (NotL d M z)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1883
  then show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1884
    apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1885
    apply(subgoal_tac "\<exists>a'::name. a'\<sharp>(N,M{x:=<a>.N},([(y,x)]\<bullet>M){y:=<a>.N})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1886
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1887
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1888
    apply(simp add: fresh_fun_simp_NotL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1889
    apply(simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1890
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1891
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1892
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1893
  case (AndR c1 M c2 M' c3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1894
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1895
    by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1896
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1897
  case OrR1
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1898
  then show ?case by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1899
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1900
  case OrR2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1901
  then show ?case by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1902
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1903
  case (AndL1 u M v)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1904
  then show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1905
    apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1906
    apply(subgoal_tac "\<exists>a'::name. a'\<sharp>(N,M{x:=<a>.N},([(y,x)]\<bullet>M){y:=<a>.N},u,v)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1907
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1908
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1909
    apply(simp add: fresh_fun_simp_AndL1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1910
    apply(simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1911
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1912
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1913
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1914
  case (AndL2 u M v)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1915
  then show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1916
    apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1917
    apply(subgoal_tac "\<exists>a'::name. a'\<sharp>(N,M{x:=<a>.N},([(y,x)]\<bullet>M){y:=<a>.N},u,v)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1918
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1919
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1920
    apply(simp add: fresh_fun_simp_AndL2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1921
    apply(simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1922
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1923
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1924
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1925
  case (OrL x1 M x2 M' x3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1926
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1927
    apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1928
    apply(subgoal_tac 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1929
    "\<exists>a'::name. a'\<sharp>(N,M{x:=<a>.N},M'{x:=<a>.N},([(y,x)]\<bullet>M){y:=<a>.N},([(y,x)]\<bullet>M'){y:=<a>.N},x1,x2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1930
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1931
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1932
    apply(simp add: fresh_fun_simp_OrL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1933
    apply(simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1934
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1935
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1936
next 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1937
  case ImpR
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1938
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1939
  by(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_left abs_supp fin_supp fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1940
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1941
  case (ImpL d M v M' u)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1942
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1943
    apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1944
    apply(subgoal_tac 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1945
    "\<exists>a'::name. a'\<sharp>(N,M{u:=<a>.N},M'{u:=<a>.N},([(y,u)]\<bullet>M){y:=<a>.N},([(y,u)]\<bullet>M'){y:=<a>.N},d,v)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1946
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1947
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1948
    apply(simp add: fresh_fun_simp_ImpL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1949
    apply(simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1950
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1951
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1952
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1953
  case (Cut d M y M')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1954
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1955
    apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1956
    apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1957
    apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1958
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1959
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1960
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1961
lemma substn_rename4:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1962
  assumes a: "c\<sharp>(N,a)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1963
  shows "M{x:=<a>.N} = M{x:=<c>.([(c,a)]\<bullet>N)}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1964
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1965
proof(nominal_induct M avoiding: x c a N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1966
  case (Ax z d)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1967
  then show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1968
    by (auto simp add: fresh_prod fresh_atm calc_atm trm.inject alpha perm_swap fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1969
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1970
  case NotR
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1971
  then show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1972
    by (auto simp add: fresh_prod fresh_atm calc_atm trm.inject alpha perm_swap fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1973
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1974
  case (NotL d M y)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1975
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1976
    apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1977
    apply(subgoal_tac "\<exists>a'::name. a'\<sharp>(N,M{x:=<c>.([(c,a)]\<bullet>N)},[(c,a)]\<bullet>N)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1978
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1979
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1980
    apply(simp add: fresh_fun_simp_NotL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1981
    apply(simp add: trm.inject alpha perm_swap fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1982
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1983
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1984
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1985
  case (OrL x1 M x2 M' x3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1986
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1987
    apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1988
    apply(subgoal_tac 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1989
       "\<exists>a'::name. a'\<sharp>(N,M{x:=<c>.([(c,a)]\<bullet>N)},M'{x:=<c>.([(c,a)]\<bullet>N)},[(c,a)]\<bullet>N,x1,x2,x3)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1990
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1991
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1992
    apply(simp add: fresh_fun_simp_OrL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1993
    apply (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh subst_fresh perm_swap fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1994
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1995
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1996
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1997
  case OrR1
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1998
  then show ?case by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  1999
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2000
  case OrR2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2001
  then show ?case by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2002
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2003
  case (AndL1 u M v)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2004
  then show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2005
    apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2006
    apply(subgoal_tac "\<exists>a'::name. a'\<sharp>(N,M{x:=<c>.([(c,a)]\<bullet>N)},[(c,a)]\<bullet>N,u,v)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2007
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2008
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2009
    apply(simp add: fresh_fun_simp_AndL1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2010
    apply(simp add: trm.inject alpha perm_swap fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2011
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2012
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2013
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2014
  case (AndL2 u M v)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2015
  then show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2016
    apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2017
    apply(subgoal_tac "\<exists>a'::name. a'\<sharp>(N,M{x:=<c>.([(c,a)]\<bullet>N)},[(c,a)]\<bullet>N,u,v)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2018
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2019
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2020
    apply(simp add: fresh_fun_simp_AndL2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2021
    apply(simp add: trm.inject alpha perm_swap fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2022
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2023
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2024
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2025
  case (AndR c1 M c2 M' c3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2026
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2027
    by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2028
next 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2029
  case ImpR
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2030
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2031
    by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2032
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2033
  case (ImpL d M y M' u)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2034
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2035
    apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2036
    apply(subgoal_tac "\<exists>a'::name. a'\<sharp>(N,M{u:=<c>.([(c,a)]\<bullet>N)},M'{u:=<c>.([(c,a)]\<bullet>N)},[(c,a)]\<bullet>N,y,u)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2037
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2038
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2039
    apply(simp add: fresh_fun_simp_ImpL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2040
    apply(simp add: trm.inject alpha perm_swap fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2041
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2042
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2043
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2044
  case (Cut d M y M')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2045
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2046
    by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left perm_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2047
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2048
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2049
lemma subst_rename5:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2050
  assumes a: "c'\<sharp>(c,N)" "x'\<sharp>(x,M)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2051
  shows "M{x:=<c>.N} = ([(x',x)]\<bullet>M){x':=<c'>.([(c',c)]\<bullet>N)}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2052
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2053
  have "M{x:=<c>.N} = ([(x',x)]\<bullet>M){x':=<c>.N}" using a by (simp add: substn_rename3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2054
  also have "\<dots> = ([(x',x)]\<bullet>M){x':=<c'>.([(c',c)]\<bullet>N)}" using a by (simp add: substn_rename4)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2055
  finally show ?thesis by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2056
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2057
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2058
lemma subst_rename6:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2059
  assumes a: "c'\<sharp>(c,M)" "x'\<sharp>(x,N)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2060
  shows "M{c:=(x).N} = ([(c',c)]\<bullet>M){c':=(x').([(x',x)]\<bullet>N)}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2061
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2062
  have "M{c:=(x).N} = ([(c',c)]\<bullet>M){c':=(x).N}" using a by (simp add: substc_rename1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2063
  also have "\<dots> = ([(c',c)]\<bullet>M){c':=(x').([(x',x)]\<bullet>N)}" using a by (simp add: substc_rename2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2064
  finally show ?thesis by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2065
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2066
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2067
lemmas subst_rename = substc_rename1 substc_rename2 substn_rename3 substn_rename4 subst_rename5 subst_rename6
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2068
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2069
lemma better_Cut_substn[simp]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2070
  assumes a: "a\<sharp>[c].P" "x\<sharp>(y,P)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2071
  shows "(Cut <a>.M (x).N){y:=<c>.P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2072
  (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}))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2073
proof -   
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2074
  obtain x'::"name"   where fs1: "x'\<sharp>(M,N,c,P,x,y)" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2075
  obtain a'::"coname" where fs2: "a'\<sharp>(M,N,c,P,a)" by (rule exists_fresh(2), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2076
  have eq1: "(Cut <a>.M (x).N) = (Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2077
    using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2078
  have eq2: "(M=Ax y a) = (([(a',a)]\<bullet>M)=Ax y a')"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2079
    apply(auto simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2080
    apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2081
    apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2082
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2083
  have "(Cut <a>.M (x).N){y:=<c>.P} = (Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N)){y:=<c>.P}" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2084
    using eq1 by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2085
  also have "\<dots> = (if ([(a',a)]\<bullet>M)=Ax y a' then Cut <c>.P (x').(([(x',x)]\<bullet>N){y:=<c>.P}) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2086
                              else Cut <a'>.(([(a',a)]\<bullet>M){y:=<c>.P}) (x').(([(x',x)]\<bullet>N){y:=<c>.P}))" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2087
    using fs1 fs2 by (auto simp add: fresh_prod fresh_left calc_atm fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2088
  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}))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2089
    using fs1 fs2 a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2090
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2091
    apply(simp only: eq2[symmetric])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2092
    apply(auto simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2093
    apply(simp_all add: alpha fresh_atm fresh_prod subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2094
    apply(simp_all add: eqvts perm_fresh_fresh calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2095
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2096
    apply(rule subst_rename)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2097
    apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2098
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2099
    apply(simp add: perm_fresh_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2100
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2101
  finally show ?thesis by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2102
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2103
    
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2104
lemma better_Cut_substc[simp]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2105
  assumes a: "a\<sharp>(c,P)" "x\<sharp>[y].P"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2106
  shows "(Cut <a>.M (x).N){c:=(y).P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2107
  (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}))" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2108
proof -   
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2109
  obtain x'::"name"   where fs1: "x'\<sharp>(M,N,c,P,x,y)" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2110
  obtain a'::"coname" where fs2: "a'\<sharp>(M,N,c,P,a)" by (rule exists_fresh(2), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2111
  have eq1: "(Cut <a>.M (x).N) = (Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2112
    using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2113
  have eq2: "(N=Ax x c) = (([(x',x)]\<bullet>N)=Ax x' c)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2114
    apply(auto simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2115
    apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2116
    apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2117
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2118
  have "(Cut <a>.M (x).N){c:=(y).P} = (Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N)){c:=(y).P}" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2119
    using eq1 by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2120
  also have "\<dots> = (if ([(x',x)]\<bullet>N)=Ax x' c then  Cut <a'>.(([(a',a)]\<bullet>M){c:=(y).P}) (y).P
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2121
                              else Cut <a'>.(([(a',a)]\<bullet>M){c:=(y).P}) (x').(([(x',x)]\<bullet>N){c:=(y).P}))" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2122
    using fs1 fs2  by (simp add: fresh_prod fresh_left calc_atm fresh_atm trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2123
  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}))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2124
    using fs1 fs2 a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2125
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2126
    apply(simp only: eq2[symmetric])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2127
    apply(auto simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2128
    apply(simp_all add: alpha fresh_atm fresh_prod subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2129
    apply(simp_all add: eqvts perm_fresh_fresh calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2130
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2131
    apply(rule subst_rename)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2132
    apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2133
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2134
    apply(simp add: perm_fresh_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2135
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2136
  finally show ?thesis by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2137
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2138
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2139
lemma better_Cut_substn':
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2140
  assumes a: "a\<sharp>[c].P" "y\<sharp>(N,x)" "M\<noteq>Ax y a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2141
  shows "(Cut <a>.M (x).N){y:=<c>.P} = Cut <a>.(M{y:=<c>.P}) (x).N"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2142
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2143
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2144
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2145
apply(subgoal_tac "Cut <a>.M (x).N = Cut <a>.M (ca).([(ca,x)]\<bullet>N)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2146
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2147
apply(subgoal_tac"y\<sharp>([(ca,x)]\<bullet>N)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2148
apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2149
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2150
apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2151
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2152
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2153
apply(simp add: alpha fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2154
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2155
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2156
lemma better_NotR_substc:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2157
  assumes a: "d\<sharp>M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2158
  shows "(NotR (x).M d){d:=(z).P} = fresh_fun (\<lambda>a'. Cut <a'>.NotR (x).M a' (z).P)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2159
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2160
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2161
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2162
apply(subgoal_tac "NotR (x).M d = NotR (c).([(c,x)]\<bullet>M) d")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2163
apply(auto simp add: fresh_left calc_atm forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2164
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2165
apply(rule_tac f="fresh_fun" in arg_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2166
apply(simp add:  expand_fun_eq)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2167
apply(rule allI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2168
apply(simp add: trm.inject alpha fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2169
apply(perm_simp add: trm.inject alpha fresh_prod fresh_atm fresh_left, auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2170
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2171
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2172
lemma better_NotL_substn:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2173
  assumes a: "y\<sharp>M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2174
  shows "(NotL <a>.M y){y:=<c>.P} = fresh_fun (\<lambda>x'. Cut <c>.P (x').NotL <a>.M x')"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2175
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2176
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2177
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2178
apply(subgoal_tac "NotL <a>.M y = NotL <ca>.([(ca,a)]\<bullet>M) y")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2179
apply(auto simp add: fresh_left calc_atm forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2180
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2181
apply(rule_tac f="fresh_fun" in arg_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2182
apply(simp add:  expand_fun_eq)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2183
apply(rule allI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2184
apply(simp add: trm.inject alpha fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2185
apply(perm_simp add: trm.inject alpha fresh_prod fresh_atm fresh_left, auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2186
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2187
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2188
lemma better_AndL1_substn:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2189
  assumes a: "y\<sharp>[x].M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2190
  shows "(AndL1 (x).M y){y:=<c>.P} = fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL1 (x).M z')"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2191
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2192
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2193
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2194
apply(subgoal_tac "AndL1 (x).M y = AndL1 (ca).([(ca,x)]\<bullet>M) y")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2195
apply(auto simp add: fresh_left calc_atm forget abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2196
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2197
apply(rule_tac f="fresh_fun" in arg_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2198
apply(simp add:  expand_fun_eq)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2199
apply(rule allI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2200
apply(simp add: trm.inject alpha fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2201
apply(rule forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2202
apply(simp add: fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2203
apply(rule_tac f="fresh_fun" in arg_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2204
apply(simp add:  expand_fun_eq)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2205
apply(rule allI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2206
apply(simp add: trm.inject alpha fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2207
apply(rule forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2208
apply(simp add: fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2209
apply(perm_simp add: trm.inject alpha fresh_left calc_atm fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2210
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2211
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2212
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2213
lemma better_AndL2_substn:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2214
  assumes a: "y\<sharp>[x].M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2215
  shows "(AndL2 (x).M y){y:=<c>.P} = fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL2 (x).M z')"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2216
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2217
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2218
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2219
apply(subgoal_tac "AndL2 (x).M y = AndL2 (ca).([(ca,x)]\<bullet>M) y")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2220
apply(auto simp add: fresh_left calc_atm forget abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2221
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2222
apply(rule_tac f="fresh_fun" in arg_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2223
apply(simp add:  expand_fun_eq)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2224
apply(rule allI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2225
apply(simp add: trm.inject alpha fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2226
apply(rule forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2227
apply(simp add: fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2228
apply(rule_tac f="fresh_fun" in arg_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2229
apply(simp add:  expand_fun_eq)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2230
apply(rule allI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2231
apply(simp add: trm.inject alpha fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2232
apply(rule forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2233
apply(simp add: fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2234
apply(perm_simp add: trm.inject alpha fresh_left calc_atm fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2235
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2236
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2237
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2238
lemma better_AndR_substc:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2239
  assumes a: "c\<sharp>([a].M,[b].N)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2240
  shows "(AndR <a>.M <b>.N c){c:=(z).P} = fresh_fun (\<lambda>a'. Cut <a'>.(AndR <a>.M <b>.N a') (z).P)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2241
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2242
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2243
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2244
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2245
apply(subgoal_tac "AndR <a>.M <b>.N c = AndR <ca>.([(ca,a)]\<bullet>M) <caa>.([(caa,b)]\<bullet>N) c")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2246
apply(auto simp add: fresh_left calc_atm forget abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2247
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2248
apply(rule substc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2249
apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2250
apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2251
apply(auto simp add: fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2252
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2253
apply(rule_tac f="fresh_fun" in arg_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2254
apply(simp add:  expand_fun_eq)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2255
apply(rule allI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2256
apply(simp add: trm.inject alpha fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2257
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2258
apply(rule forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2259
apply(auto simp add: fresh_left calc_atm abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2260
apply(rule forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2261
apply(auto simp add: fresh_left calc_atm abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2262
apply(perm_simp add: trm.inject alpha fresh_left calc_atm fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2263
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2264
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2265
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2266
lemma better_OrL_substn:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2267
  assumes a: "x\<sharp>([y].M,[z].N)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2268
  shows "(OrL (y).M (z).N x){x:=<c>.P} = fresh_fun (\<lambda>z'. Cut <c>.P (z').OrL (y).M (z).N z')"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2269
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2270
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2271
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2272
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2273
apply(subgoal_tac "OrL (y).M (z).N x = OrL (ca).([(ca,y)]\<bullet>M) (caa).([(caa,z)]\<bullet>N) x")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2274
apply(auto simp add: fresh_left calc_atm forget abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2275
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2276
apply(rule substn.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2277
apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2278
apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2279
apply(auto simp add: fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2280
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2281
apply(rule_tac f="fresh_fun" in arg_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2282
apply(simp add:  expand_fun_eq)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2283
apply(rule allI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2284
apply(simp add: trm.inject alpha fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2285
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2286
apply(rule forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2287
apply(auto simp add: fresh_left calc_atm abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2288
apply(rule forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2289
apply(auto simp add: fresh_left calc_atm abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2290
apply(perm_simp add: trm.inject alpha fresh_left calc_atm fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2291
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2292
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2293
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2294
lemma better_OrR1_substc:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2295
  assumes a: "d\<sharp>[a].M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2296
  shows "(OrR1 <a>.M d){d:=(z).P} = fresh_fun (\<lambda>a'. Cut <a'>.OrR1 <a>.M a' (z).P)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2297
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2298
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2299
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2300
apply(subgoal_tac "OrR1 <a>.M d = OrR1 <c>.([(c,a)]\<bullet>M) d")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2301
apply(auto simp add: fresh_left calc_atm forget abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2302
apply(rule_tac f="fresh_fun" in arg_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2303
apply(simp add:  expand_fun_eq)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2304
apply(rule allI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2305
apply(simp add: trm.inject alpha fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2306
apply(rule forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2307
apply(simp add: fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2308
apply(rule_tac f="fresh_fun" in arg_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2309
apply(simp add:  expand_fun_eq)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2310
apply(rule allI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2311
apply(simp add: trm.inject alpha fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2312
apply(rule forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2313
apply(simp add: fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2314
apply(perm_simp add: trm.inject alpha fresh_left calc_atm fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2315
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2316
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2317
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2318
lemma better_OrR2_substc:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2319
  assumes a: "d\<sharp>[a].M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2320
  shows "(OrR2 <a>.M d){d:=(z).P} = fresh_fun (\<lambda>a'. Cut <a'>.OrR2 <a>.M a' (z).P)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2321
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2322
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2323
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2324
apply(subgoal_tac "OrR2 <a>.M d = OrR2 <c>.([(c,a)]\<bullet>M) d")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2325
apply(auto simp add: fresh_left calc_atm forget abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2326
apply(rule_tac f="fresh_fun" in arg_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2327
apply(simp add:  expand_fun_eq)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2328
apply(rule allI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2329
apply(simp add: trm.inject alpha fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2330
apply(rule forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2331
apply(simp add: fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2332
apply(rule_tac f="fresh_fun" in arg_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2333
apply(simp add:  expand_fun_eq)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2334
apply(rule allI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2335
apply(simp add: trm.inject alpha fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2336
apply(rule forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2337
apply(simp add: fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2338
apply(perm_simp add: trm.inject alpha fresh_left calc_atm fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2339
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2340
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2341
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2342
lemma better_ImpR_substc:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2343
  assumes a: "d\<sharp>[a].M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2344
  shows "(ImpR (x).<a>.M d){d:=(z).P} = fresh_fun (\<lambda>a'. Cut <a'>.ImpR (x).<a>.M a' (z).P)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2345
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2346
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2347
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2348
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2349
apply(subgoal_tac "ImpR (x).<a>.M d = ImpR (ca).<c>.([(c,a)]\<bullet>[(ca,x)]\<bullet>M) d")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2350
apply(auto simp add: fresh_left calc_atm forget abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2351
apply(rule_tac f="fresh_fun" in arg_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2352
apply(simp add:  expand_fun_eq)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2353
apply(rule allI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2354
apply(simp add: trm.inject alpha fresh_prod fresh_atm abs_perm abs_fresh fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2355
apply(rule forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2356
apply(simp add: fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2357
apply(rule_tac f="fresh_fun" in arg_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2358
apply(simp add:  expand_fun_eq)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2359
apply(rule allI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2360
apply(simp add: trm.inject alpha fresh_prod fresh_atm abs_perm fresh_left calc_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2361
apply(rule forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2362
apply(simp add: fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2363
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2364
apply(perm_simp add: trm.inject alpha fresh_left calc_atm fresh_prod fresh_atm abs_fresh abs_perm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2365
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2366
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2367
lemma better_ImpL_substn:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2368
  assumes a: "y\<sharp>(M,[x].N)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2369
  shows "(ImpL <a>.M (x).N y){y:=<c>.P} = fresh_fun (\<lambda>z'. Cut <c>.P (z').ImpL <a>.M (x).N z')"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2370
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2371
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2372
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2373
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2374
apply(subgoal_tac "ImpL <a>.M (x).N y = ImpL <ca>.([(ca,a)]\<bullet>M) (caa).([(caa,x)]\<bullet>N) y")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2375
apply(auto simp add: fresh_left calc_atm forget abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2376
apply(rule_tac f="fresh_fun" in arg_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2377
apply(simp add:  expand_fun_eq)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2378
apply(rule allI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2379
apply(simp add: trm.inject alpha fresh_prod fresh_atm abs_perm abs_fresh fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2380
apply(rule forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2381
apply(simp add: fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2382
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2383
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2384
apply(perm_simp add: trm.inject alpha fresh_left calc_atm fresh_prod fresh_atm abs_fresh abs_perm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2385
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2386
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2387
lemma freshn_after_substc:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2388
  fixes x::"name"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2389
  assumes a: "x\<sharp>M{c:=(y).P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2390
  shows "x\<sharp>M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2391
using a supp_subst8
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2392
apply(simp add: fresh_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2393
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2394
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2395
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2396
lemma freshn_after_substn:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2397
  fixes x::"name"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2398
  assumes a: "x\<sharp>M{y:=<c>.P}" "x\<noteq>y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2399
  shows "x\<sharp>M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2400
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2401
using a supp_subst5
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2402
apply(simp add: fresh_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2403
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2404
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2405
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2406
lemma freshc_after_substc:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2407
  fixes a::"coname"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2408
  assumes a: "a\<sharp>M{c:=(y).P}" "a\<noteq>c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2409
  shows "a\<sharp>M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2410
using a supp_subst7
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2411
apply(simp add: fresh_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2412
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2413
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2414
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2415
lemma freshc_after_substn:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2416
  fixes a::"coname"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2417
  assumes a: "a\<sharp>M{y:=<c>.P}" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2418
  shows "a\<sharp>M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2419
using a supp_subst6
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2420
apply(simp add: fresh_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2421
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2422
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2423
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2424
lemma substn_crename_comm:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2425
  assumes a: "c\<noteq>a" "c\<noteq>b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2426
  shows "M{x:=<c>.P}[a\<turnstile>c>b] = M[a\<turnstile>c>b]{x:=<c>.(P[a\<turnstile>c>b])}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2427
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2428
apply(nominal_induct M avoiding: x c P a b rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2429
apply(auto simp add: subst_fresh rename_fresh trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2430
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(P,x,c)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2431
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2432
apply(subgoal_tac "Cut <c>.P (x).Ax x a = Cut <c>.P (x').Ax x' a")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2433
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2434
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2435
apply(rule crename.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2436
apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2437
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2438
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2439
apply(simp add: alpha trm.inject calc_atm fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2440
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2441
apply(simp add: alpha trm.inject calc_atm fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2442
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2443
apply(rule fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2444
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2445
apply(rule better_crename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2446
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2447
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2448
apply(simp add: crename_id)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2449
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2450
apply(rule better_crename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2451
apply(simp add: fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2452
apply(simp add: rename_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2453
apply(auto simp add: fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2454
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2455
apply(rule better_crename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2456
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2457
apply(auto simp add: fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2458
apply(drule crename_ax)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2459
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2460
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2461
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2462
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{x:=<c>.P},P,P[a\<turnstile>c>b],x,trm[a\<turnstile>c>b]{x:=<c>.P[a\<turnstile>c>b]})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2463
apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2464
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2465
apply(simp add: fresh_fun_simp_NotL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2466
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2467
apply(rule better_crename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2468
apply(simp add: fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2469
apply(simp add: rename_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2470
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2471
apply(rule fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2472
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{x:=<c>.P},P,P[a\<turnstile>c>b],name1,trm[a\<turnstile>c>b]{x:=<c>.P[a\<turnstile>c>b]})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2473
apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2474
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2475
apply(simp add: fresh_fun_simp_AndL1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2476
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2477
apply(rule better_crename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2478
apply(simp add: fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2479
apply(simp add: rename_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2480
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2481
apply(rule fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2482
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{x:=<c>.P},P,P[a\<turnstile>c>b],name1,trm[a\<turnstile>c>b]{x:=<c>.P[a\<turnstile>c>b]})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2483
apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2484
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2485
apply(simp add: fresh_fun_simp_AndL2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2486
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2487
apply(rule better_crename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2488
apply(simp add: fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2489
apply(simp add: rename_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2490
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2491
apply(rule fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2492
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{x:=<c>.P},trm2{x:=<c>.P},P,P[a\<turnstile>c>b],name1,name2,
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2493
                                  trm1[a\<turnstile>c>b]{x:=<c>.P[a\<turnstile>c>b]},trm2[a\<turnstile>c>b]{x:=<c>.P[a\<turnstile>c>b]})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2494
apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2495
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2496
apply(simp add: fresh_fun_simp_OrL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2497
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2498
apply(rule better_crename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2499
apply(simp add: fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2500
apply(simp add: rename_fresh subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2501
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2502
apply(rule fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2503
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{name2:=<c>.P},trm2{name2:=<c>.P},P,P[a\<turnstile>c>b],name1,
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2504
                                  trm1[a\<turnstile>c>b]{name2:=<c>.P[a\<turnstile>c>b]},trm2[a\<turnstile>c>b]{name2:=<c>.P[a\<turnstile>c>b]})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2505
apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2506
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2507
apply(simp add: fresh_fun_simp_ImpL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2508
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2509
apply(rule better_crename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2510
apply(simp add: fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2511
apply(simp add: rename_fresh subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2512
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2513
apply(rule fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2514
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2515
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2516
lemma substc_crename_comm:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2517
  assumes a: "c\<noteq>a" "c\<noteq>b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2518
  shows "M{c:=(x).P}[a\<turnstile>c>b] = M[a\<turnstile>c>b]{c:=(x).(P[a\<turnstile>c>b])}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2519
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2520
apply(nominal_induct M avoiding: x c P a b rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2521
apply(auto simp add: subst_fresh rename_fresh trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2522
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2523
apply(rule better_crename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2524
apply(simp add: fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2525
apply(simp add: rename_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2526
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2527
apply(rule better_crename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2528
apply(simp add: fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2529
apply(simp add: rename_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2530
apply(drule crename_ax)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2531
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2532
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2533
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2534
apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(a,b,trm{coname:=(x).P},P,P[a\<turnstile>c>b],x,trm[a\<turnstile>c>b]{coname:=(x).P[a\<turnstile>c>b]})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2535
apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2536
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2537
apply(simp add: fresh_fun_simp_NotR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2538
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2539
apply(rule better_crename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2540
apply(simp add: fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2541
apply(simp add: rename_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2542
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2543
apply(rule fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2544
apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(coname1,coname2,a,b,trm1{coname3:=(x).P},trm2{coname3:=(x).P},
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2545
                   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]})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2546
apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2547
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2548
apply(simp add: fresh_fun_simp_AndR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2549
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2550
apply(rule better_crename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2551
apply(simp add: fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2552
apply(simp add: rename_fresh subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2553
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2554
apply(rule fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2555
apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(coname1,trm{coname2:=(x).P},P,P[a\<turnstile>c>b],a,b,
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2556
                         trm[a\<turnstile>c>b]{coname2:=(x).P[a\<turnstile>c>b]})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2557
apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2558
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2559
apply(simp add: fresh_fun_simp_OrR1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2560
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2561
apply(rule better_crename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2562
apply(simp add: fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2563
apply(simp add: rename_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2564
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2565
apply(rule fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2566
apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(coname1,trm{coname2:=(x).P},P,P[a\<turnstile>c>b],a,b,
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2567
                         trm[a\<turnstile>c>b]{coname2:=(x).P[a\<turnstile>c>b]})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2568
apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2569
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2570
apply(simp add: fresh_fun_simp_OrR2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2571
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2572
apply(rule better_crename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2573
apply(simp add: fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2574
apply(simp add: rename_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2575
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2576
apply(rule fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2577
apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(coname1,trm{coname2:=(x).P},P,P[a\<turnstile>c>b],a,b,
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2578
                         trm[a\<turnstile>c>b]{coname2:=(x).P[a\<turnstile>c>b]})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2579
apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2580
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2581
apply(simp add: fresh_fun_simp_ImpR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2582
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2583
apply(rule better_crename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2584
apply(simp add: fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2585
apply(simp add: rename_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2586
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2587
apply(rule fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2588
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2589
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2590
lemma substn_nrename_comm:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2591
  assumes a: "x\<noteq>y" "x\<noteq>z"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2592
  shows "M{x:=<c>.P}[y\<turnstile>n>z] = M[y\<turnstile>n>z]{x:=<c>.(P[y\<turnstile>n>z])}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2593
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2594
apply(nominal_induct M avoiding: x c P y z rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2595
apply(auto simp add: subst_fresh rename_fresh trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2596
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2597
apply(rule better_nrename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2598
apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2599
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2600
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2601
apply(rule better_nrename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2602
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2603
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2604
apply(drule nrename_ax)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2605
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2606
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2607
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2608
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(y,z,trm{x:=<c>.P},P,P[y\<turnstile>n>z],x,trm[y\<turnstile>n>z]{x:=<c>.P[y\<turnstile>n>z]})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2609
apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2610
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2611
apply(simp add: fresh_fun_simp_NotL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2612
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2613
apply(rule better_nrename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2614
apply(simp add: fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2615
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2616
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2617
apply(rule fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2618
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{x:=<c>.P},P,P[y\<turnstile>n>z],name1,trm[y\<turnstile>n>z]{x:=<c>.P[y\<turnstile>n>z]},y,z)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2619
apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2620
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2621
apply(simp add: fresh_fun_simp_AndL1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2622
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2623
apply(rule better_nrename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2624
apply(simp add: fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2625
apply(simp add: rename_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2626
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2627
apply(rule fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2628
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(y,z,trm{x:=<c>.P},P,P[y\<turnstile>n>z],name1,trm[y\<turnstile>n>z]{x:=<c>.P[y\<turnstile>n>z]})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2629
apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2630
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2631
apply(simp add: fresh_fun_simp_AndL2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2632
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2633
apply(rule better_nrename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2634
apply(simp add: fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2635
apply(simp add: rename_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2636
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2637
apply(rule fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2638
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{x:=<c>.P},trm2{x:=<c>.P},P,P[y\<turnstile>n>z],name1,name2,y,z,
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2639
                                  trm1[y\<turnstile>n>z]{x:=<c>.P[y\<turnstile>n>z]},trm2[y\<turnstile>n>z]{x:=<c>.P[y\<turnstile>n>z]})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2640
apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2641
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2642
apply(simp add: fresh_fun_simp_OrL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2643
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2644
apply(rule better_nrename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2645
apply(simp add: fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2646
apply(simp add: rename_fresh subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2647
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2648
apply(rule fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2649
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{name2:=<c>.P},trm2{name2:=<c>.P},P,P[y\<turnstile>n>z],y,z,name1,
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2650
                                  trm1[y\<turnstile>n>z]{name2:=<c>.P[y\<turnstile>n>z]},trm2[y\<turnstile>n>z]{name2:=<c>.P[y\<turnstile>n>z]})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2651
apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2652
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2653
apply(simp add: fresh_fun_simp_ImpL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2654
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2655
apply(rule better_nrename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2656
apply(simp add: fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2657
apply(simp add: rename_fresh subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2658
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2659
apply(rule fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2660
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2661
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2662
lemma substc_nrename_comm:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2663
  assumes a: "x\<noteq>y" "x\<noteq>z"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2664
  shows "M{c:=(x).P}[y\<turnstile>n>z] = M[y\<turnstile>n>z]{c:=(x).(P[y\<turnstile>n>z])}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2665
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2666
apply(nominal_induct M avoiding: x c P y z rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2667
apply(auto simp add: subst_fresh rename_fresh trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2668
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2669
apply(rule better_nrename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2670
apply(simp add: fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2671
apply(simp add: rename_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2672
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2673
apply(rule better_nrename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2674
apply(simp add: fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2675
apply(simp add: rename_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2676
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2677
apply(rule better_nrename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2678
apply(simp add: fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2679
apply(simp add: rename_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2680
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2681
apply(rule better_nrename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2682
apply(simp add: fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2683
apply(simp add: rename_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2684
apply(drule nrename_ax)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2685
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2686
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2687
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2688
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2689
apply(rule better_nrename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2690
apply(simp add: fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2691
apply(simp add: rename_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2692
apply(drule nrename_ax)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2693
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2694
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2695
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2696
apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(y,z,trm{coname:=(x).P},P,P[y\<turnstile>n>z],x,trm[y\<turnstile>n>z]{coname:=(x).P[y\<turnstile>n>z]})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2697
apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2698
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2699
apply(simp add: fresh_fun_simp_NotR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2700
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2701
apply(rule better_nrename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2702
apply(simp add: fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2703
apply(simp add: rename_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2704
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2705
apply(rule fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2706
apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(coname1,coname2,y,z,trm1{coname3:=(x).P},trm2{coname3:=(x).P},
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2707
                   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]})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2708
apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2709
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2710
apply(simp add: fresh_fun_simp_AndR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2711
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2712
apply(rule better_nrename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2713
apply(simp add: fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2714
apply(simp add: rename_fresh subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2715
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2716
apply(rule fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2717
apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(coname1,trm{coname2:=(x).P},P,P[y\<turnstile>n>z],y,z,
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2718
                         trm[y\<turnstile>n>z]{coname2:=(x).P[y\<turnstile>n>z]})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2719
apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2720
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2721
apply(simp add: fresh_fun_simp_OrR1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2722
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2723
apply(rule better_nrename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2724
apply(simp add: fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2725
apply(simp add: rename_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2726
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2727
apply(rule fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2728
apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(coname1,trm{coname2:=(x).P},P,P[y\<turnstile>n>z],y,z,
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2729
                         trm[y\<turnstile>n>z]{coname2:=(x).P[y\<turnstile>n>z]})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2730
apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2731
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2732
apply(simp add: fresh_fun_simp_OrR2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2733
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2734
apply(rule better_nrename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2735
apply(simp add: fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2736
apply(simp add: rename_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2737
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2738
apply(rule fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2739
apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(coname1,trm{coname2:=(x).P},P,P[y\<turnstile>n>z],y,z,
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2740
                         trm[y\<turnstile>n>z]{coname2:=(x).P[y\<turnstile>n>z]})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2741
apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2742
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2743
apply(simp add: fresh_fun_simp_ImpR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2744
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2745
apply(rule better_nrename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2746
apply(simp add: fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2747
apply(simp add: rename_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2748
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2749
apply(rule fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2750
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2751
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2752
lemma substn_crename_comm':
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2753
  assumes a: "a\<noteq>c" "a\<sharp>P"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2754
  shows "M{x:=<c>.P}[a\<turnstile>c>b] = M[a\<turnstile>c>b]{x:=<c>.P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2755
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2756
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2757
  assume a1: "a\<noteq>c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2758
  assume a2: "a\<sharp>P"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2759
  obtain c'::"coname" where fs2: "c'\<sharp>(c,P,a,b)" by (rule exists_fresh(2), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2760
  have eq: "M{x:=<c>.P} = M{x:=<c'>.([(c',c)]\<bullet>P)}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2761
    using fs2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2762
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2763
    apply(rule subst_rename)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2764
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2765
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2766
   have eq': "M[a\<turnstile>c>b]{x:=<c>.P} = M[a\<turnstile>c>b]{x:=<c'>.([(c',c)]\<bullet>P)}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2767
    using fs2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2768
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2769
    apply(rule subst_rename)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2770
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2771
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2772
  have eq2: "([(c',c)]\<bullet>P)[a\<turnstile>c>b] = ([(c',c)]\<bullet>P)" using fs2 a2 a1
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2773
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2774
    apply(rule rename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2775
    apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2776
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2777
  have "M{x:=<c>.P}[a\<turnstile>c>b] = M{x:=<c'>.([(c',c)]\<bullet>P)}[a\<turnstile>c>b]" using eq by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2778
  also have "\<dots> = M[a\<turnstile>c>b]{x:=<c'>.(([(c',c)]\<bullet>P)[a\<turnstile>c>b])}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2779
    using fs2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2780
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2781
    apply(rule substn_crename_comm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2782
    apply(simp_all add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2783
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2784
  also have "\<dots> = M[a\<turnstile>c>b]{x:=<c'>.(([(c',c)]\<bullet>P))}" using eq2 by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2785
  also have "\<dots> = M[a\<turnstile>c>b]{x:=<c>.P}" using eq' by simp 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2786
  finally show ?thesis by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2787
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2788
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2789
lemma substc_crename_comm':
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2790
  assumes a: "c\<noteq>a" "c\<noteq>b" "a\<sharp>P"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2791
  shows "M{c:=(x).P}[a\<turnstile>c>b] = M[a\<turnstile>c>b]{c:=(x).P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2792
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2793
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2794
  assume a1: "c\<noteq>a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2795
  assume a1': "c\<noteq>b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2796
  assume a2: "a\<sharp>P"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2797
  obtain c'::"coname" where fs2: "c'\<sharp>(c,M,a,b)" by (rule exists_fresh(2), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2798
  have eq: "M{c:=(x).P} = ([(c',c)]\<bullet>M){c':=(x).P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2799
    using fs2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2800
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2801
    apply(rule subst_rename)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2802
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2803
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2804
   have eq': "([(c',c)]\<bullet>(M[a\<turnstile>c>b])){c':=(x).P} = M[a\<turnstile>c>b]{c:=(x).P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2805
    using fs2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2806
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2807
    apply(rule subst_rename[symmetric])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2808
    apply(simp add: rename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2809
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2810
  have eq2: "([(c',c)]\<bullet>M)[a\<turnstile>c>b] = ([(c',c)]\<bullet>(M[a\<turnstile>c>b]))" using fs2 a2 a1 a1'
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2811
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2812
    apply(simp add: rename_eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2813
    apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2814
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2815
  have "M{c:=(x).P}[a\<turnstile>c>b] = ([(c',c)]\<bullet>M){c':=(x).P}[a\<turnstile>c>b]" using eq by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2816
  also have "\<dots> = ([(c',c)]\<bullet>M)[a\<turnstile>c>b]{c':=(x).P[a\<turnstile>c>b]}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2817
    using fs2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2818
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2819
    apply(rule substc_crename_comm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2820
    apply(simp_all add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2821
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2822
  also have "\<dots> = ([(c',c)]\<bullet>(M[a\<turnstile>c>b])){c':=(x).P[a\<turnstile>c>b]}" using eq2 by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2823
  also have "\<dots> = ([(c',c)]\<bullet>(M[a\<turnstile>c>b])){c':=(x).P}" using a2 by (simp add: rename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2824
  also have "\<dots> = M[a\<turnstile>c>b]{c:=(x).P}" using eq' by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2825
  finally show ?thesis by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2826
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2827
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2828
lemma substn_nrename_comm':
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2829
  assumes a: "x\<noteq>y" "x\<noteq>z" "y\<sharp>P"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2830
  shows "M{x:=<c>.P}[y\<turnstile>n>z] = M[y\<turnstile>n>z]{x:=<c>.P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2831
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2832
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2833
  assume a1: "x\<noteq>y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2834
  assume a1': "x\<noteq>z"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2835
  assume a2: "y\<sharp>P"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2836
  obtain x'::"name" where fs2: "x'\<sharp>(x,M,y,z)" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2837
  have eq: "M{x:=<c>.P} = ([(x',x)]\<bullet>M){x':=<c>.P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2838
    using fs2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2839
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2840
    apply(rule subst_rename)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2841
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2842
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2843
   have eq': "([(x',x)]\<bullet>(M[y\<turnstile>n>z])){x':=<c>.P} = M[y\<turnstile>n>z]{x:=<c>.P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2844
    using fs2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2845
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2846
    apply(rule subst_rename[symmetric])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2847
    apply(simp add: rename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2848
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2849
  have eq2: "([(x',x)]\<bullet>M)[y\<turnstile>n>z] = ([(x',x)]\<bullet>(M[y\<turnstile>n>z]))" using fs2 a2 a1 a1'
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2850
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2851
    apply(simp add: rename_eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2852
    apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2853
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2854
  have "M{x:=<c>.P}[y\<turnstile>n>z] = ([(x',x)]\<bullet>M){x':=<c>.P}[y\<turnstile>n>z]" using eq by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2855
  also have "\<dots> = ([(x',x)]\<bullet>M)[y\<turnstile>n>z]{x':=<c>.P[y\<turnstile>n>z]}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2856
    using fs2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2857
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2858
    apply(rule substn_nrename_comm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2859
    apply(simp_all add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2860
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2861
  also have "\<dots> = ([(x',x)]\<bullet>(M[y\<turnstile>n>z])){x':=<c>.P[y\<turnstile>n>z]}" using eq2 by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2862
  also have "\<dots> = ([(x',x)]\<bullet>(M[y\<turnstile>n>z])){x':=<c>.P}" using a2 by (simp add: rename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2863
  also have "\<dots> = M[y\<turnstile>n>z]{x:=<c>.P}" using eq' by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2864
  finally show ?thesis by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2865
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2866
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2867
lemma substc_nrename_comm':
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2868
  assumes a: "x\<noteq>y" "y\<sharp>P"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2869
  shows "M{c:=(x).P}[y\<turnstile>n>z] = M[y\<turnstile>n>z]{c:=(x).P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2870
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2871
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2872
  assume a1: "x\<noteq>y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2873
  assume a2: "y\<sharp>P"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2874
  obtain x'::"name" where fs2: "x'\<sharp>(x,P,y,z)" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2875
  have eq: "M{c:=(x).P} = M{c:=(x').([(x',x)]\<bullet>P)}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2876
    using fs2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2877
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2878
    apply(rule subst_rename)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2879
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2880
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2881
   have eq': "M[y\<turnstile>n>z]{c:=(x).P} = M[y\<turnstile>n>z]{c:=(x').([(x',x)]\<bullet>P)}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2882
    using fs2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2883
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2884
    apply(rule subst_rename)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2885
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2886
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2887
  have eq2: "([(x',x)]\<bullet>P)[y\<turnstile>n>z] = ([(x',x)]\<bullet>P)" using fs2 a2 a1
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2888
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2889
    apply(rule rename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2890
    apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2891
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2892
  have "M{c:=(x).P}[y\<turnstile>n>z] = M{c:=(x').([(x',x)]\<bullet>P)}[y\<turnstile>n>z]" using eq by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2893
  also have "\<dots> = M[y\<turnstile>n>z]{c:=(x').(([(x',x)]\<bullet>P)[y\<turnstile>n>z])}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2894
    using fs2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2895
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2896
    apply(rule substc_nrename_comm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2897
    apply(simp_all add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2898
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2899
  also have "\<dots> = M[y\<turnstile>n>z]{c:=(x').(([(x',x)]\<bullet>P))}" using eq2 by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2900
  also have "\<dots> = M[y\<turnstile>n>z]{c:=(x).P}" using eq' by simp 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2901
  finally show ?thesis by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2902
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2903
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2904
lemmas subst_comm = substn_crename_comm substc_crename_comm
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2905
                    substn_nrename_comm substc_nrename_comm
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2906
lemmas subst_comm' = substn_crename_comm' substc_crename_comm'
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2907
                     substn_nrename_comm' substc_nrename_comm'
18881
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
  2908
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
  2909
text {* typing contexts *}
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
  2910
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
  2911
types 
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
  2912
  ctxtn = "(name\<times>ty) list"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
  2913
  ctxtc = "(coname\<times>ty) list"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
  2914
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
  2915
inductive2
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
  2916
  validc :: "ctxtc \<Rightarrow> bool"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
  2917
where
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
  2918
  vc1[intro]: "validc []"
23092
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2919
| vc2[intro]: "\<lbrakk>a\<sharp>\<Delta>; validc \<Delta>\<rbrakk> \<Longrightarrow> validc ((a,T)#\<Delta>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2920
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2921
equivariance validc
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
  2922
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
  2923
inductive2
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
  2924
  validn :: "ctxtn \<Rightarrow> bool"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
  2925
where
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
  2926
  vn1[intro]: "validn []"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
  2927
| vn2[intro]: "\<lbrakk>x\<sharp>\<Gamma>; validn \<Gamma>\<rbrakk> \<Longrightarrow> validn ((x,T)#\<Gamma>)"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
  2928
23092
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2929
equivariance validn
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2930
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2931
lemma fresh_ctxt:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2932
  fixes a::"coname"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2933
  and   x::"name"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2934
  and   \<Gamma>::"ctxtn"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2935
  and   \<Delta>::"ctxtc"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2936
  shows "a\<sharp>\<Gamma>" and "x\<sharp>\<Delta>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2937
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2938
  show "a\<sharp>\<Gamma>" by (induct \<Gamma>) (auto simp add: fresh_list_nil fresh_list_cons fresh_prod fresh_atm fresh_ty)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2939
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2940
  show "x\<sharp>\<Delta>" by (induct \<Delta>) (auto simp add: fresh_list_nil fresh_list_cons fresh_prod fresh_atm fresh_ty)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2941
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2942
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2943
text {* cut-reductions *}
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2944
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2945
declare abs_perm[eqvt]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2946
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2947
inductive2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2948
  fin :: "trm \<Rightarrow> name \<Rightarrow> bool"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2949
where
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2950
  [intro]: "fin (Ax x a) x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2951
| [intro]: "x\<sharp>M \<Longrightarrow> fin (NotL <a>.M x) x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2952
| [intro]: "y\<sharp>[x].M \<Longrightarrow> fin (AndL1 (x).M y) y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2953
| [intro]: "y\<sharp>[x].M \<Longrightarrow> fin (AndL2 (x).M y) y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2954
| [intro]: "\<lbrakk>z\<sharp>[x].M;z\<sharp>[y].N\<rbrakk> \<Longrightarrow> fin (OrL (x).M (y).N z) z"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2955
| [intro]: "\<lbrakk>y\<sharp>M;y\<sharp>[x].N\<rbrakk> \<Longrightarrow> fin (ImpL <a>.M (x).N y) y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2956
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2957
equivariance fin
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2958
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2959
lemma fin_Ax_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2960
  assumes a: "fin (Ax x a) y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2961
  shows "x=y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2962
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2963
apply(erule_tac fin.cases)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2964
apply(auto simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2965
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2966
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2967
lemma fin_NotL_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2968
  assumes a: "fin (NotL <a>.M x) y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2969
  shows "x=y \<and> x\<sharp>M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2970
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2971
apply(erule_tac fin.cases)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2972
apply(auto simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2973
apply(subgoal_tac "y\<sharp>[aa].Ma")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2974
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2975
apply(simp_all add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2976
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2977
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2978
lemma fin_AndL1_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2979
  assumes a: "fin (AndL1 (x).M y) z"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2980
  shows "z=y \<and> z\<sharp>[x].M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2981
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2982
apply(erule_tac fin.cases)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2983
apply(auto simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2984
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2985
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2986
lemma fin_AndL2_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2987
  assumes a: "fin (AndL2 (x).M y) z"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2988
  shows "z=y \<and> z\<sharp>[x].M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2989
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2990
apply(erule_tac fin.cases)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2991
apply(auto simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2992
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2993
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2994
lemma fin_OrL_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2995
  assumes a: "fin (OrL (x).M (y).N u) z"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2996
  shows "z=u \<and> z\<sharp>[x].M \<and> z\<sharp>[y].N"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2997
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2998
apply(erule_tac fin.cases)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  2999
apply(auto simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3000
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3001
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3002
lemma fin_ImpL_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3003
  assumes a: "fin (ImpL <a>.M (x).N z) y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3004
  shows "z=y \<and> z\<sharp>M \<and> z\<sharp>[x].N"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3005
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3006
apply(erule_tac fin.cases)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3007
apply(auto simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3008
apply(subgoal_tac "y\<sharp>[aa].Ma")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3009
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3010
apply(simp_all add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3011
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3012
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3013
lemma fin_rest_elims:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3014
  shows "fin (Cut <a>.M (x).N) y \<Longrightarrow> False"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3015
  and   "fin (NotR (x).M c) y \<Longrightarrow> False"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3016
  and   "fin (AndR <a>.M <b>.N c) y \<Longrightarrow> False"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3017
  and   "fin (OrR1 <a>.M b) y \<Longrightarrow> False"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3018
  and   "fin (OrR2 <a>.M b) y \<Longrightarrow> False"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3019
  and   "fin (ImpR (x).<a>.M b) y \<Longrightarrow> False"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3020
by (erule fin.cases, simp_all add: trm.inject)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3021
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3022
lemmas fin_elims = fin_Ax_elim fin_NotL_elim fin_AndL1_elim fin_AndL2_elim fin_OrL_elim 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3023
                   fin_ImpL_elim fin_rest_elims
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3024
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3025
lemma fin_rename:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3026
  shows "fin M x \<Longrightarrow> fin ([(x',x)]\<bullet>M) x'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3027
by (induct rule: fin.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3028
   (auto simp add: calc_atm simp add: fresh_left abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3029
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3030
lemma not_fin_subst1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3031
  assumes a: "\<not>(fin M x)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3032
  shows "\<not>(fin (M{c:=(y).P}) x)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3033
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3034
apply(nominal_induct M avoiding: x c y P rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3035
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3036
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3037
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3038
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3039
apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(trm{coname:=(y).P},P,x)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3040
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3041
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3042
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3043
apply(simp add: fresh_fun_simp_NotR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3044
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3045
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3046
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3047
apply(drule fin_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3048
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3049
apply(drule freshn_after_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3050
apply(simp add: fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3051
apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(trm1{coname3:=(y).P},trm2{coname3:=(y).P},P,coname1,coname2,coname3,x)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3052
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3053
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3054
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3055
apply(simp add: fresh_fun_simp_AndR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3056
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3057
apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(trm1{coname3:=(y).P},trm2{coname3:=(y).P},P,coname1,coname2,coname3,x)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3058
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3059
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3060
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3061
apply(simp add: fresh_fun_simp_AndR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3062
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3063
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3064
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3065
apply(drule fin_AndL1_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3066
apply(auto simp add: abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3067
apply(drule freshn_after_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3068
apply(subgoal_tac "name2\<sharp>[name1]. trm")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3069
apply(simp add: fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3070
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3071
apply(drule fin_AndL2_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3072
apply(auto simp add: abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3073
apply(drule freshn_after_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3074
apply(subgoal_tac "name2\<sharp>[name1].trm")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3075
apply(simp add: fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3076
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3077
apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(trm{coname2:=(y).P},coname1,P,x)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3078
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3079
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3080
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3081
apply(simp add: fresh_fun_simp_OrR1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3082
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3083
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3084
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3085
apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(trm{coname2:=(y).P},coname1,P,x)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3086
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3087
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3088
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3089
apply(simp add: fresh_fun_simp_OrR2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3090
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3091
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3092
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3093
apply(drule fin_OrL_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3094
apply(auto simp add: abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3095
apply(drule freshn_after_substc)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3096
apply(subgoal_tac "name3\<sharp>[name1].trm1 \<and> name3\<sharp>[name2].trm2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3097
apply(simp add: fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3098
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3099
apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(trm{coname2:=(y).P},coname1,P,x)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3100
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3101
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3102
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3103
apply(simp add: fresh_fun_simp_ImpR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3104
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3105
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3106
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3107
apply(drule fin_ImpL_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3108
apply(auto simp add: abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3109
apply(drule freshn_after_substc)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3110
apply(subgoal_tac "x\<sharp>[name1].trm2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3111
apply(simp add: fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3112
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3113
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3114
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3115
lemma not_fin_subst2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3116
  assumes a: "\<not>(fin M x)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3117
  shows "\<not>(fin (M{y:=<c>.P}) x)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3118
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3119
apply(nominal_induct M avoiding: x c y P rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3120
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3121
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3122
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3123
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3124
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3125
apply(subgoal_tac "\<exists>a'::name. a'\<sharp>(trm{y:=<c>.P},P,x)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3126
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3127
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3128
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3129
apply(simp add: fresh_fun_simp_NotL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3130
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3131
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3132
apply(drule fin_NotL_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3133
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3134
apply(drule freshn_after_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3135
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3136
apply(simp add: fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3137
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3138
apply(subgoal_tac "\<exists>a'::name. a'\<sharp>(trm{y:=<c>.P},P,name1,x)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3139
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3140
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3141
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3142
apply(simp add: fresh_fun_simp_AndL1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3143
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3144
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3145
apply(drule fin_AndL1_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3146
apply(auto simp add: abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3147
apply(drule freshn_after_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3148
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3149
apply(subgoal_tac "name2\<sharp>[name1]. trm")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3150
apply(simp add: fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3151
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3152
apply(subgoal_tac "\<exists>a'::name. a'\<sharp>(trm{y:=<c>.P},P,name1,x)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3153
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3154
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3155
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3156
apply(simp add: fresh_fun_simp_AndL2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3157
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3158
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3159
apply(drule fin_AndL2_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3160
apply(auto simp add: abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3161
apply(drule freshn_after_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3162
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3163
apply(subgoal_tac "name2\<sharp>[name1].trm")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3164
apply(simp add: fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3165
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3166
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3167
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3168
apply(subgoal_tac "\<exists>a'::name. a'\<sharp>(trm1{y:=<c>.P},trm2{y:=<c>.P},name1,name2,P,x)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3169
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3170
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3171
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3172
apply(simp add: fresh_fun_simp_OrL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3173
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3174
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3175
apply(drule fin_OrL_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3176
apply(auto simp add: abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3177
apply(drule freshn_after_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3178
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3179
apply(drule freshn_after_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3180
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3181
apply(subgoal_tac "name3\<sharp>[name1].trm1 \<and> name3\<sharp>[name2].trm2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3182
apply(simp add: fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3183
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3184
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3185
apply(subgoal_tac "\<exists>a'::name. a'\<sharp>(trm1{name2:=<c>.P},trm2{name2:=<c>.P},name1,P,x)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3186
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3187
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3188
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3189
apply(simp add: fresh_fun_simp_ImpL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3190
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3191
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3192
apply(drule fin_ImpL_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3193
apply(auto simp add: abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3194
apply(drule freshn_after_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3195
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3196
apply(drule freshn_after_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3197
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3198
apply(subgoal_tac "x\<sharp>[name1].trm2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3199
apply(simp add: fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3200
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3201
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3202
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3203
lemma fin_subst1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3204
  assumes a: "fin M x" "x\<noteq>y" "x\<sharp>P"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3205
  shows "fin (M{y:=<c>.P}) x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3206
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3207
apply(nominal_induct M avoiding: x y c P rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3208
apply(auto dest!: fin_elims simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3209
apply(rule fin.intros, simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3210
apply(rule fin.intros, simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3211
apply(rule fin.intros, simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3212
apply(rule fin.intros, simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3213
apply(rule fin.intros, simp add: subst_fresh abs_fresh, simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3214
apply(rule fin.intros, simp add: subst_fresh abs_fresh, simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3215
apply(rule fin.intros, simp add: subst_fresh abs_fresh, simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3216
apply(rule fin.intros, simp add: subst_fresh abs_fresh, simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3217
apply(rule fin.intros, simp add: subst_fresh abs_fresh, simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3218
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3219
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3220
lemma fin_subst2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3221
  assumes a: "fin M y" "x\<noteq>y" "y\<sharp>P" "M\<noteq>Ax y c" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3222
  shows "fin (M{c:=(x).P}) y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3223
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3224
apply(nominal_induct M avoiding: x y c P rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3225
apply(drule fin_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3226
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3227
apply(rule fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3228
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3229
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3230
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3231
apply(rule fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3232
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3233
apply(rule subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3234
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3235
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3236
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3237
apply(rule fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3238
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3239
apply(rule subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3240
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3241
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3242
apply(rule fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3243
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3244
apply(rule subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3245
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3246
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3247
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3248
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3249
apply(rule fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3250
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3251
apply(rule subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3252
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3253
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3254
apply(rule subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3255
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3256
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3257
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3258
apply(rule fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3259
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3260
apply(rule subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3261
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3262
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3263
apply(rule subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3264
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3265
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3266
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3267
lemma fin_substn_nrename:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3268
  assumes a: "fin M x" "x\<noteq>y" "x\<sharp>P"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3269
  shows "M[x\<turnstile>n>y]{y:=<c>.P} = Cut <c>.P (x).(M{y:=<c>.P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3270
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3271
apply(nominal_induct M avoiding: x y c P rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3272
apply(drule fin_Ax_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3273
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3274
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3275
apply(simp add: alpha calc_atm fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3276
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3277
apply(drule fin_rest_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3278
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3279
apply(drule fin_rest_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3280
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3281
apply(drule fin_NotL_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3282
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3283
apply(subgoal_tac "\<exists>z::name. z\<sharp>(trm,y,x,P,trm[x\<turnstile>n>y]{y:=<c>.P})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3284
apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3285
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3286
apply(simp add: fresh_fun_simp_NotL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3287
apply(simp add: trm.inject alpha fresh_atm calc_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3288
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3289
apply(simp add: nsubst_eqvt calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3290
apply(simp add: perm_fresh_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3291
apply(simp add: nrename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3292
apply(rule subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3293
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3294
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3295
apply(rule fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3296
apply(drule fin_rest_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3297
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3298
apply(drule fin_AndL1_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3299
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3300
apply(subgoal_tac "\<exists>z::name. z\<sharp>(name2,name1,P,trm[name2\<turnstile>n>y]{y:=<c>.P},y,P,trm)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3301
apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3302
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3303
apply(simp add: fresh_fun_simp_AndL1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3304
apply(simp add: trm.inject alpha fresh_atm calc_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3305
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3306
apply(simp add: nsubst_eqvt calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3307
apply(simp add: perm_fresh_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3308
apply(simp add: nrename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3309
apply(rule subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3310
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3311
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3312
apply(rule fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3313
apply(drule fin_AndL2_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3314
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3315
apply(subgoal_tac "\<exists>z::name. z\<sharp>(name2,name1,P,trm[name2\<turnstile>n>y]{y:=<c>.P},y,P,trm)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3316
apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3317
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3318
apply(simp add: fresh_fun_simp_AndL2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3319
apply(simp add: trm.inject alpha fresh_atm calc_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3320
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3321
apply(simp add: nsubst_eqvt calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3322
apply(simp add: perm_fresh_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3323
apply(simp add: nrename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3324
apply(rule subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3325
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3326
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3327
apply(rule fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3328
apply(drule fin_rest_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3329
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3330
apply(drule fin_rest_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3331
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3332
apply(drule fin_OrL_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3333
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3334
apply(simp add: subst_fresh rename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3335
apply(subgoal_tac "\<exists>z::name. z\<sharp>(name3,name2,name1,P,trm1[name3\<turnstile>n>y]{y:=<c>.P},trm2[name3\<turnstile>n>y]{y:=<c>.P},y,P,trm1,trm2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3336
apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3337
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3338
apply(simp add: fresh_fun_simp_OrL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3339
apply(simp add: trm.inject alpha fresh_atm calc_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3340
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3341
apply(simp add: nsubst_eqvt calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3342
apply(simp add: perm_fresh_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3343
apply(simp add: nrename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3344
apply(simp add: nsubst_eqvt calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3345
apply(simp add: perm_fresh_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3346
apply(simp add: nrename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3347
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3348
apply(rule fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3349
apply(drule fin_rest_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3350
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3351
apply(drule fin_ImpL_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3352
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3353
apply(simp add: subst_fresh rename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3354
apply(subgoal_tac "\<exists>z::name. z\<sharp>(name1,x,P,trm1[x\<turnstile>n>y]{y:=<c>.P},trm2[x\<turnstile>n>y]{y:=<c>.P},y,P,trm1,trm2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3355
apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3356
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3357
apply(simp add: fresh_fun_simp_ImpL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3358
apply(simp add: trm.inject alpha fresh_atm calc_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3359
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3360
apply(simp add: nsubst_eqvt calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3361
apply(simp add: perm_fresh_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3362
apply(simp add: nrename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3363
apply(simp add: nsubst_eqvt calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3364
apply(simp add: perm_fresh_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3365
apply(simp add: nrename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3366
apply(rule exists_fresh')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3367
apply(rule fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3368
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3369
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3370
inductive2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3371
  fic :: "trm \<Rightarrow> coname \<Rightarrow> bool"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3372
where
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3373
  [intro]: "fic (Ax x a) a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3374
| [intro]: "a\<sharp>M \<Longrightarrow> fic (NotR (x).M a) a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3375
| [intro]: "\<lbrakk>c\<sharp>[a].M;c\<sharp>[b].N\<rbrakk> \<Longrightarrow> fic (AndR <a>.M <b>.N c) c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3376
| [intro]: "b\<sharp>[a].M \<Longrightarrow> fic (OrR1 <a>.M b) b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3377
| [intro]: "b\<sharp>[a].M \<Longrightarrow> fic (OrR2 <a>.M b) b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3378
| [intro]: "\<lbrakk>b\<sharp>[a].M\<rbrakk> \<Longrightarrow> fic (ImpR (x).<a>.M b) b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3379
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3380
equivariance fic
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3381
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3382
lemma fic_Ax_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3383
  assumes a: "fic (Ax x a) b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3384
  shows "a=b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3385
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3386
apply(erule_tac fic.cases)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3387
apply(auto simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3388
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3389
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3390
lemma fic_NotR_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3391
  assumes a: "fic (NotR (x).M a) b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3392
  shows "a=b \<and> b\<sharp>M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3393
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3394
apply(erule_tac fic.cases)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3395
apply(auto simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3396
apply(subgoal_tac "b\<sharp>[xa].Ma")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3397
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3398
apply(simp_all add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3399
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3400
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3401
lemma fic_OrR1_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3402
  assumes a: "fic (OrR1 <a>.M b) c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3403
  shows "b=c \<and> c\<sharp>[a].M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3404
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3405
apply(erule_tac fic.cases)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3406
apply(auto simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3407
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3408
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3409
lemma fic_OrR2_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3410
  assumes a: "fic (OrR2 <a>.M b) c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3411
  shows "b=c \<and> c\<sharp>[a].M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3412
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3413
apply(erule_tac fic.cases)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3414
apply(auto simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3415
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3416
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3417
lemma fic_AndR_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3418
  assumes a: "fic (AndR <a>.M <b>.N c) d"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3419
  shows "c=d \<and> d\<sharp>[a].M \<and> d\<sharp>[b].N"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3420
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3421
apply(erule_tac fic.cases)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3422
apply(auto simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3423
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3424
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3425
lemma fic_ImpR_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3426
  assumes a: "fic (ImpR (x).<a>.M b) c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3427
  shows "b=c \<and> b\<sharp>[a].M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3428
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3429
apply(erule_tac fic.cases)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3430
apply(auto simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3431
apply(subgoal_tac "c\<sharp>[xa].[aa].Ma")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3432
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3433
apply(simp_all add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3434
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3435
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3436
lemma fic_rest_elims:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3437
  shows "fic (Cut <a>.M (x).N) d \<Longrightarrow> False"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3438
  and   "fic (NotL <a>.M x) d \<Longrightarrow> False"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3439
  and   "fic (OrL (x).M (y).N z) d \<Longrightarrow> False"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3440
  and   "fic (AndL1 (x).M y) d \<Longrightarrow> False"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3441
  and   "fic (AndL2 (x).M y) d \<Longrightarrow> False"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3442
  and   "fic (ImpL <a>.M (x).N y) d \<Longrightarrow> False"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3443
by (erule fic.cases, simp_all add: trm.inject)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3444
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3445
lemmas fic_elims = fic_Ax_elim fic_NotR_elim fic_OrR1_elim fic_OrR2_elim fic_AndR_elim 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3446
                   fic_ImpR_elim fic_rest_elims
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3447
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3448
lemma fic_rename:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3449
  shows "fic M a \<Longrightarrow> fic ([(a',a)]\<bullet>M) a'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3450
by (induct rule: fic.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3451
   (auto simp add: calc_atm simp add: fresh_left abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3452
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3453
lemma not_fic_subst1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3454
  assumes a: "\<not>(fic M a)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3455
  shows "\<not>(fic (M{y:=<c>.P}) a)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3456
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3457
apply(nominal_induct M avoiding: a c y P rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3458
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3459
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3460
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3461
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3462
apply(drule fic_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3463
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3464
apply(drule freshc_after_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3465
apply(simp add: fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3466
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<c>.P},P,a)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3467
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3468
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3469
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3470
apply(simp add: fresh_fun_simp_NotL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3471
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3472
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3473
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3474
apply(drule fic_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3475
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3476
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3477
apply(drule freshc_after_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3478
apply(drule freshc_after_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3479
apply(simp add: fic.intros abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3480
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<c>.P},P,name1,a)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3481
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3482
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3483
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3484
apply(simp add: fresh_fun_simp_AndL1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3485
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3486
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3487
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3488
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{y:=<c>.P},P,name1,a)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3489
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3490
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3491
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3492
apply(simp add: fresh_fun_simp_AndL2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3493
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3494
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3495
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3496
apply(drule fic_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3497
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3498
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3499
apply(drule freshc_after_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3500
apply(simp add: fic.intros abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3501
apply(drule fic_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3502
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3503
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3504
apply(drule freshc_after_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3505
apply(simp add: fic.intros abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3506
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{y:=<c>.P},trm2{y:=<c>.P},P,name1,name2,a)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3507
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3508
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3509
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3510
apply(simp add: fresh_fun_simp_OrL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3511
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3512
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3513
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3514
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3515
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3516
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3517
apply(drule freshc_after_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3518
apply(simp add: fic.intros abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3519
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{name2:=<c>.P},trm2{name2:=<c>.P},P,name1,name2,a)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3520
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3521
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3522
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3523
apply(simp add: fresh_fun_simp_ImpL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3524
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3525
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3526
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3527
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3528
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3529
lemma not_fic_subst2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3530
  assumes a: "\<not>(fic M a)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3531
  shows "\<not>(fic (M{c:=(y).P}) a)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3532
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3533
apply(nominal_induct M avoiding: a c y P rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3534
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3535
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3536
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3537
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3538
apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(trm{coname:=(y).P},P,a)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3539
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3540
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3541
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3542
apply(simp add: fresh_fun_simp_NotR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3543
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3544
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3545
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3546
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3547
apply(drule freshc_after_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3548
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3549
apply(simp add: fic.intros abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3550
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3551
apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(trm1{coname3:=(y).P},trm2{coname3:=(y).P},P,coname1,coname2,a)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3552
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3553
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3554
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3555
apply(simp add: fresh_fun_simp_AndR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3556
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3557
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3558
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3559
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3560
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3561
apply(drule freshc_after_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3562
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3563
apply(drule freshc_after_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3564
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3565
apply(simp add: fic.intros abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3566
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3567
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3568
apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(trm{coname2:=(y).P},P,coname1,a)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3569
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3570
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3571
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3572
apply(simp add: fresh_fun_simp_OrR1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3573
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3574
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3575
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3576
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3577
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3578
apply(drule freshc_after_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3579
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3580
apply(simp add: fic.intros abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3581
apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(trm{coname2:=(y).P},P,coname1,a)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3582
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3583
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3584
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3585
apply(simp add: fresh_fun_simp_OrR2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3586
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3587
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3588
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3589
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3590
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3591
apply(drule freshc_after_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3592
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3593
apply(simp add: fic.intros abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3594
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3595
apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(trm{coname2:=(y).P},P,coname1,a)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3596
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3597
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3598
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3599
apply(simp add: fresh_fun_simp_ImpR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3600
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3601
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3602
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3603
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3604
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3605
apply(drule freshc_after_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3606
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3607
apply(simp add: fic.intros abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3608
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3609
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3610
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3611
lemma fic_subst1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3612
  assumes a: "fic M a" "a\<noteq>b" "a\<sharp>P"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3613
  shows "fic (M{b:=(x).P}) a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3614
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3615
apply(nominal_induct M avoiding: x b a P rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3616
apply(drule fic_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3617
apply(simp add: fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3618
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3619
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3620
apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3621
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3622
apply(rule subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3623
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3624
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3625
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3626
apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3627
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3628
apply(rule subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3629
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3630
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3631
apply(rule subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3632
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3633
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3634
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3635
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3636
apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3637
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3638
apply(rule subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3639
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3640
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3641
apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3642
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3643
apply(rule subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3644
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3645
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3646
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3647
apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3648
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3649
apply(rule subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3650
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3651
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3652
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3653
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3654
lemma fic_subst2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3655
  assumes a: "fic M a" "c\<noteq>a" "a\<sharp>P" "M\<noteq>Ax x a" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3656
  shows "fic (M{x:=<c>.P}) a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3657
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3658
apply(nominal_induct M avoiding: x a c P rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3659
apply(drule fic_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3660
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3661
apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3662
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3663
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3664
apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3665
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3666
apply(rule subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3667
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3668
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3669
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3670
apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3671
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3672
apply(rule subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3673
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3674
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3675
apply(rule subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3676
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3677
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3678
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3679
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3680
apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3681
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3682
apply(rule subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3683
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3684
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3685
apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3686
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3687
apply(rule subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3688
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3689
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3690
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3691
apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3692
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3693
apply(rule subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3694
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3695
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3696
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3697
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3698
lemma fic_substc_crename:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3699
  assumes a: "fic M a" "a\<noteq>b" "a\<sharp>P"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3700
  shows "M[a\<turnstile>c>b]{b:=(y).P} = Cut <a>.(M{b:=(y).P}) (y).P"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3701
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3702
apply(nominal_induct M avoiding: a b  y P rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3703
apply(drule fic_Ax_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3704
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3705
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3706
apply(simp add: alpha calc_atm fresh_atm trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3707
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3708
apply(drule fic_rest_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3709
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3710
apply(drule fic_NotR_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3711
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3712
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3713
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3714
apply(simp add: trm.inject alpha fresh_atm fresh_prod fresh_atm calc_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3715
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3716
apply(simp add: csubst_eqvt calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3717
apply(simp add: perm_fresh_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3718
apply(simp add: crename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3719
apply(rule subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3720
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3721
apply(drule fic_rest_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3722
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3723
apply(drule fic_AndR_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3724
apply(simp add: abs_fresh fresh_atm subst_fresh rename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3725
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3726
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3727
apply(simp add: trm.inject alpha fresh_atm calc_atm abs_fresh fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3728
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3729
apply(simp add: csubst_eqvt calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3730
apply(simp add: perm_fresh_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3731
apply(simp add: csubst_eqvt calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3732
apply(simp add: perm_fresh_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3733
apply(simp add: subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3734
apply(drule fic_rest_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3735
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3736
apply(drule fic_rest_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3737
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3738
apply(drule fic_OrR1_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3739
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3740
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3741
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3742
apply(simp add: trm.inject alpha fresh_atm calc_atm abs_fresh fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3743
apply(simp add: csubst_eqvt calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3744
apply(simp add: perm_fresh_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3745
apply(simp add: subst_fresh rename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3746
apply(drule fic_OrR2_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3747
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3748
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3749
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3750
apply(simp add: trm.inject alpha fresh_atm calc_atm abs_fresh fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3751
apply(simp add: csubst_eqvt calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3752
apply(simp add: perm_fresh_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3753
apply(simp add: subst_fresh rename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3754
apply(drule fic_rest_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3755
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3756
apply(drule fic_ImpR_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3757
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3758
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3759
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3760
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3761
apply(simp add: trm.inject alpha fresh_atm calc_atm abs_fresh fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3762
apply(simp add: csubst_eqvt calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3763
apply(simp add: perm_fresh_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3764
apply(simp add: subst_fresh rename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3765
apply(drule fic_rest_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3766
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3767
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3768
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3769
inductive2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3770
  l_redu :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<longrightarrow>\<^isub>l _" [100,100] 100)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3771
where
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3772
  LAxR:  "\<lbrakk>x\<sharp>M; a\<sharp>b; fic M a\<rbrakk> \<Longrightarrow> Cut <a>.M (x).(Ax x b) \<longrightarrow>\<^isub>l M[a\<turnstile>c>b]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3773
| LAxL:  "\<lbrakk>a\<sharp>M; x\<sharp>y; fin M x\<rbrakk> \<Longrightarrow> Cut <a>.(Ax y a) (x).M \<longrightarrow>\<^isub>l M[x\<turnstile>n>y]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3774
| 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>
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3775
          Cut <a>.(NotR (x).M a) (y).(NotL <b>.N y) \<longrightarrow>\<^isub>l Cut <b>.N (x).M" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3776
| 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>
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3777
          Cut <b>.(AndR <a1>.M1 <a2>.M2 b) (y).(AndL1 (x).N y) \<longrightarrow>\<^isub>l Cut <a1>.M1 (x).N"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3778
| 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>
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3779
          Cut <b>.(AndR <a1>.M1 <a2>.M2 b) (y).(AndL2 (x).N y) \<longrightarrow>\<^isub>l Cut <a2>.M2 (x).N"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3780
| 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>
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3781
          Cut <b>.(OrR1 <a>.M b) (y).(OrL (x1).N1 (x2).N2 y) \<longrightarrow>\<^isub>l Cut <a>.M (x1).N1"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3782
| 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>
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3783
          Cut <b>.(OrR2 <a>.M b) (y).(OrL (x1).N1 (x2).N2 y) \<longrightarrow>\<^isub>l Cut <a>.M (x2).N2"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3784
| 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); 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3785
          c\<sharp>(P,[a].M,b,a); a\<sharp>([c].N,P); y\<sharp>(N,[x].M)\<rbrakk> \<Longrightarrow>
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3786
          Cut <b>.(ImpR (x).<a>.M b) (z).(ImpL <c>.N (y).P z) \<longrightarrow>\<^isub>l Cut <a>.(Cut <c>.N (x).M) (y).P"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3787
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3788
equivariance l_redu
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3789
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3790
lemma l_redu_eqvt':
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3791
  fixes pi1::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3792
  and   pi2::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3793
  shows "(pi1\<bullet>M) \<longrightarrow>\<^isub>l (pi1\<bullet>M') \<Longrightarrow> M \<longrightarrow>\<^isub>l M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3794
  and   "(pi2\<bullet>M) \<longrightarrow>\<^isub>l (pi2\<bullet>M') \<Longrightarrow> M \<longrightarrow>\<^isub>l M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3795
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3796
apply(drule_tac pi="rev pi1" in l_redu.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3797
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3798
apply(drule_tac pi="rev pi2" in l_redu.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3799
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3800
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3801
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3802
nominal_inductive l_redu
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3803
  apply(simp_all add: abs_fresh fresh_atm rename_fresh fresh_prod abs_supp fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3804
  apply(force)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3805
  done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3806
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3807
lemma fresh_l_redu:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3808
  fixes x::"name"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3809
  and   a::"coname"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3810
  shows "M \<longrightarrow>\<^isub>l M' \<Longrightarrow> x\<sharp>M \<Longrightarrow> x\<sharp>M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3811
  and   "M \<longrightarrow>\<^isub>l M' \<Longrightarrow> a\<sharp>M \<Longrightarrow> a\<sharp>M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3812
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3813
apply(induct rule: l_redu.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3814
apply(auto simp add: abs_fresh rename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3815
apply(case_tac "xa=x")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3816
apply(simp add: rename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3817
apply(simp add: rename_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3818
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3819
apply(induct rule: l_redu.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3820
apply(auto simp add: abs_fresh rename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3821
apply(case_tac "aa=a")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3822
apply(simp add: rename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3823
apply(simp add: rename_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3824
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3825
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3826
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3827
lemma better_LAxR_intro[intro]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3828
  shows "fic M a \<Longrightarrow> Cut <a>.M (x).(Ax x b) \<longrightarrow>\<^isub>l M[a\<turnstile>c>b]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3829
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3830
  assume fin: "fic M a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3831
  obtain x'::"name" where fs1: "x'\<sharp>(M,x)" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3832
  obtain a'::"coname" where fs2: "a'\<sharp>(a,M,b)" by (rule exists_fresh(2), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3833
  have "Cut <a>.M (x).(Ax x b) =  Cut <a'>.([(a',a)]\<bullet>M) (x').(Ax x' b)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3834
    using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3835
  also have "\<dots> \<longrightarrow>\<^isub>l ([(a',a)]\<bullet>M)[a'\<turnstile>c>b]" using fs1 fs2 fin
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3836
    by (auto intro: l_redu.intros simp add: fresh_left calc_atm fic_rename)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3837
  also have "\<dots> = M[a\<turnstile>c>b]" using fs1 fs2 by (simp add: crename_rename)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3838
  finally show ?thesis by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3839
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3840
    
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3841
lemma better_LAxL_intro[intro]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3842
  shows "fin M x \<Longrightarrow> Cut <a>.(Ax y a) (x).M \<longrightarrow>\<^isub>l M[x\<turnstile>n>y]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3843
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3844
  assume fin: "fin M x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3845
  obtain x'::"name" where fs1: "x'\<sharp>(y,M,x)" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3846
  obtain a'::"coname" where fs2: "a'\<sharp>(a,M)" by (rule exists_fresh(2), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3847
  have "Cut <a>.(Ax y a) (x).M = Cut <a'>.(Ax y a') (x').([(x',x)]\<bullet>M)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3848
    using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3849
  also have "\<dots> \<longrightarrow>\<^isub>l ([(x',x)]\<bullet>M)[x'\<turnstile>n>y]" using fs1 fs2 fin
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3850
    by (auto intro: l_redu.intros simp add: fresh_left calc_atm fin_rename)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3851
  also have "\<dots> = M[x\<turnstile>n>y]" using fs1 fs2 by (simp add: nrename_rename)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3852
  finally show ?thesis by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3853
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3854
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3855
lemma better_LNot_intro[intro]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3856
  shows "\<lbrakk>y\<sharp>N; a\<sharp>M\<rbrakk> \<Longrightarrow> Cut <a>.(NotR (x).M a) (y).(NotL <b>.N y) \<longrightarrow>\<^isub>l Cut <b>.N (x).M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3857
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3858
  assume fs: "y\<sharp>N" "a\<sharp>M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3859
  obtain x'::"name" where f1: "x'\<sharp>(y,N,M,x)" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3860
  obtain y'::"name" where f2: "y'\<sharp>(y,N,M,x,x')" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3861
  obtain a'::"coname" where f3: "a'\<sharp>(a,M,N,b)" by (rule exists_fresh(2), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3862
  obtain b'::"coname" where f4: "b'\<sharp>(a,M,N,b,a')" by (rule exists_fresh(2), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3863
  have "Cut <a>.(NotR (x).M a) (y).(NotL <b>.N y) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3864
                      = Cut <a'>.(NotR (x).([(a',a)]\<bullet>M) a') (y').(NotL <b>.([(y',y)]\<bullet>N) y')"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3865
    using f1 f2 f3 f4 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3866
    by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3867
  also have "\<dots> = Cut <a'>.(NotR (x).M a') (y').(NotL <b>.N y')"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3868
    using f1 f2 f3 f4 fs by (perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3869
  also have "\<dots> = Cut <a'>.(NotR (x').([(x',x)]\<bullet>M) a') (y').(NotL <b'>.([(b',b)]\<bullet>N) y')"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3870
    using f1 f2 f3 f4 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3871
    by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3872
  also have "\<dots> \<longrightarrow>\<^isub>l Cut <b'>.([(b',b)]\<bullet>N) (x').([(x',x)]\<bullet>M)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3873
    using f1 f2 f3 f4 fs
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3874
    by (auto intro:  l_redu.intros simp add: fresh_prod fresh_left calc_atm fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3875
  also have "\<dots> = Cut <b>.N (x).M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3876
    using f1 f2 f3 f4 by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3877
  finally show ?thesis by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3878
qed 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3879
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3880
lemma better_LAnd1_intro[intro]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3881
  shows "\<lbrakk>a\<sharp>([b1].M1,[b2].M2); y\<sharp>[x].N\<rbrakk> 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3882
         \<Longrightarrow> Cut <a>.(AndR <b1>.M1 <b2>.M2 a) (y).(AndL1 (x).N y) \<longrightarrow>\<^isub>l Cut <b1>.M1 (x).N"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3883
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3884
  assume fs: "a\<sharp>([b1].M1,[b2].M2)" "y\<sharp>[x].N"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3885
  obtain x'::"name" where f1: "x'\<sharp>(y,N,M1,M2,x)" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3886
  obtain y'::"name" where f2: "y'\<sharp>(y,N,M1,M2,x,x')" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3887
  obtain a'::"coname" where f3: "a'\<sharp>(a,M1,M2,N,b1,b2)" by (rule exists_fresh(2), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3888
  obtain b1'::"coname" where f4:"b1'\<sharp>(a,M1,M2,N,b1,b2,a')" by (rule exists_fresh(2), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3889
  obtain b2'::"coname" where f5:"b2'\<sharp>(a,M1,M2,N,b1,b2,a',b1')" by (rule exists_fresh(2),rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3890
  have "Cut <a>.(AndR <b1>.M1 <b2>.M2 a) (y).(AndL1 (x).N y)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3891
                      = Cut <a'>.(AndR <b1>.M1 <b2>.M2 a') (y').(AndL1 (x).N y')"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3892
    using f1 f2 f3 f4 fs
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3893
    apply(rule_tac sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3894
    apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3895
    apply(auto simp add: perm_fresh_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3896
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3897
  also have "\<dots> = Cut <a'>.(AndR <b1'>.([(b1',b1)]\<bullet>M1) <b2'>.([(b2',b2)]\<bullet>M2) a') 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3898
                                                               (y').(AndL1 (x').([(x',x)]\<bullet>N) y')"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3899
    using f1 f2 f3 f4 f5 fs 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3900
    apply(rule_tac sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3901
    apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3902
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3903
  also have "\<dots> \<longrightarrow>\<^isub>l Cut <b1'>.([(b1',b1)]\<bullet>M1) (x').([(x',x)]\<bullet>N)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3904
    using f1 f2 f3 f4 f5 fs
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3905
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3906
    apply(rule l_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3907
    apply(auto simp add: abs_fresh fresh_prod fresh_left calc_atm fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3908
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3909
  also have "\<dots> = Cut <b1>.M1 (x).N"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3910
    using f1 f2 f3 f4 f5 fs by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3911
  finally show ?thesis by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3912
qed 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3913
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3914
lemma better_LAnd2_intro[intro]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3915
  shows "\<lbrakk>a\<sharp>([b1].M1,[b2].M2); y\<sharp>[x].N\<rbrakk> 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3916
         \<Longrightarrow> Cut <a>.(AndR <b1>.M1 <b2>.M2 a) (y).(AndL2 (x).N y) \<longrightarrow>\<^isub>l Cut <b2>.M2 (x).N"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3917
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3918
  assume fs: "a\<sharp>([b1].M1,[b2].M2)" "y\<sharp>[x].N"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3919
  obtain x'::"name" where f1: "x'\<sharp>(y,N,M1,M2,x)" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3920
  obtain y'::"name" where f2: "y'\<sharp>(y,N,M1,M2,x,x')" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3921
  obtain a'::"coname" where f3: "a'\<sharp>(a,M1,M2,N,b1,b2)" by (rule exists_fresh(2), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3922
  obtain b1'::"coname" where f4:"b1'\<sharp>(a,M1,M2,N,b1,b2,a')" by (rule exists_fresh(2), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3923
  obtain b2'::"coname" where f5:"b2'\<sharp>(a,M1,M2,N,b1,b2,a',b1')" by (rule exists_fresh(2),rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3924
  have "Cut <a>.(AndR <b1>.M1 <b2>.M2 a) (y).(AndL2 (x).N y)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3925
                      = Cut <a'>.(AndR <b1>.M1 <b2>.M2 a') (y').(AndL2 (x).N y')"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3926
    using f1 f2 f3 f4 fs
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3927
    apply(rule_tac sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3928
    apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3929
    apply(auto simp add: perm_fresh_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3930
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3931
  also have "\<dots> = Cut <a'>.(AndR <b1'>.([(b1',b1)]\<bullet>M1) <b2'>.([(b2',b2)]\<bullet>M2) a') 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3932
                                                               (y').(AndL2 (x').([(x',x)]\<bullet>N) y')"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3933
    using f1 f2 f3 f4 f5 fs 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3934
    apply(rule_tac sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3935
    apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3936
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3937
  also have "\<dots> \<longrightarrow>\<^isub>l Cut <b2'>.([(b2',b2)]\<bullet>M2) (x').([(x',x)]\<bullet>N)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3938
    using f1 f2 f3 f4 f5 fs
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3939
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3940
    apply(rule l_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3941
    apply(auto simp add: abs_fresh fresh_prod fresh_left calc_atm fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3942
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3943
  also have "\<dots> = Cut <b2>.M2 (x).N"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3944
    using f1 f2 f3 f4 f5 fs by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3945
  finally show ?thesis by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3946
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3947
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3948
lemma better_LOr1_intro[intro]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3949
  shows "\<lbrakk>y\<sharp>([x1].N1,[x2].N2); b\<sharp>[a].M\<rbrakk> 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3950
         \<Longrightarrow> Cut <b>.(OrR1 <a>.M b) (y).(OrL (x1).N1 (x2).N2 y) \<longrightarrow>\<^isub>l Cut <a>.M (x1).N1"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3951
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3952
  assume fs: "y\<sharp>([x1].N1,[x2].N2)" "b\<sharp>[a].M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3953
  obtain y'::"name" where f1: "y'\<sharp>(y,M,N1,N2,x1,x2)" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3954
  obtain x1'::"name" where f2: "x1'\<sharp>(y,M,N1,N2,x1,x2,y')" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3955
  obtain x2'::"name" where f3: "x2'\<sharp>(y,M,N1,N2,x1,x2,y',x1')" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3956
  obtain a'::"coname" where f4: "a'\<sharp>(a,N1,N2,M,b)" by (rule exists_fresh(2), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3957
  obtain b'::"coname" where f5: "b'\<sharp>(a,N1,N2,M,b,a')" by (rule exists_fresh(2),rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3958
  have "Cut <b>.(OrR1 <a>.M b) (y).(OrL (x1).N1 (x2).N2 y)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3959
                      = Cut <b'>.(OrR1 <a>.M b') (y').(OrL (x1).N1 (x2).N2 y')"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3960
    using f1 f2 f3 f4 f5 fs
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3961
    apply(rule_tac sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3962
    apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3963
    apply(auto simp add: perm_fresh_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3964
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3965
  also have "\<dots> = Cut <b'>.(OrR1 <a'>.([(a',a)]\<bullet>M) b') 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3966
              (y').(OrL (x1').([(x1',x1)]\<bullet>N1) (x2').([(x2',x2)]\<bullet>N2) y')"   
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3967
    using f1 f2 f3 f4 f5 fs 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3968
    apply(rule_tac sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3969
    apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3970
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3971
  also have "\<dots> \<longrightarrow>\<^isub>l Cut <a'>.([(a',a)]\<bullet>M) (x1').([(x1',x1)]\<bullet>N1)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3972
    using f1 f2 f3 f4 f5 fs
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3973
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3974
    apply(rule l_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3975
    apply(auto simp add: abs_fresh fresh_prod fresh_left calc_atm fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3976
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3977
  also have "\<dots> = Cut <a>.M (x1).N1"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3978
    using f1 f2 f3 f4 f5 fs by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3979
  finally show ?thesis by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3980
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3981
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3982
lemma better_LOr2_intro[intro]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3983
  shows "\<lbrakk>y\<sharp>([x1].N1,[x2].N2); b\<sharp>[a].M\<rbrakk> 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3984
         \<Longrightarrow> Cut <b>.(OrR2 <a>.M b) (y).(OrL (x1).N1 (x2).N2 y) \<longrightarrow>\<^isub>l Cut <a>.M (x2).N2"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3985
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3986
  assume fs: "y\<sharp>([x1].N1,[x2].N2)" "b\<sharp>[a].M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3987
  obtain y'::"name" where f1: "y'\<sharp>(y,M,N1,N2,x1,x2)" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3988
  obtain x1'::"name" where f2: "x1'\<sharp>(y,M,N1,N2,x1,x2,y')" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3989
  obtain x2'::"name" where f3: "x2'\<sharp>(y,M,N1,N2,x1,x2,y',x1')" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3990
  obtain a'::"coname" where f4: "a'\<sharp>(a,N1,N2,M,b)" by (rule exists_fresh(2), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3991
  obtain b'::"coname" where f5: "b'\<sharp>(a,N1,N2,M,b,a')" by (rule exists_fresh(2),rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3992
  have "Cut <b>.(OrR2 <a>.M b) (y).(OrL (x1).N1 (x2).N2 y)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3993
                      = Cut <b'>.(OrR2 <a>.M b') (y').(OrL (x1).N1 (x2).N2 y')"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3994
    using f1 f2 f3 f4 f5 fs
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3995
    apply(rule_tac sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3996
    apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3997
    apply(auto simp add: perm_fresh_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3998
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  3999
  also have "\<dots> = Cut <b'>.(OrR2 <a'>.([(a',a)]\<bullet>M) b') 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4000
              (y').(OrL (x1').([(x1',x1)]\<bullet>N1) (x2').([(x2',x2)]\<bullet>N2) y')"   
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4001
    using f1 f2 f3 f4 f5 fs 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4002
    apply(rule_tac sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4003
    apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4004
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4005
  also have "\<dots> \<longrightarrow>\<^isub>l Cut <a'>.([(a',a)]\<bullet>M) (x2').([(x2',x2)]\<bullet>N2)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4006
    using f1 f2 f3 f4 f5 fs
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4007
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4008
    apply(rule l_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4009
    apply(auto simp add: abs_fresh fresh_prod fresh_left calc_atm fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4010
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4011
  also have "\<dots> = Cut <a>.M (x2).N2"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4012
    using f1 f2 f3 f4 f5 fs by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4013
  finally show ?thesis by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4014
qed 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4015
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4016
lemma better_LImp_intro[intro]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4017
  shows "\<lbrakk>z\<sharp>(N,[y].P); b\<sharp>[a].M; a\<sharp>N\<rbrakk> 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4018
         \<Longrightarrow> Cut <b>.(ImpR (x).<a>.M b) (z).(ImpL <c>.N (y).P z) \<longrightarrow>\<^isub>l Cut <a>.(Cut <c>.N (x).M) (y).P"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4019
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4020
  assume fs: "z\<sharp>(N,[y].P)" "b\<sharp>[a].M" "a\<sharp>N"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4021
  obtain y'::"name" where f1: "y'\<sharp>(y,M,N,P,z,x)" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4022
  obtain x'::"name" where f2: "x'\<sharp>(y,M,N,P,z,x,y')" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4023
  obtain z'::"name" where f3: "z'\<sharp>(y,M,N,P,z,x,y',x')" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4024
  obtain a'::"coname" where f4: "a'\<sharp>(a,N,P,M,b)" by (rule exists_fresh(2), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4025
  obtain b'::"coname" where f5: "b'\<sharp>(a,N,P,M,b,c,a')" by (rule exists_fresh(2),rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4026
  obtain c'::"coname" where f6: "c'\<sharp>(a,N,P,M,b,c,a',b')" by (rule exists_fresh(2),rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4027
  have " Cut <b>.(ImpR (x).<a>.M b) (z).(ImpL <c>.N (y).P z)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4028
                      =  Cut <b'>.(ImpR (x).<a>.M b') (z').(ImpL <c>.N (y).P z')"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4029
    using f1 f2 f3 f4 f5 fs
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4030
    apply(rule_tac sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4031
    apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4032
    apply(auto simp add: perm_fresh_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4033
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4034
  also have "\<dots> = Cut <b'>.(ImpR (x').<a'>.([(a',a)]\<bullet>([(x',x)]\<bullet>M)) b') 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4035
                           (z').(ImpL <c'>.([(c',c)]\<bullet>N) (y').([(y',y)]\<bullet>P) z')"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4036
    using f1 f2 f3 f4 f5 f6 fs 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4037
    apply(rule_tac sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4038
    apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4039
    apply(simp add: alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4040
    apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4041
    apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4042
    apply(simp add: alpha fresh_prod fresh_atm abs_perm calc_atm fresh_left abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4043
    apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4044
    apply(simp add: alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4045
    apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4046
    apply(simp add: alpha fresh_prod fresh_atm abs_perm calc_atm fresh_left abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4047
    apply(simp add: alpha fresh_prod fresh_atm abs_perm calc_atm fresh_left abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4048
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4049
  also have "\<dots> \<longrightarrow>\<^isub>l Cut <a'>.(Cut <c'>.([(c',c)]\<bullet>N) (x').([(a',a)]\<bullet>[(x',x)]\<bullet>M)) (y').([(y',y)]\<bullet>P)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4050
    using f1 f2 f3 f4 f5 f6 fs
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4051
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4052
    apply(rule l_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4053
    apply(auto simp add: abs_fresh fresh_prod fresh_left calc_atm fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4054
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4055
  also have "\<dots> = Cut <a>.(Cut <c>.N (x).M) (y).P"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4056
    using f1 f2 f3 f4 f5 f6 fs 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4057
    apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4058
    apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4059
    apply(simp add: alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4060
    apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4061
    apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4062
    apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4063
    apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4064
    apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4065
    apply(perm_simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4066
    apply(auto simp add: fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4067
    apply(perm_simp add: alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4068
    apply(perm_simp add: alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4069
    apply(perm_simp add: alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4070
    apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4071
    apply(perm_simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4072
    apply(rule_tac pi="[(a',a)]" in pt_bij4[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4073
    apply(perm_simp add: abs_perm calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4074
    apply(perm_simp add: alpha fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4075
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4076
    apply(perm_simp add: alpha fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4077
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4078
  finally show ?thesis by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4079
qed 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4080
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4081
lemma alpha_coname:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4082
  fixes M::"trm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4083
  and   a::"coname"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4084
  assumes a: "[a].M = [b].N" "c\<sharp>(a,b,M,N)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4085
  shows "M = [(a,c)]\<bullet>[(b,c)]\<bullet>N"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4086
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4087
apply(auto simp add: alpha' fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4088
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4089
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4090
done 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4091
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4092
lemma alpha_name:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4093
  fixes M::"trm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4094
  and   x::"name"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4095
  assumes a: "[x].M = [y].N" "z\<sharp>(x,y,M,N)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4096
  shows "M = [(x,z)]\<bullet>[(y,z)]\<bullet>N"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4097
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4098
apply(auto simp add: alpha' fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4099
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4100
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4101
done 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4102
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4103
lemma alpha_name_coname:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4104
  fixes M::"trm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4105
  and   x::"name"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4106
  and   a::"coname"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4107
  assumes a: "[x].[b].M = [y].[c].N" "z\<sharp>(x,y,M,N)" "a\<sharp>(b,c,M,N)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4108
  shows "M = [(x,z)]\<bullet>[(b,a)]\<bullet>[(c,a)]\<bullet>[(y,z)]\<bullet>N"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4109
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4110
apply(auto simp add: alpha' fresh_prod fresh_atm abs_supp fin_supp abs_fresh abs_perm fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4111
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4112
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4113
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4114
done 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4115
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4116
lemma Cut_l_redu_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4117
  assumes a: "Cut <a>.M (x).N \<longrightarrow>\<^isub>l R"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4118
  shows "(\<exists>b. R = M[a\<turnstile>c>b]) \<or> (\<exists>y. R = N[x\<turnstile>n>y]) \<or>
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4119
  (\<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>
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4120
  (\<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' 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4121
                                                                            \<and> fic M a \<and> fin N x) \<or>
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4122
  (\<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' 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4123
                                                                            \<and> fic M a \<and> fin N x) \<or>
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4124
  (\<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 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4125
                                                                            \<and> fic M a \<and> fin N x) \<or>
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4126
  (\<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 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4127
                                                                            \<and> fic M a \<and> fin N x) \<or>
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4128
  (\<exists>z b M' c N1 y N2. M = ImpR (z).<b>.M' a \<and> N = ImpL <c>.N1 (y).N2 x \<and> 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4129
            R = Cut <b>.(Cut <c>.N1 (z).M') (y).N2 \<and> b\<sharp>(c,N1) \<and> fic M a \<and> fin N x)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4130
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4131
apply(erule_tac l_redu.cases)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4132
apply(rule disjI1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4133
(* ax case *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4134
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4135
apply(rule_tac x="b" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4136
apply(erule conjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4137
apply(simp add: alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4138
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4139
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4140
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4141
apply(simp add: rename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4142
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4143
apply(rule disjI1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4144
(* ax case *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4145
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4146
apply(rule_tac x="y" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4147
apply(erule conjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4148
apply(thin_tac "[a].M = [aa].Ax y aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4149
apply(simp add: alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4150
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4151
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4152
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4153
apply(simp add: rename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4154
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4155
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4156
apply(rule disjI1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4157
(* not case *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4158
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4159
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4160
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4161
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4162
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4163
apply(drule_tac c="c" in  alpha_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4164
apply(simp add: fresh_prod fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4165
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4166
apply(rule exI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4167
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4168
apply(rule refl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4169
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4170
apply(simp add: calc_atm abs_fresh fresh_prod fresh_atm fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4171
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4172
apply(drule_tac z="ca" in  alpha_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4173
apply(simp add: fresh_prod fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4174
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4175
apply(rule exI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4176
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4177
apply(rule refl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4178
apply(auto simp add: calc_atm abs_fresh fresh_left)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4179
apply(case_tac "y=x")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4180
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4181
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4182
apply(case_tac "aa=a")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4183
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4184
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4185
(* and1 case *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4186
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4187
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4188
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4189
apply(rule disjI1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4190
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4191
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4192
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4193
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4194
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4195
apply(drule_tac c="c" in  alpha_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4196
apply(simp add: fresh_prod fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4197
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4198
apply(rule exI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4199
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4200
apply(rule exI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4201
apply(rule_tac s="a" and t="[(a,c)]\<bullet>[(b,c)]\<bullet>b" in subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4202
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4203
apply(rule refl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4204
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4205
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4206
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4207
apply(drule_tac z="ca" in  alpha_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4208
apply(simp add: fresh_prod fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4209
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4210
apply(rule exI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4211
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4212
apply(rule_tac s="x" and t="[(x,ca)]\<bullet>[(y,ca)]\<bullet>y" in subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4213
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4214
apply(rule refl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4215
apply(auto simp add: fresh_left calc_atm abs_fresh split: if_splits)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4216
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4217
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4218
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4219
apply(drule_tac z="cb" in  alpha_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4220
apply(simp add: fresh_prod fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4221
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4222
apply(rule exI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4223
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4224
apply(rule_tac s="x" and t="[(x,cb)]\<bullet>[(y,cb)]\<bullet>y" in subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4225
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4226
apply(rule refl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4227
apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4228
apply(perm_simp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4229
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4230
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4231
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4232
apply(drule_tac z="cb" in  alpha_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4233
apply(simp add: fresh_prod fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4234
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4235
apply(rule exI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4236
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4237
apply(rule_tac s="x" and t="[(x,cb)]\<bullet>[(y,cb)]\<bullet>y" in subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4238
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4239
apply(rule refl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4240
apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4241
apply(perm_simp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4242
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4243
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4244
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4245
apply(drule_tac z="cb" in  alpha_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4246
apply(simp add: fresh_prod fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4247
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4248
apply(rule exI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4249
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4250
apply(rule_tac s="x" and t="[(x,cb)]\<bullet>[(y,cb)]\<bullet>y" in subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4251
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4252
apply(rule refl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4253
apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4254
apply(perm_simp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4255
(* and2 case *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4256
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4257
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4258
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4259
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4260
apply(rule disjI1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4261
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4262
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4263
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4264
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4265
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4266
apply(drule_tac c="c" in  alpha_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4267
apply(simp add: fresh_prod fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4268
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4269
apply(rule exI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4270
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4271
apply(rule_tac s="a" and t="[(a,c)]\<bullet>[(b,c)]\<bullet>b" in subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4272
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4273
apply(rule refl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4274
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4275
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4276
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4277
apply(drule_tac z="ca" in  alpha_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4278
apply(simp add: fresh_prod fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4279
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4280
apply(rule exI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4281
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4282
apply(rule_tac s="x" and t="[(x,ca)]\<bullet>[(y,ca)]\<bullet>y" in subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4283
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4284
apply(rule refl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4285
apply(auto simp add: fresh_left calc_atm abs_fresh split: if_splits)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4286
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4287
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4288
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4289
apply(drule_tac z="cb" in  alpha_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4290
apply(simp add: fresh_prod fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4291
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4292
apply(rule exI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4293
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4294
apply(rule_tac s="x" and t="[(x,cb)]\<bullet>[(y,cb)]\<bullet>y" in subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4295
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4296
apply(rule refl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4297
apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4298
apply(perm_simp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4299
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4300
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4301
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4302
apply(drule_tac z="cb" in  alpha_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4303
apply(simp add: fresh_prod fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4304
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4305
apply(rule exI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4306
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4307
apply(rule_tac s="x" and t="[(x,cb)]\<bullet>[(y,cb)]\<bullet>y" in subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4308
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4309
apply(rule refl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4310
apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4311
apply(perm_simp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4312
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4313
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4314
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4315
apply(drule_tac z="cb" in  alpha_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4316
apply(simp add: fresh_prod fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4317
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4318
apply(rule exI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4319
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4320
apply(rule_tac s="x" and t="[(x,cb)]\<bullet>[(y,cb)]\<bullet>y" in subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4321
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4322
apply(rule refl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4323
apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4324
apply(perm_simp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4325
(* or1 case *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4326
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4327
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4328
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4329
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4330
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4331
apply(rule disjI1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4332
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4333
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4334
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4335
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4336
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4337
apply(drule_tac c="c" in  alpha_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4338
apply(simp add: fresh_prod fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4339
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4340
apply(rule exI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4341
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4342
apply(rule_tac s="a" and t="[(a,c)]\<bullet>[(b,c)]\<bullet>b" in subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4343
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4344
apply(rule refl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4345
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4346
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4347
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4348
apply(drule_tac z="ca" in  alpha_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4349
apply(simp add: fresh_prod fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4350
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4351
apply(rule exI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4352
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4353
apply(rule exI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4354
apply(rule_tac s="x" and t="[(x,ca)]\<bullet>[(y,ca)]\<bullet>y" in subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4355
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4356
apply(rule refl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4357
apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4358
apply(perm_simp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4359
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4360
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4361
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4362
apply(drule_tac z="cb" in  alpha_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4363
apply(simp add: fresh_prod fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4364
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4365
apply(rule exI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4366
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4367
apply(rule exI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4368
apply(rule_tac s="x" and t="[(x,cb)]\<bullet>[(y,cb)]\<bullet>y" in subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4369
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4370
apply(rule refl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4371
apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4372
apply(perm_simp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4373
(* or2 case *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4374
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4375
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4376
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4377
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4378
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4379
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4380
apply(rule disjI1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4381
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4382
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4383
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4384
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4385
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4386
apply(drule_tac c="c" in  alpha_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4387
apply(simp add: fresh_prod fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4388
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4389
apply(rule exI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4390
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4391
apply(rule_tac s="a" and t="[(a,c)]\<bullet>[(b,c)]\<bullet>b" in subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4392
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4393
apply(rule refl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4394
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4395
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4396
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4397
apply(drule_tac z="ca" in  alpha_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4398
apply(simp add: fresh_prod fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4399
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4400
apply(rule exI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4401
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4402
apply(rule_tac s="x" and t="[(x,ca)]\<bullet>[(y,ca)]\<bullet>y" in subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4403
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4404
apply(rule refl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4405
apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4406
apply(perm_simp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4407
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4408
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4409
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4410
apply(drule_tac z="cb" in  alpha_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4411
apply(simp add: fresh_prod fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4412
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4413
apply(rule exI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4414
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4415
apply(rule_tac s="x" and t="[(x,cb)]\<bullet>[(y,cb)]\<bullet>y" in subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4416
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4417
apply(rule refl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4418
apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4419
apply(perm_simp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4420
(* imp-case *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4421
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4422
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4423
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4424
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4425
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4426
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4427
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4428
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4429
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4430
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4431
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4432
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4433
apply(drule_tac c="ca" in  alpha_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4434
apply(simp add: fresh_prod fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4435
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4436
apply(rule exI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4437
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4438
apply(rule_tac s="a" and t="[(a,ca)]\<bullet>[(b,ca)]\<bullet>b" in subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4439
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4440
apply(rule refl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4441
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4442
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4443
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4444
apply(drule_tac z="cb" in  alpha_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4445
apply(simp add: fresh_prod fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4446
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4447
apply(rule exI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4448
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4449
apply(rule_tac s="x" and t="[(x,cb)]\<bullet>[(z,cb)]\<bullet>z" in subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4450
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4451
apply(rule refl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4452
apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4453
apply(perm_simp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4454
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4455
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4456
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4457
apply(drule_tac z="cc" in  alpha_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4458
apply(simp add: fresh_prod fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4459
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4460
apply(rule exI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4461
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4462
apply(rule_tac s="x" and t="[(x,cc)]\<bullet>[(z,cc)]\<bullet>z" in subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4463
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4464
apply(rule refl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4465
apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4466
apply(perm_simp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4467
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4468
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4469
inductive2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4470
  c_redu :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<longrightarrow>\<^isub>c _" [100,100] 100)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4471
where
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4472
  left[intro]:  "\<lbrakk>\<not>fic M a; a\<sharp>N; x\<sharp>M\<rbrakk> \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>c M{a:=(x).N}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4473
| right[intro]: "\<lbrakk>\<not>fin N x; a\<sharp>N; x\<sharp>M\<rbrakk> \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>c N{x:=<a>.M}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4474
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4475
equivariance c_redu
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4476
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4477
nominal_inductive c_redu
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4478
 by (simp_all add: abs_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4479
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4480
lemma better_left[intro]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4481
  shows "\<not>fic M a \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>c M{a:=(x).N}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4482
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4483
  assume not_fic: "\<not>fic M a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4484
  obtain x'::"name" where fs1: "x'\<sharp>(N,M,x)" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4485
  obtain a'::"coname" where fs2: "a'\<sharp>(a,M,N)" by (rule exists_fresh(2), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4486
  have "Cut <a>.M (x).N =  Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4487
    using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4488
  also have "\<dots> \<longrightarrow>\<^isub>c ([(a',a)]\<bullet>M){a':=(x').([(x',x)]\<bullet>N)}" using fs1 fs2 not_fic
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4489
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4490
    apply(rule left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4491
    apply(clarify)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4492
    apply(drule_tac a'="a" in fic_rename)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4493
    apply(simp add: perm_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4494
    apply(simp add: fresh_left calc_atm)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4495
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4496
  also have "\<dots> = M{a:=(x).N}" using fs1 fs2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4497
    by (simp add: subst_rename[symmetric] fresh_atm fresh_prod fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4498
  finally show ?thesis by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4499
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4500
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4501
lemma better_right[intro]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4502
  shows "\<not>fin N x \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>c N{x:=<a>.M}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4503
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4504
  assume not_fin: "\<not>fin N x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4505
  obtain x'::"name" where fs1: "x'\<sharp>(N,M,x)" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4506
  obtain a'::"coname" where fs2: "a'\<sharp>(a,M,N)" by (rule exists_fresh(2), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4507
  have "Cut <a>.M (x).N =  Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4508
    using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4509
  also have "\<dots> \<longrightarrow>\<^isub>c ([(x',x)]\<bullet>N){x':=<a'>.([(a',a)]\<bullet>M)}" using fs1 fs2 not_fin
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4510
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4511
    apply(rule right)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4512
    apply(clarify)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4513
    apply(drule_tac x'="x" in fin_rename)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4514
    apply(simp add: perm_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4515
    apply(simp add: fresh_left calc_atm)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4516
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4517
  also have "\<dots> = N{x:=<a>.M}" using fs1 fs2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4518
    by (simp add: subst_rename[symmetric] fresh_atm fresh_prod fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4519
  finally show ?thesis by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4520
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4521
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4522
lemma fresh_c_redu:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4523
  fixes x::"name"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4524
  and   c::"coname"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4525
  shows "M \<longrightarrow>\<^isub>c M' \<Longrightarrow> x\<sharp>M \<Longrightarrow> x\<sharp>M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4526
  and   "M \<longrightarrow>\<^isub>c M' \<Longrightarrow> c\<sharp>M \<Longrightarrow> c\<sharp>M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4527
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4528
apply(induct rule: c_redu.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4529
apply(auto simp add: abs_fresh rename_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4530
apply(induct rule: c_redu.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4531
apply(auto simp add: abs_fresh rename_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4532
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4533
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4534
inductive2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4535
  a_redu :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<longrightarrow>\<^isub>a _" [100,100] 100)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4536
where
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4537
  al_redu[intro]: "M\<longrightarrow>\<^isub>l M' \<Longrightarrow> M \<longrightarrow>\<^isub>a M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4538
| ac_redu[intro]: "M\<longrightarrow>\<^isub>c M' \<Longrightarrow> M \<longrightarrow>\<^isub>a M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4539
| a_Cut_l: "\<lbrakk>a\<sharp>N; x\<sharp>M; M\<longrightarrow>\<^isub>a M'\<rbrakk> \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>a Cut <a>.M' (x).N"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4540
| a_Cut_r: "\<lbrakk>a\<sharp>N; x\<sharp>M; N\<longrightarrow>\<^isub>a N'\<rbrakk> \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>a Cut <a>.M (x).N'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4541
| a_NotL[intro]: "M\<longrightarrow>\<^isub>a M' \<Longrightarrow> NotL <a>.M x \<longrightarrow>\<^isub>a NotL <a>.M' x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4542
| a_NotR[intro]: "M\<longrightarrow>\<^isub>a M' \<Longrightarrow> NotR (x).M a \<longrightarrow>\<^isub>a NotR (x).M' a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4543
| a_AndR_l: "\<lbrakk>a\<sharp>(N,c); b\<sharp>(M,c); b\<noteq>a; M\<longrightarrow>\<^isub>a M'\<rbrakk> \<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^isub>a AndR <a>.M' <b>.N c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4544
| a_AndR_r: "\<lbrakk>a\<sharp>(N,c); b\<sharp>(M,c); b\<noteq>a; N\<longrightarrow>\<^isub>a N'\<rbrakk> \<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^isub>a AndR <a>.M <b>.N' c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4545
| a_AndL1: "\<lbrakk>x\<sharp>y; M\<longrightarrow>\<^isub>a M'\<rbrakk> \<Longrightarrow> AndL1 (x).M y \<longrightarrow>\<^isub>a AndL1 (x).M' y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4546
| a_AndL2: "\<lbrakk>x\<sharp>y; M\<longrightarrow>\<^isub>a M'\<rbrakk> \<Longrightarrow> AndL2 (x).M y \<longrightarrow>\<^isub>a AndL2 (x).M' y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4547
| a_OrL_l: "\<lbrakk>x\<sharp>(N,z); y\<sharp>(M,z); y\<noteq>x; M\<longrightarrow>\<^isub>a M'\<rbrakk> \<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^isub>a OrL (x).M' (y).N z"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4548
| a_OrL_r: "\<lbrakk>x\<sharp>(N,z); y\<sharp>(M,z); y\<noteq>x; N\<longrightarrow>\<^isub>a N'\<rbrakk> \<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^isub>a OrL (x).M (y).N' z"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4549
| a_OrR1: "\<lbrakk>a\<sharp>b; M\<longrightarrow>\<^isub>a M'\<rbrakk> \<Longrightarrow> OrR1 <a>.M b \<longrightarrow>\<^isub>a OrR1 <a>.M' b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4550
| a_OrR2: "\<lbrakk>a\<sharp>b; M\<longrightarrow>\<^isub>a M'\<rbrakk> \<Longrightarrow> OrR2 <a>.M b \<longrightarrow>\<^isub>a OrR2 <a>.M' b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4551
| a_ImpL_l: "\<lbrakk>a\<sharp>N; x\<sharp>(M,y); M\<longrightarrow>\<^isub>a M'\<rbrakk> \<Longrightarrow> ImpL <a>.M (x).N y \<longrightarrow>\<^isub>a ImpL <a>.M' (x).N y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4552
| a_ImpL_r: "\<lbrakk>a\<sharp>N; x\<sharp>(M,y); N\<longrightarrow>\<^isub>a N'\<rbrakk> \<Longrightarrow> ImpL <a>.M (x).N y \<longrightarrow>\<^isub>a ImpL <a>.M (x).N' y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4553
| a_ImpR: "\<lbrakk>a\<sharp>b; M\<longrightarrow>\<^isub>a M'\<rbrakk> \<Longrightarrow> ImpR (x).<a>.M b \<longrightarrow>\<^isub>a ImpR (x).<a>.M' b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4554
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4555
lemma fresh_a_redu:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4556
  fixes x::"name"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4557
  and   c::"coname"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4558
  shows "M \<longrightarrow>\<^isub>a M' \<Longrightarrow> x\<sharp>M \<Longrightarrow> x\<sharp>M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4559
  and   "M \<longrightarrow>\<^isub>a M' \<Longrightarrow> c\<sharp>M \<Longrightarrow> c\<sharp>M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4560
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4561
apply(induct rule: a_redu.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4562
apply(simp add: fresh_l_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4563
apply(simp add: fresh_c_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4564
apply(auto simp add: abs_fresh abs_supp fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4565
apply(induct rule: a_redu.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4566
apply(simp add: fresh_l_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4567
apply(simp add: fresh_c_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4568
apply(auto simp add: abs_fresh abs_supp fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4569
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4570
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4571
equivariance a_redu
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4572
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4573
nominal_inductive a_redu
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4574
  by (simp_all add: abs_fresh fresh_atm fresh_prod abs_supp fin_supp fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4575
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4576
lemma better_CutL_intro[intro]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4577
  shows "M\<longrightarrow>\<^isub>a M' \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>a Cut <a>.M' (x).N"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4578
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4579
  assume red: "M\<longrightarrow>\<^isub>a M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4580
  obtain x'::"name"   where fs1: "x'\<sharp>(M,N,x)" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4581
  obtain a'::"coname" where fs2: "a'\<sharp>(M,N,a)" by (rule exists_fresh(2), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4582
  have "Cut <a>.M (x).N =  Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4583
    using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4584
  also have "\<dots> \<longrightarrow>\<^isub>a  Cut <a'>.([(a',a)]\<bullet>M') (x').([(x',x)]\<bullet>N)" using fs1 fs2 red
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4585
    by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4586
  also have "\<dots> = Cut <a>.M' (x).N" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4587
    using fs1 fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4588
  finally show ?thesis by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4589
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4590
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4591
lemma better_CutR_intro[intro]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4592
  shows "N\<longrightarrow>\<^isub>a N' \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>a Cut <a>.M (x).N'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4593
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4594
  assume red: "N\<longrightarrow>\<^isub>a N'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4595
  obtain x'::"name"   where fs1: "x'\<sharp>(M,N,x)" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4596
  obtain a'::"coname" where fs2: "a'\<sharp>(M,N,a)" by (rule exists_fresh(2), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4597
  have "Cut <a>.M (x).N =  Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4598
    using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4599
  also have "\<dots> \<longrightarrow>\<^isub>a  Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N')" using fs1 fs2 red
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4600
    by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4601
  also have "\<dots> = Cut <a>.M (x).N'" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4602
    using fs1 fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4603
  finally show ?thesis by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4604
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4605
    
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4606
lemma better_AndRL_intro[intro]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4607
  shows "M\<longrightarrow>\<^isub>a M' \<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^isub>a AndR <a>.M' <b>.N c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4608
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4609
  assume red: "M\<longrightarrow>\<^isub>a M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4610
  obtain b'::"coname" where fs1: "b'\<sharp>(M,N,a,b,c)" by (rule exists_fresh(2), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4611
  obtain a'::"coname" where fs2: "a'\<sharp>(M,N,a,b,c,b')" by (rule exists_fresh(2), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4612
  have "AndR <a>.M <b>.N c =  AndR <a'>.([(a',a)]\<bullet>M) <b'>.([(b',b)]\<bullet>N) c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4613
    using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4614
  also have "\<dots> \<longrightarrow>\<^isub>a  AndR <a'>.([(a',a)]\<bullet>M') <b'>.([(b',b)]\<bullet>N) c" using fs1 fs2 red
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4615
    by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4616
  also have "\<dots> = AndR <a>.M' <b>.N c" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4617
    using fs1 fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4618
  finally show ?thesis by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4619
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4620
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4621
lemma better_AndRR_intro[intro]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4622
  shows "N\<longrightarrow>\<^isub>a N' \<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^isub>a AndR <a>.M <b>.N' c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4623
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4624
  assume red: "N\<longrightarrow>\<^isub>a N'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4625
  obtain b'::"coname" where fs1: "b'\<sharp>(M,N,a,b,c)" by (rule exists_fresh(2), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4626
  obtain a'::"coname" where fs2: "a'\<sharp>(M,N,a,b,c,b')" by (rule exists_fresh(2), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4627
  have "AndR <a>.M <b>.N c =  AndR <a'>.([(a',a)]\<bullet>M) <b'>.([(b',b)]\<bullet>N) c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4628
    using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4629
  also have "\<dots> \<longrightarrow>\<^isub>a  AndR <a'>.([(a',a)]\<bullet>M) <b'>.([(b',b)]\<bullet>N') c" using fs1 fs2 red
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4630
    by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4631
  also have "\<dots> = AndR <a>.M <b>.N' c" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4632
    using fs1 fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4633
  finally show ?thesis by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4634
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4635
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4636
lemma better_AndL1_intro[intro]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4637
  shows "M\<longrightarrow>\<^isub>a M' \<Longrightarrow> AndL1 (x).M y \<longrightarrow>\<^isub>a AndL1 (x).M' y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4638
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4639
  assume red: "M\<longrightarrow>\<^isub>a M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4640
  obtain x'::"name" where fs1: "x'\<sharp>(M,y,x)" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4641
  have "AndL1 (x).M y = AndL1 (x').([(x',x)]\<bullet>M) y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4642
    using fs1 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4643
  also have "\<dots> \<longrightarrow>\<^isub>a AndL1 (x').([(x',x)]\<bullet>M') y" using fs1 red
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4644
    by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4645
  also have "\<dots> = AndL1 (x).M' y" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4646
    using fs1 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4647
  finally show ?thesis by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4648
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4649
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4650
lemma better_AndL2_intro[intro]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4651
  shows "M\<longrightarrow>\<^isub>a M' \<Longrightarrow> AndL2 (x).M y \<longrightarrow>\<^isub>a AndL2 (x).M' y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4652
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4653
  assume red: "M\<longrightarrow>\<^isub>a M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4654
  obtain x'::"name" where fs1: "x'\<sharp>(M,y,x)" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4655
  have "AndL2 (x).M y = AndL2 (x').([(x',x)]\<bullet>M) y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4656
    using fs1 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4657
  also have "\<dots> \<longrightarrow>\<^isub>a AndL2 (x').([(x',x)]\<bullet>M') y" using fs1 red
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4658
    by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4659
  also have "\<dots> = AndL2 (x).M' y" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4660
    using fs1 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4661
  finally show ?thesis by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4662
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4663
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4664
lemma better_OrLL_intro[intro]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4665
  shows "M\<longrightarrow>\<^isub>a M' \<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^isub>a OrL (x).M' (y).N z"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4666
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4667
  assume red: "M\<longrightarrow>\<^isub>a M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4668
  obtain x'::"name" where fs1: "x'\<sharp>(M,N,x,y,z)" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4669
  obtain y'::"name" where fs2: "y'\<sharp>(M,N,x,y,z,x')" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4670
  have "OrL (x).M (y).N z =  OrL (x').([(x',x)]\<bullet>M) (y').([(y',y)]\<bullet>N) z"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4671
    using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4672
  also have "\<dots> \<longrightarrow>\<^isub>a OrL (x').([(x',x)]\<bullet>M') (y').([(y',y)]\<bullet>N) z" using fs1 fs2 red
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4673
    by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4674
  also have "\<dots> = OrL (x).M' (y).N z" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4675
    using fs1 fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4676
  finally show ?thesis by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4677
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4678
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4679
lemma better_OrLR_intro[intro]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4680
  shows "N\<longrightarrow>\<^isub>a N' \<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^isub>a OrL (x).M (y).N' z"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4681
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4682
  assume red: "N\<longrightarrow>\<^isub>a N'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4683
  obtain x'::"name" where fs1: "x'\<sharp>(M,N,x,y,z)" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4684
  obtain y'::"name" where fs2: "y'\<sharp>(M,N,x,y,z,x')" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4685
  have "OrL (x).M (y).N z =  OrL (x').([(x',x)]\<bullet>M) (y').([(y',y)]\<bullet>N) z"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4686
    using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4687
  also have "\<dots> \<longrightarrow>\<^isub>a OrL (x').([(x',x)]\<bullet>M) (y').([(y',y)]\<bullet>N') z" using fs1 fs2 red
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4688
    by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4689
  also have "\<dots> = OrL (x).M (y).N' z" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4690
    using fs1 fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4691
  finally show ?thesis by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4692
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4693
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4694
lemma better_OrR1_intro[intro]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4695
  shows "M\<longrightarrow>\<^isub>a M' \<Longrightarrow> OrR1 <a>.M b \<longrightarrow>\<^isub>a OrR1 <a>.M' b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4696
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4697
  assume red: "M\<longrightarrow>\<^isub>a M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4698
  obtain a'::"coname" where fs1: "a'\<sharp>(M,b,a)" by (rule exists_fresh(2), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4699
  have "OrR1 <a>.M b = OrR1 <a'>.([(a',a)]\<bullet>M) b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4700
    using fs1 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4701
  also have "\<dots> \<longrightarrow>\<^isub>a OrR1 <a'>.([(a',a)]\<bullet>M') b" using fs1 red
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4702
    by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4703
  also have "\<dots> = OrR1 <a>.M' b" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4704
    using fs1 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4705
  finally show ?thesis by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4706
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4707
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4708
lemma better_OrR2_intro[intro]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4709
  shows "M\<longrightarrow>\<^isub>a M' \<Longrightarrow> OrR2 <a>.M b \<longrightarrow>\<^isub>a OrR2 <a>.M' b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4710
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4711
  assume red: "M\<longrightarrow>\<^isub>a M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4712
  obtain a'::"coname" where fs1: "a'\<sharp>(M,b,a)" by (rule exists_fresh(2), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4713
  have "OrR2 <a>.M b = OrR2 <a'>.([(a',a)]\<bullet>M) b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4714
    using fs1 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4715
  also have "\<dots> \<longrightarrow>\<^isub>a OrR2 <a'>.([(a',a)]\<bullet>M') b" using fs1 red
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4716
    by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4717
  also have "\<dots> = OrR2 <a>.M' b" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4718
    using fs1 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4719
  finally show ?thesis by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4720
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4721
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4722
lemma better_ImpLL_intro[intro]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4723
  shows "M\<longrightarrow>\<^isub>a M' \<Longrightarrow> ImpL <a>.M (x).N y \<longrightarrow>\<^isub>a ImpL <a>.M' (x).N y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4724
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4725
  assume red: "M\<longrightarrow>\<^isub>a M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4726
  obtain x'::"name"   where fs1: "x'\<sharp>(M,N,x,y)" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4727
  obtain a'::"coname" where fs2: "a'\<sharp>(M,N,a)" by (rule exists_fresh(2), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4728
  have "ImpL <a>.M (x).N y =  ImpL <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N) y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4729
    using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4730
  also have "\<dots> \<longrightarrow>\<^isub>a  ImpL <a'>.([(a',a)]\<bullet>M') (x').([(x',x)]\<bullet>N) y" using fs1 fs2 red
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4731
    by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4732
  also have "\<dots> = ImpL <a>.M' (x).N y" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4733
    using fs1 fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4734
  finally show ?thesis by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4735
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4736
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4737
lemma better_ImpLR_intro[intro]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4738
  shows "N\<longrightarrow>\<^isub>a N' \<Longrightarrow> ImpL <a>.M (x).N y \<longrightarrow>\<^isub>a ImpL <a>.M (x).N' y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4739
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4740
  assume red: "N\<longrightarrow>\<^isub>a N'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4741
  obtain x'::"name"   where fs1: "x'\<sharp>(M,N,x,y)" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4742
  obtain a'::"coname" where fs2: "a'\<sharp>(M,N,a)" by (rule exists_fresh(2), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4743
  have "ImpL <a>.M (x).N y =  ImpL <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N) y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4744
    using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4745
  also have "\<dots> \<longrightarrow>\<^isub>a  ImpL <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N') y" using fs1 fs2 red
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4746
    by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4747
  also have "\<dots> = ImpL <a>.M (x).N' y" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4748
    using fs1 fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4749
  finally show ?thesis by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4750
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4751
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4752
lemma better_ImpR_intro[intro]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4753
  shows "M\<longrightarrow>\<^isub>a M' \<Longrightarrow> ImpR (x).<a>.M b \<longrightarrow>\<^isub>a ImpR (x).<a>.M' b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4754
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4755
  assume red: "M\<longrightarrow>\<^isub>a M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4756
  obtain a'::"coname" where fs2: "a'\<sharp>(M,a,b)" by (rule exists_fresh(2), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4757
  have "ImpR (x).<a>.M b = ImpR (x).<a'>.([(a',a)]\<bullet>M) b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4758
    using fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4759
  also have "\<dots> \<longrightarrow>\<^isub>a ImpR (x).<a'>.([(a',a)]\<bullet>M') b" using fs2 red
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4760
    by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4761
  also have "\<dots> = ImpR (x).<a>.M' b" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4762
    using fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4763
  finally show ?thesis by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4764
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4765
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4766
text {* axioms do not reduce *}
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4767
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4768
lemma ax_do_not_l_reduce:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4769
  shows "Ax x a \<longrightarrow>\<^isub>l M \<Longrightarrow> False"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4770
by (erule_tac l_redu.cases) (simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4771
 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4772
lemma ax_do_not_c_reduce:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4773
  shows "Ax x a \<longrightarrow>\<^isub>c M \<Longrightarrow> False"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4774
by (erule_tac c_redu.cases) (simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4775
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4776
lemma ax_do_not_a_reduce:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4777
  shows "Ax x a \<longrightarrow>\<^isub>a M \<Longrightarrow> False"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4778
apply(erule_tac a_redu.cases) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4779
apply(auto simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4780
apply(drule ax_do_not_l_reduce)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4781
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4782
apply(drule ax_do_not_c_reduce)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4783
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4784
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4785
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4786
lemma a_redu_NotL_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4787
  assumes a: "NotL <a>.M x \<longrightarrow>\<^isub>a R"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4788
  shows "\<exists>M'. R = NotL <a>.M' x \<and> M\<longrightarrow>\<^isub>aM'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4789
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4790
apply(erule_tac a_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4791
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4792
apply(erule_tac c_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4793
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4794
apply(rotate_tac 1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4795
apply(erule_tac a_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4796
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4797
apply(erule_tac c_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4798
apply(auto simp add: alpha a_redu.eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4799
apply(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4800
apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4801
apply(simp add: perm_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4802
apply(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4803
apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4804
apply(simp add: perm_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4805
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4806
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4807
lemma a_redu_NotR_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4808
  assumes a: "NotR (x).M a \<longrightarrow>\<^isub>a R"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4809
  shows "\<exists>M'. R = NotR (x).M' a \<and> M\<longrightarrow>\<^isub>aM'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4810
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4811
apply(erule_tac a_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4812
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4813
apply(erule_tac c_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4814
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4815
apply(rotate_tac 1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4816
apply(erule_tac a_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4817
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4818
apply(erule_tac c_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4819
apply(auto simp add: alpha a_redu.eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4820
apply(rule_tac x="([(x,xa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4821
apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4822
apply(simp add: perm_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4823
apply(rule_tac x="([(x,xaa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4824
apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4825
apply(simp add: perm_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4826
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4827
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4828
lemma a_redu_AndR_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4829
  assumes a: "AndR <a>.M <b>.N c\<longrightarrow>\<^isub>a R"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4830
  shows "(\<exists>M'. R = AndR <a>.M' <b>.N c \<and> M\<longrightarrow>\<^isub>aM') \<or> (\<exists>N'. R = AndR <a>.M <b>.N' c \<and> N\<longrightarrow>\<^isub>aN')"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4831
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4832
apply(erule_tac a_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4833
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4834
apply(erule_tac c_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4835
apply(rotate_tac 6)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4836
apply(erule_tac a_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4837
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4838
apply(erule_tac c_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4839
apply(rule disjI1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4840
apply(auto simp add: alpha a_redu.eqvt)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4841
apply(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4842
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4843
apply(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4844
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4845
apply(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4846
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4847
apply(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4848
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4849
apply(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4850
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4851
apply(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4852
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4853
apply(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4854
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4855
apply(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4856
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4857
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4858
apply(auto simp add: alpha a_redu.eqvt)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4859
apply(rule_tac x="([(b,ba)]\<bullet>N')" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4860
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4861
apply(rule_tac x="([(b,baa)]\<bullet>N')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4862
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4863
apply(rule_tac x="([(b,ba)]\<bullet>N')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4864
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4865
apply(rule_tac x="([(b,baa)]\<bullet>N')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4866
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4867
apply(rule_tac x="([(b,ba)]\<bullet>N')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4868
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4869
apply(rule_tac x="([(b,baa)]\<bullet>N')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4870
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4871
apply(rule_tac x="([(b,ba)]\<bullet>N')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4872
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4873
apply(rule_tac x="([(b,baa)]\<bullet>N')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4874
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4875
apply(rotate_tac 6)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4876
apply(erule_tac a_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4877
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4878
apply(erule_tac c_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4879
apply(rule disjI1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4880
apply(auto simp add: alpha a_redu.eqvt)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4881
apply(rule_tac x="([(a,aa)]\<bullet>M')" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4882
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4883
apply(rule_tac x="([(a,aa)]\<bullet>M')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4884
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4885
apply(rule_tac x="([(a,aa)]\<bullet>M')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4886
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4887
apply(rule_tac x="([(a,aa)]\<bullet>M')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4888
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4889
apply(rule_tac x="([(a,aaa)]\<bullet>M')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4890
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4891
apply(rule_tac x="([(a,aaa)]\<bullet>M')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4892
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4893
apply(rule_tac x="([(a,aaa)]\<bullet>M')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4894
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4895
apply(rule_tac x="([(a,aaa)]\<bullet>M')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4896
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4897
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4898
apply(auto simp add: alpha a_redu.eqvt)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4899
apply(rule_tac x="([(b,ba)]\<bullet>N'a)" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4900
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4901
apply(rule_tac x="([(b,ba)]\<bullet>N'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4902
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4903
apply(rule_tac x="([(b,ba)]\<bullet>N'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4904
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4905
apply(rule_tac x="([(b,ba)]\<bullet>N'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4906
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4907
apply(rule_tac x="([(b,baa)]\<bullet>N'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4908
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4909
apply(rule_tac x="([(b,baa)]\<bullet>N'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4910
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4911
apply(rule_tac x="([(b,baa)]\<bullet>N'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4912
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4913
apply(rule_tac x="([(b,baa)]\<bullet>N'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4914
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4915
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4916
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4917
lemma a_redu_AndL1_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4918
  assumes a: "AndL1 (x).M y \<longrightarrow>\<^isub>a R"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4919
  shows "\<exists>M'. R = AndL1 (x).M' y \<and> M\<longrightarrow>\<^isub>aM'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4920
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4921
apply(erule_tac a_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4922
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4923
apply(erule_tac c_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4924
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4925
apply(rotate_tac 2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4926
apply(erule_tac a_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4927
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4928
apply(erule_tac c_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4929
apply(auto simp add: alpha a_redu.eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4930
apply(rule_tac x="([(x,xa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4931
apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4932
apply(simp add: perm_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4933
apply(rule_tac x="([(x,xaa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4934
apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4935
apply(simp add: perm_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4936
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4937
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4938
lemma a_redu_AndL2_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4939
  assumes a: "AndL2 (x).M y \<longrightarrow>\<^isub>a R"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4940
  shows "\<exists>M'. R = AndL2 (x).M' y \<and> M\<longrightarrow>\<^isub>aM'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4941
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4942
apply(erule_tac a_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4943
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4944
apply(erule_tac c_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4945
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4946
apply(rotate_tac 2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4947
apply(erule_tac a_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4948
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4949
apply(erule_tac c_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4950
apply(auto simp add: alpha a_redu.eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4951
apply(rule_tac x="([(x,xa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4952
apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4953
apply(simp add: perm_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4954
apply(rule_tac x="([(x,xaa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4955
apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4956
apply(simp add: perm_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4957
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4958
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4959
lemma a_redu_OrL_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4960
  assumes a: "OrL (x).M (y).N z\<longrightarrow>\<^isub>a R"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4961
  shows "(\<exists>M'. R = OrL (x).M' (y).N z \<and> M\<longrightarrow>\<^isub>aM') \<or> (\<exists>N'. R = OrL (x).M (y).N' z \<and> N\<longrightarrow>\<^isub>aN')"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4962
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4963
apply(erule_tac a_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4964
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4965
apply(erule_tac c_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4966
apply(rotate_tac 6)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4967
apply(erule_tac a_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4968
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4969
apply(erule_tac c_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4970
apply(rule disjI1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4971
apply(auto simp add: alpha a_redu.eqvt)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4972
apply(rule_tac x="([(x,xa)]\<bullet>M'a)" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4973
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4974
apply(rule_tac x="([(x,xa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4975
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4976
apply(rule_tac x="([(x,xa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4977
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4978
apply(rule_tac x="([(x,xa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4979
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4980
apply(rule_tac x="([(x,xaa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4981
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4982
apply(rule_tac x="([(x,xaa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4983
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4984
apply(rule_tac x="([(x,xaa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4985
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4986
apply(rule_tac x="([(x,xaa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4987
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4988
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4989
apply(auto simp add: alpha a_redu.eqvt)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4990
apply(rule_tac x="([(y,ya)]\<bullet>N')" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4991
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4992
apply(rule_tac x="([(y,yaa)]\<bullet>N')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4993
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4994
apply(rule_tac x="([(y,ya)]\<bullet>N')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4995
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4996
apply(rule_tac x="([(y,yaa)]\<bullet>N')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4997
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4998
apply(rule_tac x="([(y,ya)]\<bullet>N')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  4999
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5000
apply(rule_tac x="([(y,yaa)]\<bullet>N')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5001
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5002
apply(rule_tac x="([(y,ya)]\<bullet>N')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5003
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5004
apply(rule_tac x="([(y,yaa)]\<bullet>N')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5005
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5006
apply(rotate_tac 6)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5007
apply(erule_tac a_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5008
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5009
apply(erule_tac c_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5010
apply(rule disjI1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5011
apply(auto simp add: alpha a_redu.eqvt)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5012
apply(rule_tac x="([(x,xa)]\<bullet>M')" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5013
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5014
apply(rule_tac x="([(x,xa)]\<bullet>M')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5015
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5016
apply(rule_tac x="([(x,xa)]\<bullet>M')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5017
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5018
apply(rule_tac x="([(x,xa)]\<bullet>M')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5019
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5020
apply(rule_tac x="([(x,xaa)]\<bullet>M')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5021
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5022
apply(rule_tac x="([(x,xaa)]\<bullet>M')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5023
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5024
apply(rule_tac x="([(x,xaa)]\<bullet>M')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5025
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5026
apply(rule_tac x="([(x,xaa)]\<bullet>M')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5027
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5028
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5029
apply(auto simp add: alpha a_redu.eqvt)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5030
apply(rule_tac x="([(y,ya)]\<bullet>N'a)" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5031
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5032
apply(rule_tac x="([(y,ya)]\<bullet>N'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5033
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5034
apply(rule_tac x="([(y,ya)]\<bullet>N'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5035
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5036
apply(rule_tac x="([(y,ya)]\<bullet>N'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5037
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5038
apply(rule_tac x="([(y,yaa)]\<bullet>N'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5039
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5040
apply(rule_tac x="([(y,yaa)]\<bullet>N'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5041
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5042
apply(rule_tac x="([(y,yaa)]\<bullet>N'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5043
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5044
apply(rule_tac x="([(y,yaa)]\<bullet>N'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5045
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5046
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5047
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5048
lemma a_redu_OrR1_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5049
  assumes a: "OrR1 <a>.M b \<longrightarrow>\<^isub>a R"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5050
  shows "\<exists>M'. R = OrR1 <a>.M' b \<and> M\<longrightarrow>\<^isub>aM'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5051
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5052
apply(erule_tac a_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5053
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5054
apply(erule_tac c_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5055
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5056
apply(rotate_tac 2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5057
apply(erule_tac a_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5058
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5059
apply(erule_tac c_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5060
apply(auto simp add: alpha a_redu.eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5061
apply(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5062
apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5063
apply(simp add: perm_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5064
apply(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5065
apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5066
apply(simp add: perm_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5067
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5068
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5069
lemma a_redu_OrR2_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5070
  assumes a: "OrR2 <a>.M b \<longrightarrow>\<^isub>a R"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5071
  shows "\<exists>M'. R = OrR2 <a>.M' b \<and> M\<longrightarrow>\<^isub>aM'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5072
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5073
apply(erule_tac a_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5074
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5075
apply(erule_tac c_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5076
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5077
apply(rotate_tac 2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5078
apply(erule_tac a_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5079
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5080
apply(erule_tac c_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5081
apply(auto simp add: alpha a_redu.eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5082
apply(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5083
apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5084
apply(simp add: perm_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5085
apply(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5086
apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5087
apply(simp add: perm_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5088
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5089
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5090
lemma a_redu_ImpL_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5091
  assumes a: "ImpL <a>.M (y).N z\<longrightarrow>\<^isub>a R"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5092
  shows "(\<exists>M'. R = ImpL <a>.M' (y).N z \<and> M\<longrightarrow>\<^isub>aM') \<or> (\<exists>N'. R = ImpL <a>.M (y).N' z \<and> N\<longrightarrow>\<^isub>aN')"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5093
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5094
apply(erule_tac a_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5095
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5096
apply(erule_tac c_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5097
apply(rotate_tac 5)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5098
apply(erule_tac a_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5099
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5100
apply(erule_tac c_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5101
apply(rule disjI1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5102
apply(auto simp add: alpha a_redu.eqvt)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5103
apply(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5104
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5105
apply(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5106
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5107
apply(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5108
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5109
apply(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5110
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5111
apply(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5112
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5113
apply(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5114
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5115
apply(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5116
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5117
apply(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5118
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5119
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5120
apply(auto simp add: alpha a_redu.eqvt)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5121
apply(rule_tac x="([(y,xa)]\<bullet>N')" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5122
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5123
apply(rule_tac x="([(y,xa)]\<bullet>N')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5124
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5125
apply(rule_tac x="([(y,xa)]\<bullet>N')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5126
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5127
apply(rule_tac x="([(y,xa)]\<bullet>N')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5128
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5129
apply(rule_tac x="([(y,xa)]\<bullet>N')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5130
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5131
apply(rule_tac x="([(y,xa)]\<bullet>N')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5132
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5133
apply(rule_tac x="([(y,xa)]\<bullet>N')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5134
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5135
apply(rule_tac x="([(y,xa)]\<bullet>N')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5136
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5137
apply(rotate_tac 5)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5138
apply(erule_tac a_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5139
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5140
apply(erule_tac c_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5141
apply(rule disjI1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5142
apply(auto simp add: alpha a_redu.eqvt)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5143
apply(rule_tac x="([(a,aa)]\<bullet>M')" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5144
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5145
apply(rule_tac x="([(a,aa)]\<bullet>M')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5146
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5147
apply(rule_tac x="([(a,aa)]\<bullet>M')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5148
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5149
apply(rule_tac x="([(a,aa)]\<bullet>M')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5150
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5151
apply(rule_tac x="([(a,aaa)]\<bullet>M')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5152
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5153
apply(rule_tac x="([(a,aaa)]\<bullet>M')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5154
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5155
apply(rule_tac x="([(a,aaa)]\<bullet>M')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5156
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5157
apply(rule_tac x="([(a,aaa)]\<bullet>M')" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5158
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5159
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5160
apply(auto simp add: alpha a_redu.eqvt)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5161
apply(rule_tac x="([(y,xa)]\<bullet>N'a)" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5162
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5163
apply(rule_tac x="([(y,xa)]\<bullet>N'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5164
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5165
apply(rule_tac x="([(y,xa)]\<bullet>N'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5166
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5167
apply(rule_tac x="([(y,xa)]\<bullet>N'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5168
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5169
apply(rule_tac x="([(y,xa)]\<bullet>N'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5170
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5171
apply(rule_tac x="([(y,xa)]\<bullet>N'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5172
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5173
apply(rule_tac x="([(y,xa)]\<bullet>N'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5174
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5175
apply(rule_tac x="([(y,xa)]\<bullet>N'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5176
apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5177
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5178
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5179
lemma a_redu_ImpR_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5180
  assumes a: "ImpR (x).<a>.M b \<longrightarrow>\<^isub>a R"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5181
  shows "\<exists>M'. R = ImpR (x).<a>.M' b \<and> M\<longrightarrow>\<^isub>aM'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5182
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5183
apply(erule_tac a_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5184
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5185
apply(erule_tac c_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5186
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5187
apply(rotate_tac 2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5188
apply(erule_tac a_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5189
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5190
apply(erule_tac c_redu.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5191
apply(auto simp add: alpha a_redu.eqvt abs_perm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5192
apply(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5193
apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5194
apply(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5195
apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5196
apply(rule_tac x="([(a,aa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5197
apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5198
apply(rule_tac x="([(a,aaa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5199
apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5200
apply(rule_tac x="([(x,xa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5201
apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5202
apply(rule_tac x="([(x,xa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5203
apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5204
apply(rule_tac x="([(a,aa)]\<bullet>[(x,xa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5205
apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5206
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5207
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5208
apply(rule perm_compose)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5209
apply(simp add: calc_atm perm_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5210
apply(rule_tac x="([(a,aaa)]\<bullet>[(x,xa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5211
apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5212
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5213
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5214
apply(rule perm_compose)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5215
apply(simp add: calc_atm perm_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5216
apply(rule_tac x="([(x,xaa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5217
apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5218
apply(rule_tac x="([(x,xaa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5219
apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5220
apply(rule_tac x="([(a,aa)]\<bullet>[(x,xaa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5221
apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5222
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5223
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5224
apply(rule perm_compose)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5225
apply(simp add: calc_atm perm_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5226
apply(rule_tac x="([(a,aaa)]\<bullet>[(x,xaa)]\<bullet>M'a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5227
apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5228
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5229
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5230
apply(rule perm_compose)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5231
apply(simp add: calc_atm perm_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5232
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5233
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5234
text {* Transitive Closure*}
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5235
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5236
abbreviation
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5237
 a_star_redu :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<longrightarrow>\<^isub>a* _" [100,100] 100)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5238
where
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5239
  "M \<longrightarrow>\<^isub>a* M' \<equiv> (a_redu)^** M M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5240
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5241
lemma a_starI:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5242
  assumes a: "M \<longrightarrow>\<^isub>a M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5243
  shows "M \<longrightarrow>\<^isub>a* M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5244
using a by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5245
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5246
lemma a_star_refl:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5247
  shows "M \<longrightarrow>\<^isub>a* M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5248
  by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5249
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5250
lemma a_starE:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5251
  assumes a: "M \<longrightarrow>\<^isub>a* M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5252
  shows "M = M' \<or> (\<exists>N. M \<longrightarrow>\<^isub>a N \<and> N \<longrightarrow>\<^isub>a* M')"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5253
using a 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5254
by (induct) (auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5255
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5256
lemma a_star_refl:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5257
  shows "M \<longrightarrow>\<^isub>a* M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5258
  by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5259
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5260
lemma a_star_trans[trans]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5261
  assumes a1: "M1\<longrightarrow>\<^isub>a* M2"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5262
  and     a2: "M2\<longrightarrow>\<^isub>a* M3"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5263
  shows "M1 \<longrightarrow>\<^isub>a* M3"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5264
using a2 a1
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5265
by (induct) (auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5266
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5267
text {* congruence rules for \<longrightarrow>\<^isub>a* *}
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5268
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5269
lemma ax_do_not_a_star_reduce:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5270
  shows "Ax x a \<longrightarrow>\<^isub>a* M \<Longrightarrow> M = Ax x a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5271
apply(induct set: rtrancl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5272
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5273
apply(drule  ax_do_not_a_reduce)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5274
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5275
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5276
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5277
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5278
lemma a_star_CutL:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5279
    "M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>a* Cut <a>.M' (x).N"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5280
by (induct set: rtrancl) (blast intro: rtrancl.rtrancl_into_rtrancl)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5281
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5282
lemma a_star_CutR:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5283
    "N \<longrightarrow>\<^isub>a* N'\<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>a* Cut <a>.M (x).N'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5284
by (induct set: rtrancl) (blast intro: rtrancl.rtrancl_into_rtrancl)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5285
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5286
lemma a_star_Cut:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5287
    "\<lbrakk>M \<longrightarrow>\<^isub>a* M'; N \<longrightarrow>\<^isub>a* N'\<rbrakk> \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>a* Cut <a>.M' (x).N'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5288
by (blast intro!: a_star_CutL a_star_CutR intro: rtrancl_trans')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5289
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5290
lemma a_star_NotR:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5291
    "M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> NotR (x).M a \<longrightarrow>\<^isub>a* NotR (x).M' a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5292
by (induct set: rtrancl) (blast intro: rtrancl.rtrancl_into_rtrancl)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5293
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5294
lemma a_star_NotL:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5295
    "M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> NotL <a>.M x \<longrightarrow>\<^isub>a* NotL <a>.M' x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5296
by (induct set: rtrancl) (blast intro: rtrancl.rtrancl_into_rtrancl)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5297
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5298
lemma a_star_AndRL:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5299
    "M \<longrightarrow>\<^isub>a* M'\<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^isub>a* AndR <a>.M' <b>.N c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5300
by (induct set: rtrancl) (blast intro: rtrancl.rtrancl_into_rtrancl)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5301
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5302
lemma a_star_AndRR:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5303
    "N \<longrightarrow>\<^isub>a* N'\<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^isub>a* AndR <a>.M <b>.N' c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5304
by (induct set: rtrancl) (blast intro: rtrancl.rtrancl_into_rtrancl)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5305
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5306
lemma a_star_AndR:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5307
    "\<lbrakk>M \<longrightarrow>\<^isub>a* M'; N \<longrightarrow>\<^isub>a* N'\<rbrakk> \<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^isub>a* AndR <a>.M' <b>.N' c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5308
by (blast intro!: a_star_AndRL a_star_AndRR intro: rtrancl_trans')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5309
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5310
lemma a_star_AndL1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5311
    "M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> AndL1 (x).M y \<longrightarrow>\<^isub>a* AndL1 (x).M' y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5312
by (induct set: rtrancl) (blast intro: rtrancl.rtrancl_into_rtrancl)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5313
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5314
lemma a_star_AndL2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5315
    "M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> AndL2 (x).M y \<longrightarrow>\<^isub>a* AndL2 (x).M' y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5316
by (induct set: rtrancl) (blast intro: rtrancl.rtrancl_into_rtrancl)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5317
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5318
lemma a_star_OrLL:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5319
    "M \<longrightarrow>\<^isub>a* M'\<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^isub>a* OrL (x).M' (y).N z"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5320
by (induct set: rtrancl) (blast intro: rtrancl.rtrancl_into_rtrancl)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5321
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5322
lemma a_star_OrLR:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5323
    "N \<longrightarrow>\<^isub>a* N'\<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^isub>a* OrL (x).M (y).N' z"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5324
by (induct set: rtrancl) (blast intro: rtrancl.rtrancl_into_rtrancl)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5325
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5326
lemma a_star_OrL:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5327
    "\<lbrakk>M \<longrightarrow>\<^isub>a* M'; N \<longrightarrow>\<^isub>a* N'\<rbrakk> \<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^isub>a* OrL (x).M' (y).N' z"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5328
by (blast intro!: a_star_OrLL a_star_OrLR intro: rtrancl_trans')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5329
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5330
lemma a_star_OrR1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5331
    "M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> OrR1 <a>.M b \<longrightarrow>\<^isub>a* OrR1 <a>.M' b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5332
by (induct set: rtrancl) (blast intro: rtrancl.rtrancl_into_rtrancl)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5333
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5334
lemma a_star_OrR2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5335
    "M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> OrR2 <a>.M b \<longrightarrow>\<^isub>a* OrR2 <a>.M' b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5336
by (induct set: rtrancl) (blast intro: rtrancl.rtrancl_into_rtrancl)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5337
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5338
lemma a_star_ImpLL:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5339
    "M \<longrightarrow>\<^isub>a* M'\<Longrightarrow> ImpL <a>.M (y).N z \<longrightarrow>\<^isub>a* ImpL <a>.M' (y).N z"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5340
by (induct set: rtrancl) (blast intro: rtrancl.rtrancl_into_rtrancl)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5341
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5342
lemma a_star_ImpLR:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5343
    "N \<longrightarrow>\<^isub>a* N'\<Longrightarrow> ImpL <a>.M (y).N z \<longrightarrow>\<^isub>a* ImpL <a>.M (y).N' z"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5344
by (induct set: rtrancl) (blast intro: rtrancl.rtrancl_into_rtrancl)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5345
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5346
lemma a_star_ImpL:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5347
    "\<lbrakk>M \<longrightarrow>\<^isub>a* M'; N \<longrightarrow>\<^isub>a* N'\<rbrakk> \<Longrightarrow> ImpL <a>.M (y).N z \<longrightarrow>\<^isub>a* ImpL <a>.M' (y).N' z"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5348
by (blast intro!: a_star_ImpLL a_star_ImpLR intro: rtrancl_trans')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5349
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5350
lemma a_star_ImpR:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5351
    "M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> ImpR (x).<a>.M b \<longrightarrow>\<^isub>a* ImpR (x).<a>.M' b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5352
by (induct set: rtrancl) (blast intro: rtrancl.rtrancl_into_rtrancl)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5353
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5354
lemmas a_star_congs = a_star_Cut a_star_NotR a_star_NotL a_star_AndR a_star_AndL1 a_star_AndL2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5355
                      a_star_OrL a_star_OrR1 a_star_OrR2 a_star_ImpL a_star_ImpR
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5356
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5357
lemma a_star_redu_NotL_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5358
  assumes a: "NotL <a>.M x \<longrightarrow>\<^isub>a* R"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5359
  shows "\<exists>M'. R = NotL <a>.M' x \<and> M \<longrightarrow>\<^isub>a* M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5360
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5361
apply(induct set: rtrancl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5362
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5363
apply(drule a_redu_NotL_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5364
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5365
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5366
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5367
lemma a_star_redu_NotR_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5368
  assumes a: "NotR (x).M a \<longrightarrow>\<^isub>a* R"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5369
  shows "\<exists>M'. R = NotR (x).M' a \<and> M \<longrightarrow>\<^isub>a* M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5370
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5371
apply(induct set: rtrancl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5372
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5373
apply(drule a_redu_NotR_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5374
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5375
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5376
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5377
lemma a_star_redu_AndR_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5378
  assumes a: "AndR <a>.M <b>.N c\<longrightarrow>\<^isub>a* R"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5379
  shows "(\<exists>M' N'. R = AndR <a>.M' <b>.N' c \<and> M \<longrightarrow>\<^isub>a* M' \<and> N \<longrightarrow>\<^isub>a* N')"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5380
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5381
apply(induct set: rtrancl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5382
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5383
apply(drule a_redu_AndR_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5384
apply(auto simp add: alpha trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5385
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5386
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5387
lemma a_star_redu_AndL1_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5388
  assumes a: "AndL1 (x).M y \<longrightarrow>\<^isub>a* R"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5389
  shows "\<exists>M'. R = AndL1 (x).M' y \<and> M \<longrightarrow>\<^isub>a* M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5390
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5391
apply(induct set: rtrancl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5392
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5393
apply(drule a_redu_AndL1_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5394
apply(auto simp add: alpha trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5395
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5396
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5397
lemma a_star_redu_AndL2_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5398
  assumes a: "AndL2 (x).M y \<longrightarrow>\<^isub>a* R"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5399
  shows "\<exists>M'. R = AndL2 (x).M' y \<and> M \<longrightarrow>\<^isub>a* M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5400
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5401
apply(induct set: rtrancl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5402
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5403
apply(drule a_redu_AndL2_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5404
apply(auto simp add: alpha trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5405
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5406
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5407
lemma a_star_redu_OrL_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5408
  assumes a: "OrL (x).M (y).N z \<longrightarrow>\<^isub>a* R"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5409
  shows "(\<exists>M' N'. R = OrL (x).M' (y).N' z \<and> M \<longrightarrow>\<^isub>a* M' \<and> N \<longrightarrow>\<^isub>a* N')"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5410
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5411
apply(induct set: rtrancl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5412
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5413
apply(drule a_redu_OrL_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5414
apply(auto simp add: alpha trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5415
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5416
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5417
lemma a_star_redu_OrR1_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5418
  assumes a: "OrR1 <a>.M y \<longrightarrow>\<^isub>a* R"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5419
  shows "\<exists>M'. R = OrR1 <a>.M' y \<and> M \<longrightarrow>\<^isub>a* M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5420
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5421
apply(induct set: rtrancl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5422
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5423
apply(drule a_redu_OrR1_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5424
apply(auto simp add: alpha trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5425
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5426
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5427
lemma a_star_redu_OrR2_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5428
  assumes a: "OrR2 <a>.M y \<longrightarrow>\<^isub>a* R"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5429
  shows "\<exists>M'. R = OrR2 <a>.M' y \<and> M \<longrightarrow>\<^isub>a* M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5430
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5431
apply(induct set: rtrancl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5432
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5433
apply(drule a_redu_OrR2_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5434
apply(auto simp add: alpha trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5435
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5436
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5437
lemma a_star_redu_ImpR_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5438
  assumes a: "ImpR (x).<a>.M y \<longrightarrow>\<^isub>a* R"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5439
  shows "\<exists>M'. R = ImpR (x).<a>.M' y \<and> M \<longrightarrow>\<^isub>a* M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5440
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5441
apply(induct set: rtrancl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5442
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5443
apply(drule a_redu_ImpR_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5444
apply(auto simp add: alpha trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5445
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5446
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5447
lemma a_star_redu_ImpL_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5448
  assumes a: "ImpL <a>.M (y).N z \<longrightarrow>\<^isub>a* R"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5449
  shows "(\<exists>M' N'. R = ImpL <a>.M' (y).N' z \<and> M \<longrightarrow>\<^isub>a* M' \<and> N \<longrightarrow>\<^isub>a* N')"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5450
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5451
apply(induct set: rtrancl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5452
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5453
apply(drule a_redu_ImpL_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5454
apply(auto simp add: alpha trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5455
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5456
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5457
text {* Substitution *}
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5458
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5459
lemma subst_not_fin1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5460
  shows "\<not>fin(M{x:=<c>.P}) x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5461
apply(nominal_induct M avoiding: x c P rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5462
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5463
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5464
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5465
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5466
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5467
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5468
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{x:=<c>.P},P)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5469
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5470
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5471
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5472
apply(simp add: fresh_fun_simp_NotL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5473
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5474
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5475
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5476
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5477
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{x:=<c>.P},P,name1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5478
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5479
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5480
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5481
apply(simp add: fresh_fun_simp_AndL1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5482
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5483
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5484
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5485
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{x:=<c>.P},P,name1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5486
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5487
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5488
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5489
apply(simp add: fresh_fun_simp_AndL2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5490
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5491
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5492
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5493
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5494
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5495
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{x:=<c>.P},P,name1,trm2{x:=<c>.P},name2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5496
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5497
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5498
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5499
apply(simp add: fresh_fun_simp_OrL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5500
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5501
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5502
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5503
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5504
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{name2:=<c>.P},P,name1,trm2{name2:=<c>.P})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5505
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5506
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5507
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5508
apply(simp add: fresh_fun_simp_ImpL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5509
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5510
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5511
apply(erule fin.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5512
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5513
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5514
lemma subst_not_fin2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5515
  assumes a: "\<not>fin M y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5516
  shows "\<not>fin(M{c:=(x).P}) y" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5517
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5518
apply(nominal_induct M avoiding: x c P y rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5519
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5520
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5521
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5522
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5523
apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(trm{coname:=(x).P},P)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5524
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5525
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5526
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5527
apply(simp add: fresh_fun_simp_NotR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5528
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5529
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5530
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5531
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5532
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5533
apply(drule freshn_after_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5534
apply(simp add: fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5535
apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(trm1{coname3:=(x).P},P,coname1,trm2{coname3:=(x).P},coname2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5536
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5537
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5538
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5539
apply(simp add: fresh_fun_simp_AndR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5540
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5541
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5542
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5543
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5544
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5545
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5546
apply(drule freshn_after_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5547
apply(simp add: fin.intros abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5548
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5549
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5550
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5551
apply(drule freshn_after_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5552
apply(simp add: fin.intros abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5553
apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(trm{coname2:=(x).P},P,coname1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5554
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5555
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5556
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5557
apply(simp add: fresh_fun_simp_OrR1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5558
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5559
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5560
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5561
apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(trm{coname2:=(x).P},P,coname1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5562
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5563
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5564
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5565
apply(simp add: fresh_fun_simp_OrR2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5566
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5567
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5568
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5569
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5570
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5571
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5572
apply(drule freshn_after_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5573
apply(drule freshn_after_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5574
apply(simp add: fin.intros abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5575
apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(trm{coname2:=(x).P},P,coname1,coname2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5576
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5577
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5578
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5579
apply(simp add: fresh_fun_simp_ImpR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5580
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5581
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5582
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5583
apply(drule fin_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5584
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5585
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5586
apply(drule freshn_after_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5587
apply(drule freshn_after_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5588
apply(simp add: fin.intros abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5589
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5590
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5591
lemma subst_not_fic1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5592
  shows "\<not>fic (M{a:=(x).P}) a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5593
apply(nominal_induct M avoiding: a x P rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5594
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5595
apply(erule fic.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5596
apply(erule fic.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5597
apply(erule fic.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5598
apply(erule fic.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5599
apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(trm{coname:=(x).P},P)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5600
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5601
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5602
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5603
apply(simp add: fresh_fun_simp_NotR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5604
apply(erule fic.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5605
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5606
apply(erule fic.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5607
apply(erule fic.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5608
apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5609
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5610
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5611
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5612
apply(simp add: fresh_fun_simp_AndR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5613
apply(erule fic.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5614
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5615
apply(erule fic.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5616
apply(erule fic.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5617
apply(erule fic.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5618
apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(trm{coname2:=(x).P},P,coname1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5619
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5620
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5621
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5622
apply(simp add: fresh_fun_simp_OrR1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5623
apply(erule fic.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5624
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5625
apply(erule fic.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5626
apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(trm{coname2:=(x).P},P,coname1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5627
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5628
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5629
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5630
apply(simp add: fresh_fun_simp_OrR2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5631
apply(erule fic.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5632
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5633
apply(erule fic.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5634
apply(erule fic.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5635
apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(trm{coname2:=(x).P},P,coname1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5636
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5637
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5638
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5639
apply(simp add: fresh_fun_simp_ImpR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5640
apply(erule fic.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5641
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5642
apply(erule fic.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5643
apply(erule fic.cases, simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5644
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5645
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5646
lemma subst_not_fic2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5647
  assumes a: "\<not>fic M a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5648
  shows "\<not>fic(M{x:=<b>.P}) a" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5649
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5650
apply(nominal_induct M avoiding: x a P b rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5651
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5652
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5653
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5654
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5655
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5656
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5657
apply(drule freshc_after_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5658
apply(simp add: fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5659
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{x:=<b>.P},P)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5660
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5661
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5662
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5663
apply(simp add: fresh_fun_simp_NotL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5664
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5665
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5666
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5667
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5668
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5669
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5670
apply(drule freshc_after_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5671
apply(drule freshc_after_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5672
apply(simp add: fic.intros abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5673
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{x:=<b>.P},P,name1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5674
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5675
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5676
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5677
apply(simp add: fresh_fun_simp_AndL1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5678
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5679
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5680
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5681
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{x:=<b>.P},P,name1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5682
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5683
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5684
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5685
apply(simp add: fresh_fun_simp_AndL2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5686
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5687
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5688
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5689
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5690
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5691
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5692
apply(drule freshc_after_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5693
apply(simp add: fic.intros abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5694
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5695
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5696
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5697
apply(drule freshc_after_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5698
apply(simp add: fic.intros abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5699
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{x:=<b>.P},P,name1,trm2{x:=<b>.P},name2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5700
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5701
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5702
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5703
apply(simp add: fresh_fun_simp_OrL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5704
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5705
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5706
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5707
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5708
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5709
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5710
apply(drule freshc_after_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5711
apply(simp add: fic.intros abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5712
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{name2:=<b>.P},trm2{name2:=<b>.P},P,name1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5713
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5714
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5715
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5716
apply(simp add: fresh_fun_simp_ImpL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5717
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5718
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5719
apply(drule fic_elims, simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5720
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5721
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5722
text {* Reductions *}
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5723
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5724
lemma fin_l_reduce:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5725
  assumes  a: "fin M x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5726
  and      b: "M \<longrightarrow>\<^isub>l M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5727
  shows "fin M' x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5728
using b a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5729
apply(induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5730
apply(erule fin.cases)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5731
apply(simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5732
apply(rotate_tac 3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5733
apply(erule fin.cases)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5734
apply(simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5735
apply(erule fin.cases, simp_all add: trm.inject)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5736
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5737
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5738
lemma fin_c_reduce:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5739
  assumes  a: "fin M x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5740
  and      b: "M \<longrightarrow>\<^isub>c M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5741
  shows "fin M' x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5742
using b a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5743
apply(induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5744
apply(erule fin.cases, simp_all add: trm.inject)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5745
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5746
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5747
lemma fin_a_reduce:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5748
  assumes  a: "fin M x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5749
  and      b: "M \<longrightarrow>\<^isub>a M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5750
  shows "fin M' x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5751
using a b
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5752
apply(induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5753
apply(drule ax_do_not_a_reduce)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5754
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5755
apply(drule a_redu_NotL_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5756
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5757
apply(rule fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5758
apply(simp add: fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5759
apply(drule a_redu_AndL1_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5760
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5761
apply(rule fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5762
apply(force simp add: abs_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5763
apply(drule a_redu_AndL2_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5764
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5765
apply(rule fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5766
apply(force simp add: abs_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5767
apply(drule a_redu_OrL_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5768
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5769
apply(rule fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5770
apply(force simp add: abs_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5771
apply(force simp add: abs_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5772
apply(rule fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5773
apply(force simp add: abs_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5774
apply(force simp add: abs_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5775
apply(drule a_redu_ImpL_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5776
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5777
apply(rule fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5778
apply(force simp add: abs_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5779
apply(force simp add: abs_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5780
apply(rule fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5781
apply(force simp add: abs_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5782
apply(force simp add: abs_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5783
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5784
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5785
lemma fin_a_star_reduce:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5786
  assumes  a: "fin M x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5787
  and      b: "M \<longrightarrow>\<^isub>a* M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5788
  shows "fin M' x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5789
using b a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5790
apply(induct set: rtrancl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5791
apply(auto simp add: fin_a_reduce)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5792
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5793
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5794
lemma fic_l_reduce:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5795
  assumes  a: "fic M x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5796
  and      b: "M \<longrightarrow>\<^isub>l M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5797
  shows "fic M' x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5798
using b a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5799
apply(induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5800
apply(erule fic.cases)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5801
apply(simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5802
apply(rotate_tac 3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5803
apply(erule fic.cases)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5804
apply(simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5805
apply(erule fic.cases, simp_all add: trm.inject)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5806
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5807
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5808
lemma fic_c_reduce:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5809
  assumes a: "fic M x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5810
  and     b: "M \<longrightarrow>\<^isub>c M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5811
  shows "fic M' x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5812
using b a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5813
apply(induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5814
apply(erule fic.cases, simp_all add: trm.inject)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5815
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5816
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5817
lemma fic_a_reduce:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5818
  assumes a: "fic M x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5819
  and     b: "M \<longrightarrow>\<^isub>a M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5820
  shows "fic M' x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5821
using a b
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5822
apply(induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5823
apply(drule ax_do_not_a_reduce)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5824
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5825
apply(drule a_redu_NotR_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5826
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5827
apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5828
apply(simp add: fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5829
apply(drule a_redu_AndR_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5830
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5831
apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5832
apply(force simp add: abs_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5833
apply(force simp add: abs_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5834
apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5835
apply(force simp add: abs_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5836
apply(force simp add: abs_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5837
apply(drule a_redu_OrR1_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5838
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5839
apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5840
apply(force simp add: abs_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5841
apply(drule a_redu_OrR2_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5842
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5843
apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5844
apply(force simp add: abs_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5845
apply(drule a_redu_ImpR_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5846
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5847
apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5848
apply(force simp add: abs_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5849
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5850
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5851
lemma fic_a_star_reduce:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5852
  assumes  a: "fic M x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5853
  and      b: "M \<longrightarrow>\<^isub>a* M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5854
  shows "fic M' x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5855
using b a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5856
apply(induct set: rtrancl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5857
apply(auto simp add: fic_a_reduce)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5858
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5859
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5860
text {* substitution properties *}
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5861
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5862
lemma subst_with_ax1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5863
  shows "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5864
proof(nominal_induct M avoiding: x a y rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5865
  case (Ax z b x a y)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5866
  show "(Ax z b){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (Ax z b)[x\<turnstile>n>y]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5867
  proof (cases "z=x")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5868
    case True
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5869
    assume eq: "z=x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5870
    have "(Ax z b){x:=<a>.Ax y a} = Cut <a>.Ax y a (x).Ax x b" using eq by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5871
    also have "\<dots> \<longrightarrow>\<^isub>a* (Ax x b)[x\<turnstile>n>y]" by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5872
    finally show "Ax z b{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (Ax z b)[x\<turnstile>n>y]" using eq by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5873
  next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5874
    case False
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5875
    assume neq: "z\<noteq>x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5876
    then show "(Ax z b){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (Ax z b)[x\<turnstile>n>y]" using neq by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5877
  qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5878
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5879
  case (Cut b M z N x a y)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5880
  have fs: "b\<sharp>x" "b\<sharp>a" "b\<sharp>y" "b\<sharp>N" "z\<sharp>x" "z\<sharp>a" "z\<sharp>y" "z\<sharp>M" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5881
  have ih1: "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5882
  have ih2: "N{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* N[x\<turnstile>n>y]" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5883
  show "(Cut <b>.M (z).N){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (Cut <b>.M (z).N)[x\<turnstile>n>y]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5884
  proof (cases "M = Ax x b")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5885
    case True
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5886
    assume eq: "M = Ax x b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5887
    have "(Cut <b>.M (z).N){x:=<a>.Ax y a} = Cut <a>.Ax y a (z).(N{x:=<a>.Ax y a})" using fs eq by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5888
    also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.Ax y a (z).(N[x\<turnstile>n>y])" using ih2 a_star_congs by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5889
    also have "\<dots> = Cut <b>.(M[x\<turnstile>n>y]) (z).(N[x\<turnstile>n>y])" using eq
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5890
      by (simp add: trm.inject alpha calc_atm fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5891
    finally show "(Cut <b>.M (z).N){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (Cut <b>.M (z).N)[x\<turnstile>n>y]" using fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5892
  next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5893
    case False
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5894
    assume neq: "M \<noteq> Ax x b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5895
    have "(Cut <b>.M (z).N){x:=<a>.Ax y a} = Cut <b>.(M{x:=<a>.Ax y a}) (z).(N{x:=<a>.Ax y a})" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5896
      using fs neq by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5897
    also have "\<dots> \<longrightarrow>\<^isub>a* Cut <b>.(M[x\<turnstile>n>y]) (z).(N[x\<turnstile>n>y])" using ih1 ih2 a_star_congs by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5898
    finally show "(Cut <b>.M (z).N){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (Cut <b>.M (z).N)[x\<turnstile>n>y]" using fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5899
  qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5900
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5901
  case (NotR z M b x a y)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5902
  have fs: "z\<sharp>x" "z\<sharp>a" "z\<sharp>y" "z\<sharp>b" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5903
  have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5904
  have "(NotR (z).M b){x:=<a>.Ax y a} = NotR (z).(M{x:=<a>.Ax y a}) b" using fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5905
  also have "\<dots> \<longrightarrow>\<^isub>a* NotR (z).(M[x\<turnstile>n>y]) b" using ih by (auto intro: a_star_congs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5906
  finally show "(NotR (z).M b){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (NotR (z).M b)[x\<turnstile>n>y]" using fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5907
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5908
  case (NotL b M z x a y)  
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5909
  have fs: "b\<sharp>x" "b\<sharp>a" "b\<sharp>y" "b\<sharp>z" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5910
  have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5911
  show "(NotL <b>.M z){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (NotL <b>.M z)[x\<turnstile>n>y]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5912
  proof(cases "z=x")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5913
    case True
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5914
    assume eq: "z=x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5915
    obtain x'::"name" where new: "x'\<sharp>(Ax y a,M{x:=<a>.Ax y a})" by (rule exists_fresh(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5916
    have "(NotL <b>.M z){x:=<a>.Ax y a} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5917
                        fresh_fun (\<lambda>x'. Cut <a>.Ax y a (x').NotL <b>.(M{x:=<a>.Ax y a}) x')"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5918
      using eq fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5919
    also have "\<dots> = Cut <a>.Ax y a (x').NotL <b>.(M{x:=<a>.Ax y a}) x'" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5920
      using new by (simp add: fresh_fun_simp_NotL fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5921
    also have "\<dots> \<longrightarrow>\<^isub>a* (NotL <b>.(M{x:=<a>.Ax y a}) x')[x'\<turnstile>n>y]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5922
      using new 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5923
      apply(rule_tac a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5924
      apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5925
      apply(rule better_LAxL_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5926
      apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5927
      done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5928
    also have "\<dots> = NotL <b>.(M{x:=<a>.Ax y a}) y" using new by (simp add: nrename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5929
    also have "\<dots> \<longrightarrow>\<^isub>a* NotL <b>.(M[x\<turnstile>n>y]) y" using ih by (auto intro: a_star_congs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5930
    also have "\<dots> = (NotL <b>.M z)[x\<turnstile>n>y]" using eq by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5931
    finally show "(NotL <b>.M z){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (NotL <b>.M z)[x\<turnstile>n>y]" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5932
  next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5933
    case False
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5934
    assume neq: "z\<noteq>x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5935
    have "(NotL <b>.M z){x:=<a>.Ax y a} = NotL <b>.(M{x:=<a>.Ax y a}) z" using fs neq by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5936
    also have "\<dots> \<longrightarrow>\<^isub>a* NotL <b>.(M[x\<turnstile>n>y]) z" using ih by (auto intro: a_star_congs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5937
    finally show "(NotL <b>.M z){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (NotL <b>.M z)[x\<turnstile>n>y]" using neq by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5938
  qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5939
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5940
  case (AndR c M d N e x a y)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5941
  have fs: "c\<sharp>x" "c\<sharp>a" "c\<sharp>y" "d\<sharp>x" "d\<sharp>a" "d\<sharp>y" "d\<noteq>c" "c\<sharp>N" "c\<sharp>e" "d\<sharp>M" "d\<sharp>e" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5942
  have ih1: "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5943
  have ih2: "N{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* N[x\<turnstile>n>y]" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5944
  have "(AndR <c>.M <d>.N e){x:=<a>.Ax y a} = AndR <c>.(M{x:=<a>.Ax y a}) <d>.(N{x:=<a>.Ax y a}) e"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5945
    using fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5946
  also have "\<dots> \<longrightarrow>\<^isub>a* AndR <c>.(M[x\<turnstile>n>y]) <d>.(N[x\<turnstile>n>y]) e" using ih1 ih2 by (auto intro: a_star_congs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5947
  finally show "(AndR <c>.M <d>.N e){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (AndR <c>.M <d>.N e)[x\<turnstile>n>y]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5948
    using fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5949
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5950
  case (AndL1 u M v x a y)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5951
  have fs: "u\<sharp>x" "u\<sharp>a" "u\<sharp>y" "u\<sharp>v" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5952
  have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5953
  show "(AndL1 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (AndL1 (u).M v)[x\<turnstile>n>y]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5954
  proof(cases "v=x")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5955
    case True
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5956
    assume eq: "v=x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5957
    obtain v'::"name" where new: "v'\<sharp>(Ax y a,M{x:=<a>.Ax y a},u)" by (rule exists_fresh(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5958
    have "(AndL1 (u).M v){x:=<a>.Ax y a} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5959
                        fresh_fun (\<lambda>v'. Cut <a>.Ax y a (v').AndL1 (u).(M{x:=<a>.Ax y a}) v')"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5960
      using eq fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5961
    also have "\<dots> = Cut <a>.Ax y a (v').AndL1 (u).(M{x:=<a>.Ax y a}) v'" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5962
      using new by (simp add: fresh_fun_simp_AndL1 fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5963
    also have "\<dots> \<longrightarrow>\<^isub>a* (AndL1 (u).(M{x:=<a>.Ax y a}) v')[v'\<turnstile>n>y]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5964
      using new 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5965
      apply(rule_tac a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5966
      apply(rule a_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5967
      apply(rule better_LAxL_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5968
      apply(rule fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5969
      apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5970
      done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5971
    also have "\<dots> = AndL1 (u).(M{x:=<a>.Ax y a}) y" using fs new
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5972
      by (auto simp add: fresh_prod fresh_atm nrename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5973
    also have "\<dots> \<longrightarrow>\<^isub>a* AndL1 (u).(M[x\<turnstile>n>y]) y" using ih by (auto intro: a_star_congs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5974
    also have "\<dots> = (AndL1 (u).M v)[x\<turnstile>n>y]" using eq fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5975
    finally show "(AndL1 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (AndL1 (u).M v)[x\<turnstile>n>y]" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5976
  next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5977
    case False
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5978
    assume neq: "v\<noteq>x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5979
    have "(AndL1 (u).M v){x:=<a>.Ax y a} = AndL1 (u).(M{x:=<a>.Ax y a}) v" using fs neq by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5980
    also have "\<dots> \<longrightarrow>\<^isub>a* AndL1 (u).(M[x\<turnstile>n>y]) v" using ih by (auto intro: a_star_congs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5981
    finally show "(AndL1 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (AndL1 (u).M v)[x\<turnstile>n>y]" using fs neq by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5982
  qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5983
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5984
  case (AndL2 u M v x a y)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5985
  have fs: "u\<sharp>x" "u\<sharp>a" "u\<sharp>y" "u\<sharp>v" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5986
  have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5987
  show "(AndL2 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (AndL2 (u).M v)[x\<turnstile>n>y]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5988
  proof(cases "v=x")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5989
    case True
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5990
    assume eq: "v=x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5991
    obtain v'::"name" where new: "v'\<sharp>(Ax y a,M{x:=<a>.Ax y a},u)" by (rule exists_fresh(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5992
    have "(AndL2 (u).M v){x:=<a>.Ax y a} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5993
                        fresh_fun (\<lambda>v'. Cut <a>.Ax y a (v').AndL2 (u).(M{x:=<a>.Ax y a}) v')"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5994
      using eq fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5995
    also have "\<dots> = Cut <a>.Ax y a (v').AndL2 (u).(M{x:=<a>.Ax y a}) v'" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5996
      using new by (simp add: fresh_fun_simp_AndL2 fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5997
    also have "\<dots> \<longrightarrow>\<^isub>a* (AndL2 (u).(M{x:=<a>.Ax y a}) v')[v'\<turnstile>n>y]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5998
      using new 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  5999
      apply(rule_tac a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6000
      apply(rule a_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6001
      apply(rule better_LAxL_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6002
      apply(rule fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6003
      apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6004
      done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6005
    also have "\<dots> = AndL2 (u).(M{x:=<a>.Ax y a}) y" using fs new
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6006
      by (auto simp add: fresh_prod fresh_atm nrename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6007
    also have "\<dots> \<longrightarrow>\<^isub>a* AndL2 (u).(M[x\<turnstile>n>y]) y" using ih by (auto intro: a_star_congs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6008
    also have "\<dots> = (AndL2 (u).M v)[x\<turnstile>n>y]" using eq fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6009
    finally show "(AndL2 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (AndL2 (u).M v)[x\<turnstile>n>y]" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6010
  next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6011
    case False
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6012
    assume neq: "v\<noteq>x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6013
    have "(AndL2 (u).M v){x:=<a>.Ax y a} = AndL2 (u).(M{x:=<a>.Ax y a}) v" using fs neq by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6014
    also have "\<dots> \<longrightarrow>\<^isub>a* AndL2 (u).(M[x\<turnstile>n>y]) v" using ih by (auto intro: a_star_congs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6015
    finally show "(AndL2 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (AndL2 (u).M v)[x\<turnstile>n>y]" using fs neq by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6016
  qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6017
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6018
  case (OrR1 c M d  x a y)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6019
  have fs: "c\<sharp>x" "c\<sharp>a" "c\<sharp>y" "c\<sharp>d" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6020
  have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6021
  have "(OrR1 <c>.M d){x:=<a>.Ax y a} = OrR1 <c>.(M{x:=<a>.Ax y a}) d" using fs by (simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6022
  also have "\<dots> \<longrightarrow>\<^isub>a* OrR1 <c>.(M[x\<turnstile>n>y]) d" using ih by (auto intro: a_star_congs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6023
  finally show "(OrR1 <c>.M d){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (OrR1 <c>.M d)[x\<turnstile>n>y]" using fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6024
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6025
  case (OrR2 c M d  x a y)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6026
  have fs: "c\<sharp>x" "c\<sharp>a" "c\<sharp>y" "c\<sharp>d" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6027
  have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6028
  have "(OrR2 <c>.M d){x:=<a>.Ax y a} = OrR2 <c>.(M{x:=<a>.Ax y a}) d" using fs by (simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6029
  also have "\<dots> \<longrightarrow>\<^isub>a* OrR2 <c>.(M[x\<turnstile>n>y]) d" using ih by (auto intro: a_star_congs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6030
  finally show "(OrR2 <c>.M d){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (OrR2 <c>.M d)[x\<turnstile>n>y]" using fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6031
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6032
  case (OrL u M v N z x a y)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6033
  have fs: "u\<sharp>x" "u\<sharp>a" "u\<sharp>y" "v\<sharp>x" "v\<sharp>a" "v\<sharp>y" "v\<noteq>u" "u\<sharp>N" "u\<sharp>z" "v\<sharp>M" "v\<sharp>z" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6034
  have ih1: "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6035
  have ih2: "N{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* N[x\<turnstile>n>y]" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6036
  show "(OrL (u).M (v).N z){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (OrL (u).M (v).N z)[x\<turnstile>n>y]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6037
  proof(cases "z=x")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6038
    case True
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6039
    assume eq: "z=x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6040
    obtain z'::"name" where new: "z'\<sharp>(Ax y a,M{x:=<a>.Ax y a},N{x:=<a>.Ax y a},u,v,y,a)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6041
      by (rule exists_fresh(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6042
    have "(OrL (u).M (v).N z){x:=<a>.Ax y a} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6043
                 fresh_fun (\<lambda>z'. Cut <a>.Ax y a (z').OrL (u).(M{x:=<a>.Ax y a}) (v).(N{x:=<a>.Ax y a}) z')"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6044
      using eq fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6045
    also have "\<dots> = Cut <a>.Ax y a (z').OrL (u).(M{x:=<a>.Ax y a}) (v).(N{x:=<a>.Ax y a}) z'" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6046
      using new by (simp add: fresh_fun_simp_OrL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6047
    also have "\<dots> \<longrightarrow>\<^isub>a* (OrL (u).(M{x:=<a>.Ax y a}) (v).(N{x:=<a>.Ax y a}) z')[z'\<turnstile>n>y]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6048
      using new 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6049
      apply(rule_tac a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6050
      apply(rule a_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6051
      apply(rule better_LAxL_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6052
      apply(rule fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6053
      apply(simp_all add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6054
      done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6055
    also have "\<dots> = OrL (u).(M{x:=<a>.Ax y a}) (v).(N{x:=<a>.Ax y a}) y" using fs new
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6056
      by (auto simp add: fresh_prod fresh_atm nrename_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6057
    also have "\<dots> \<longrightarrow>\<^isub>a* OrL (u).(M[x\<turnstile>n>y]) (v).(N[x\<turnstile>n>y]) y" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6058
      using ih1 ih2 by (auto intro: a_star_congs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6059
    also have "\<dots> = (OrL (u).M (v).N z)[x\<turnstile>n>y]" using eq fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6060
    finally show "(OrL (u).M (v).N z){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (OrL (u).M (v).N z)[x\<turnstile>n>y]" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6061
  next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6062
    case False
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6063
    assume neq: "z\<noteq>x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6064
    have "(OrL (u).M (v).N z){x:=<a>.Ax y a} = OrL (u).(M{x:=<a>.Ax y a}) (v).(N{x:=<a>.Ax y a}) z" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6065
      using fs neq by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6066
    also have "\<dots> \<longrightarrow>\<^isub>a* OrL (u).(M[x\<turnstile>n>y]) (v).(N[x\<turnstile>n>y]) z" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6067
      using ih1 ih2 by (auto intro: a_star_congs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6068
    finally show "(OrL (u).M (v).N z){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (OrL (u).M (v).N z)[x\<turnstile>n>y]" using fs neq by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6069
  qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6070
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6071
  case (ImpR z c M d x a y)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6072
  have fs: "z\<sharp>x" "z\<sharp>a" "z\<sharp>y" "c\<sharp>x" "c\<sharp>a" "c\<sharp>y" "z\<sharp>d" "c\<sharp>d" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6073
  have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6074
  have "(ImpR (z).<c>.M d){x:=<a>.Ax y a} = ImpR (z).<c>.(M{x:=<a>.Ax y a}) d" using fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6075
  also have "\<dots> \<longrightarrow>\<^isub>a* ImpR (z).<c>.(M[x\<turnstile>n>y]) d" using ih by (auto intro: a_star_congs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6076
  finally show "(ImpR (z).<c>.M d){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (ImpR (z).<c>.M d)[x\<turnstile>n>y]" using fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6077
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6078
  case (ImpL c M u N v x a y)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6079
  have fs: "c\<sharp>x" "c\<sharp>a" "c\<sharp>y" "u\<sharp>x" "u\<sharp>a" "u\<sharp>y" "c\<sharp>N" "c\<sharp>v" "u\<sharp>M" "u\<sharp>v" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6080
  have ih1: "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6081
  have ih2: "N{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* N[x\<turnstile>n>y]" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6082
  show "(ImpL <c>.M (u).N v){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (ImpL <c>.M (u).N v)[x\<turnstile>n>y]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6083
  proof(cases "v=x")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6084
    case True
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6085
    assume eq: "v=x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6086
    obtain v'::"name" where new: "v'\<sharp>(Ax y a,M{x:=<a>.Ax y a},N{x:=<a>.Ax y a},y,a,u)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6087
      by (rule exists_fresh(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6088
    have "(ImpL <c>.M (u).N v){x:=<a>.Ax y a} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6089
                 fresh_fun (\<lambda>v'. Cut <a>.Ax y a (v').ImpL <c>.(M{x:=<a>.Ax y a}) (u).(N{x:=<a>.Ax y a}) v')"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6090
      using eq fs by simp 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6091
    also have "\<dots> = Cut <a>.Ax y a (v').ImpL <c>.(M{x:=<a>.Ax y a}) (u).(N{x:=<a>.Ax y a}) v'" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6092
      using new by (simp add: fresh_fun_simp_ImpL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6093
    also have "\<dots> \<longrightarrow>\<^isub>a* (ImpL <c>.(M{x:=<a>.Ax y a}) (u).(N{x:=<a>.Ax y a}) v')[v'\<turnstile>n>y]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6094
      using new 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6095
      apply(rule_tac a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6096
      apply(rule a_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6097
      apply(rule better_LAxL_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6098
      apply(rule fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6099
      apply(simp_all add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6100
      done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6101
    also have "\<dots> = ImpL <c>.(M{x:=<a>.Ax y a}) (u).(N{x:=<a>.Ax y a}) y" using fs new
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6102
      by (auto simp add: fresh_prod subst_fresh fresh_atm trm.inject alpha rename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6103
    also have "\<dots> \<longrightarrow>\<^isub>a* ImpL <c>.(M[x\<turnstile>n>y]) (u).(N[x\<turnstile>n>y]) y" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6104
      using ih1 ih2 by (auto intro: a_star_congs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6105
    also have "\<dots> = (ImpL <c>.M (u).N v)[x\<turnstile>n>y]" using eq fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6106
    finally show "(ImpL <c>.M (u).N v){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (ImpL <c>.M (u).N v)[x\<turnstile>n>y]" using fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6107
  next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6108
    case False
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6109
    assume neq: "v\<noteq>x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6110
    have "(ImpL <c>.M (u).N v){x:=<a>.Ax y a} = ImpL <c>.(M{x:=<a>.Ax y a}) (u).(N{x:=<a>.Ax y a}) v" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6111
      using fs neq by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6112
    also have "\<dots> \<longrightarrow>\<^isub>a* ImpL <c>.(M[x\<turnstile>n>y]) (u).(N[x\<turnstile>n>y]) v" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6113
      using ih1 ih2 by (auto intro: a_star_congs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6114
    finally show "(ImpL <c>.M (u).N v){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (ImpL <c>.M (u).N v)[x\<turnstile>n>y]" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6115
      using fs neq by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6116
  qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6117
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6118
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6119
lemma subst_with_ax2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6120
  shows "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6121
proof(nominal_induct M avoiding: b a x rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6122
  case (Ax z c b a x)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6123
  show "(Ax z c){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (Ax z c)[b\<turnstile>c>a]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6124
  proof (cases "c=b")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6125
    case True
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6126
    assume eq: "c=b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6127
    have "(Ax z c){b:=(x).Ax x a} = Cut <b>.Ax z c (x).Ax x a" using eq by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6128
    also have "\<dots> \<longrightarrow>\<^isub>a* (Ax z c)[b\<turnstile>c>a]" using eq by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6129
    finally show "(Ax z c){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (Ax z c)[b\<turnstile>c>a]" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6130
  next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6131
    case False
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6132
    assume neq: "c\<noteq>b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6133
    then show "(Ax z c){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (Ax z c)[b\<turnstile>c>a]" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6134
  qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6135
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6136
  case (Cut c M z N b a x)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6137
  have fs: "c\<sharp>b" "c\<sharp>a" "c\<sharp>x" "c\<sharp>N" "z\<sharp>b" "z\<sharp>a" "z\<sharp>x" "z\<sharp>M" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6138
  have ih1: "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6139
  have ih2: "N{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* N[b\<turnstile>c>a]" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6140
  show "(Cut <c>.M (z).N){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (Cut <c>.M (z).N)[b\<turnstile>c>a]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6141
  proof (cases "N = Ax z b")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6142
    case True
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6143
    assume eq: "N = Ax z b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6144
    have "(Cut <c>.M (z).N){b:=(x).Ax x a} = Cut <c>.(M{b:=(x).Ax x a}) (x).Ax x a" using eq fs by simp 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6145
    also have "\<dots> \<longrightarrow>\<^isub>a* Cut <c>.(M[b\<turnstile>c>a]) (x).Ax x a" using ih1 a_star_congs by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6146
    also have "\<dots> = Cut <c>.(M[b\<turnstile>c>a]) (z).(N[b\<turnstile>c>a])" using eq fs
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6147
      by (simp add: trm.inject alpha calc_atm fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6148
    finally show "(Cut <c>.M (z).N){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (Cut <c>.M (z).N)[b\<turnstile>c>a]" using fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6149
  next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6150
    case False
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6151
    assume neq: "N \<noteq> Ax z b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6152
    have "(Cut <c>.M (z).N){b:=(x).Ax x a} = Cut <c>.(M{b:=(x).Ax x a}) (z).(N{b:=(x).Ax x a})" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6153
      using fs neq by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6154
    also have "\<dots> \<longrightarrow>\<^isub>a* Cut <c>.(M[b\<turnstile>c>a]) (z).(N[b\<turnstile>c>a])" using ih1 ih2 a_star_congs by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6155
    finally show "(Cut <c>.M (z).N){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (Cut <c>.M (z).N)[b\<turnstile>c>a]" using fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6156
  qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6157
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6158
  case (NotR z M c b a x)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6159
  have fs: "z\<sharp>b" "z\<sharp>a" "z\<sharp>x" "z\<sharp>c" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6160
  have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6161
  show "(NotR (z).M c){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (NotR (z).M c)[b\<turnstile>c>a]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6162
  proof (cases "c=b")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6163
    case True
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6164
    assume eq: "c=b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6165
    obtain a'::"coname" where new: "a'\<sharp>(Ax x a,M{b:=(x).Ax x a})" by (rule exists_fresh(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6166
    have "(NotR (z).M c){b:=(x).Ax x a} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6167
                        fresh_fun (\<lambda>a'. Cut <a'>.NotR (z).M{b:=(x).Ax x a} a' (x).Ax x a)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6168
      using eq fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6169
    also have "\<dots> = Cut <a'>.NotR (z).M{b:=(x).Ax x a} a' (x).Ax x a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6170
      using new by (simp add: fresh_fun_simp_NotR fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6171
    also have "\<dots> \<longrightarrow>\<^isub>a* (NotR (z).(M{b:=(x).Ax x a}) a')[a'\<turnstile>c>a]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6172
      using new 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6173
      apply(rule_tac a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6174
      apply(rule a_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6175
      apply(rule better_LAxR_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6176
      apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6177
      apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6178
      done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6179
    also have "\<dots> = NotR (z).(M{b:=(x).Ax x a}) a" using new by (simp add: crename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6180
    also have "\<dots> \<longrightarrow>\<^isub>a* NotR (z).(M[b\<turnstile>c>a]) a" using ih by (auto intro: a_star_congs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6181
    also have "\<dots> = (NotR (z).M c)[b\<turnstile>c>a]" using eq by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6182
    finally show "(NotR (z).M c){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (NotR (z).M c)[b\<turnstile>c>a]" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6183
  next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6184
    case False
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6185
    assume neq: "c\<noteq>b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6186
    have "(NotR (z).M c){b:=(x).Ax x a} = NotR (z).(M{b:=(x).Ax x a}) c" using fs neq by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6187
    also have "\<dots> \<longrightarrow>\<^isub>a* NotR (z).(M[b\<turnstile>c>a]) c" using ih by (auto intro: a_star_congs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6188
    finally show "(NotR (z).M c){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (NotR (z).M c)[b\<turnstile>c>a]" using neq by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6189
  qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6190
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6191
  case (NotL c M z b a x)  
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6192
  have fs: "c\<sharp>b" "c\<sharp>a" "c\<sharp>x" "c\<sharp>z" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6193
  have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6194
  have "(NotL <c>.M z){b:=(x).Ax x a} = NotL <c>.(M{b:=(x).Ax x a}) z" using fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6195
  also have "\<dots> \<longrightarrow>\<^isub>a* NotL <c>.(M[b\<turnstile>c>a]) z" using ih by (auto intro: a_star_congs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6196
  finally show "(NotL <c>.M z){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (NotL <c>.M z)[b\<turnstile>c>a]" using fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6197
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6198
  case (AndR c M d N e b a x)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6199
  have fs: "c\<sharp>b" "c\<sharp>a" "c\<sharp>x" "d\<sharp>b" "d\<sharp>a" "d\<sharp>x" "d\<noteq>c" "c\<sharp>N" "c\<sharp>e" "d\<sharp>M" "d\<sharp>e" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6200
  have ih1: "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6201
  have ih2: "N{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* N[b\<turnstile>c>a]" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6202
  show "(AndR <c>.M <d>.N e){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (AndR <c>.M <d>.N e)[b\<turnstile>c>a]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6203
  proof(cases "e=b")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6204
    case True
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6205
    assume eq: "e=b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6206
    obtain e'::"coname" where new: "e'\<sharp>(Ax x a,M{b:=(x).Ax x a},N{b:=(x).Ax x a},c,d)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6207
      by (rule exists_fresh(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6208
    have "(AndR <c>.M <d>.N e){b:=(x).Ax x a} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6209
               fresh_fun (\<lambda>e'. Cut <e'>.AndR <c>.(M{b:=(x).Ax x a}) <d>.(N{b:=(x).Ax x a}) e' (x).Ax x a)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6210
      using eq fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6211
    also have "\<dots> = Cut <e'>.AndR <c>.(M{b:=(x).Ax x a}) <d>.(N{b:=(x).Ax x a}) e' (x).Ax x a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6212
      using new by (simp add: fresh_fun_simp_AndR fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6213
    also have "\<dots> \<longrightarrow>\<^isub>a* (AndR <c>.(M{b:=(x).Ax x a}) <d>.(N{b:=(x).Ax x a}) e')[e'\<turnstile>c>a]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6214
      using new 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6215
      apply(rule_tac a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6216
      apply(rule a_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6217
      apply(rule better_LAxR_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6218
      apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6219
      apply(simp_all add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6220
      done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6221
    also have "\<dots> = AndR <c>.(M{b:=(x).Ax x a}) <d>.(N{b:=(x).Ax x a}) a" using fs new
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6222
      by (auto simp add: fresh_prod fresh_atm subst_fresh crename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6223
    also have "\<dots> \<longrightarrow>\<^isub>a* AndR <c>.(M[b\<turnstile>c>a]) <d>.(N[b\<turnstile>c>a]) a" using ih1 ih2 by (auto intro: a_star_congs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6224
    also have "\<dots> = (AndR <c>.M <d>.N e)[b\<turnstile>c>a]" using eq fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6225
    finally show "(AndR <c>.M <d>.N e){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (AndR <c>.M <d>.N e)[b\<turnstile>c>a]" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6226
  next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6227
    case False
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6228
    assume neq: "e\<noteq>b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6229
    have "(AndR <c>.M <d>.N e){b:=(x).Ax x a} = AndR <c>.(M{b:=(x).Ax x a}) <d>.(N{b:=(x).Ax x a}) e"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6230
      using fs neq by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6231
    also have "\<dots> \<longrightarrow>\<^isub>a* AndR <c>.(M[b\<turnstile>c>a]) <d>.(N[b\<turnstile>c>a]) e" using ih1 ih2 by (auto intro: a_star_congs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6232
    finally show "(AndR <c>.M <d>.N e){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (AndR <c>.M <d>.N e)[b\<turnstile>c>a]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6233
      using fs neq by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6234
  qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6235
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6236
  case (AndL1 u M v b a x)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6237
  have fs: "u\<sharp>b" "u\<sharp>a" "u\<sharp>x" "u\<sharp>v" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6238
  have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6239
  have "(AndL1 (u).M v){b:=(x).Ax x a} = AndL1 (u).(M{b:=(x).Ax x a}) v" using fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6240
  also have "\<dots> \<longrightarrow>\<^isub>a* AndL1 (u).(M[b\<turnstile>c>a]) v" using ih by (auto intro: a_star_congs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6241
  finally show "(AndL1 (u).M v){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (AndL1 (u).M v)[b\<turnstile>c>a]" using fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6242
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6243
  case (AndL2 u M v b a x)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6244
  have fs: "u\<sharp>b" "u\<sharp>a" "u\<sharp>x" "u\<sharp>v" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6245
  have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6246
  have "(AndL2 (u).M v){b:=(x).Ax x a} = AndL2 (u).(M{b:=(x).Ax x a}) v" using fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6247
  also have "\<dots> \<longrightarrow>\<^isub>a* AndL2 (u).(M[b\<turnstile>c>a]) v" using ih by (auto intro: a_star_congs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6248
  finally show "(AndL2 (u).M v){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (AndL2 (u).M v)[b\<turnstile>c>a]" using fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6249
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6250
  case (OrR1 c M d  b a x)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6251
  have fs: "c\<sharp>b" "c\<sharp>a" "c\<sharp>x" "c\<sharp>d" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6252
  have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6253
  show "(OrR1 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (OrR1 <c>.M d)[b\<turnstile>c>a]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6254
  proof(cases "d=b")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6255
    case True
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6256
    assume eq: "d=b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6257
    obtain a'::"coname" where new: "a'\<sharp>(Ax x a,M{b:=(x).Ax x a},c,x,a)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6258
      by (rule exists_fresh(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6259
    have "(OrR1 <c>.M d){b:=(x).Ax x a} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6260
             fresh_fun (\<lambda>a'. Cut <a'>.OrR1 <c>.M{b:=(x).Ax x a} a' (x).Ax x a)" using fs eq by (simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6261
    also have "\<dots> = Cut <a'>.OrR1 <c>.M{b:=(x).Ax x a} a' (x).Ax x a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6262
      using new by (simp add: fresh_fun_simp_OrR1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6263
    also have "\<dots> \<longrightarrow>\<^isub>a* (OrR1 <c>.M{b:=(x).Ax x a} a')[a'\<turnstile>c>a]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6264
      using new 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6265
      apply(rule_tac a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6266
      apply(rule a_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6267
      apply(rule better_LAxR_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6268
      apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6269
      apply(simp_all add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6270
      done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6271
    also have "\<dots> = OrR1 <c>.M{b:=(x).Ax x a} a" using fs new
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6272
      by (auto simp add: fresh_prod fresh_atm crename_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6273
    also have "\<dots> \<longrightarrow>\<^isub>a* OrR1 <c>.(M[b\<turnstile>c>a]) a" using ih by (auto intro: a_star_congs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6274
    also have "\<dots> = (OrR1 <c>.M d)[b\<turnstile>c>a]" using eq fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6275
    finally show "(OrR1 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (OrR1 <c>.M d)[b\<turnstile>c>a]" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6276
  next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6277
    case False
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6278
    assume neq: "d\<noteq>b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6279
    have "(OrR1 <c>.M d){b:=(x).Ax x a} = OrR1 <c>.(M{b:=(x).Ax x a}) d" using fs neq by (simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6280
    also have "\<dots> \<longrightarrow>\<^isub>a* OrR1 <c>.(M[b\<turnstile>c>a]) d" using ih by (auto intro: a_star_congs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6281
    finally show "(OrR1 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (OrR1 <c>.M d)[b\<turnstile>c>a]" using fs neq by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6282
  qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6283
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6284
  case (OrR2 c M d  b a x)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6285
  have fs: "c\<sharp>b" "c\<sharp>a" "c\<sharp>x" "c\<sharp>d" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6286
  have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6287
  show "(OrR2 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (OrR2 <c>.M d)[b\<turnstile>c>a]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6288
  proof(cases "d=b")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6289
    case True
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6290
    assume eq: "d=b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6291
    obtain a'::"coname" where new: "a'\<sharp>(Ax x a,M{b:=(x).Ax x a},c,x,a)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6292
      by (rule exists_fresh(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6293
    have "(OrR2 <c>.M d){b:=(x).Ax x a} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6294
             fresh_fun (\<lambda>a'. Cut <a'>.OrR2 <c>.M{b:=(x).Ax x a} a' (x).Ax x a)" using fs eq by (simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6295
    also have "\<dots> = Cut <a'>.OrR2 <c>.M{b:=(x).Ax x a} a' (x).Ax x a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6296
      using new by (simp add: fresh_fun_simp_OrR2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6297
    also have "\<dots> \<longrightarrow>\<^isub>a* (OrR2 <c>.M{b:=(x).Ax x a} a')[a'\<turnstile>c>a]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6298
      using new 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6299
      apply(rule_tac a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6300
      apply(rule a_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6301
      apply(rule better_LAxR_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6302
      apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6303
      apply(simp_all add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6304
      done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6305
    also have "\<dots> = OrR2 <c>.M{b:=(x).Ax x a} a" using fs new
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6306
      by (auto simp add: fresh_prod fresh_atm crename_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6307
    also have "\<dots> \<longrightarrow>\<^isub>a* OrR2 <c>.(M[b\<turnstile>c>a]) a" using ih by (auto intro: a_star_congs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6308
    also have "\<dots> = (OrR2 <c>.M d)[b\<turnstile>c>a]" using eq fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6309
    finally show "(OrR2 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (OrR2 <c>.M d)[b\<turnstile>c>a]" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6310
  next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6311
    case False
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6312
    assume neq: "d\<noteq>b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6313
    have "(OrR2 <c>.M d){b:=(x).Ax x a} = OrR2 <c>.(M{b:=(x).Ax x a}) d" using fs neq by (simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6314
    also have "\<dots> \<longrightarrow>\<^isub>a* OrR2 <c>.(M[b\<turnstile>c>a]) d" using ih by (auto intro: a_star_congs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6315
    finally show "(OrR2 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (OrR2 <c>.M d)[b\<turnstile>c>a]" using fs neq by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6316
  qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6317
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6318
  case (OrL u M v N z b a x)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6319
  have fs: "u\<sharp>b" "u\<sharp>a" "u\<sharp>x" "v\<sharp>b" "v\<sharp>a" "v\<sharp>x" "v\<noteq>u" "u\<sharp>N" "u\<sharp>z" "v\<sharp>M" "v\<sharp>z" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6320
  have ih1: "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6321
  have ih2: "N{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* N[b\<turnstile>c>a]" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6322
  have "(OrL (u).M (v).N z){b:=(x).Ax x a} = OrL (u).(M{b:=(x).Ax x a}) (v).(N{b:=(x).Ax x a}) z" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6323
    using fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6324
  also have "\<dots> \<longrightarrow>\<^isub>a* OrL (u).(M[b\<turnstile>c>a]) (v).(N[b\<turnstile>c>a]) z" using ih1 ih2 by (auto intro: a_star_congs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6325
  finally show "(OrL (u).M (v).N z){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (OrL (u).M (v).N z)[b\<turnstile>c>a]" using fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6326
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6327
  case (ImpR z c M d b a x)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6328
  have fs: "z\<sharp>b" "z\<sharp>a" "z\<sharp>x" "c\<sharp>b" "c\<sharp>a" "c\<sharp>x" "z\<sharp>d" "c\<sharp>d" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6329
  have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6330
  show "(ImpR (z).<c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (ImpR (z).<c>.M d)[b\<turnstile>c>a]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6331
  proof(cases "b=d")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6332
    case True
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6333
    assume eq: "b=d"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6334
    obtain a'::"coname" where new: "a'\<sharp>(Ax x a,M{b:=(x).Ax x a},x,a,c)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6335
      by (rule exists_fresh(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6336
    have "(ImpR (z).<c>.M d){b:=(x).Ax x a} =
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6337
                fresh_fun (\<lambda>a'. Cut <a'>.ImpR z.<c>.M{b:=(x).Ax x a} a' (x).Ax x a)" using fs eq by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6338
    also have "\<dots> = Cut <a'>.ImpR z.<c>.M{b:=(x).Ax x a} a' (x).Ax x a" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6339
      using new by (simp add: fresh_fun_simp_ImpR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6340
    also have "\<dots> \<longrightarrow>\<^isub>a* (ImpR z.<c>.M{b:=(x).Ax x a} a')[a'\<turnstile>c>a]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6341
      using new 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6342
      apply(rule_tac a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6343
      apply(rule a_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6344
      apply(rule better_LAxR_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6345
      apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6346
      apply(simp_all add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6347
      done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6348
    also have "\<dots> = ImpR z.<c>.M{b:=(x).Ax x a} a" using fs new
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6349
      by (auto simp add: fresh_prod crename_fresh subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6350
    also have "\<dots> \<longrightarrow>\<^isub>a* ImpR z.<c>.(M[b\<turnstile>c>a]) a" using ih by (auto intro: a_star_congs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6351
    also have "\<dots> = (ImpR z.<c>.M b)[b\<turnstile>c>a]" using eq fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6352
    finally show "(ImpR (z).<c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (ImpR (z).<c>.M d)[b\<turnstile>c>a]" using eq by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6353
  next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6354
    case False
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6355
    assume neq: "b\<noteq>d"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6356
    have "(ImpR (z).<c>.M d){b:=(x).Ax x a} = ImpR (z).<c>.(M{b:=(x).Ax x a}) d" using fs neq by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6357
    also have "\<dots> \<longrightarrow>\<^isub>a* ImpR (z).<c>.(M[b\<turnstile>c>a]) d" using ih by (auto intro: a_star_congs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6358
    finally show "(ImpR (z).<c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (ImpR (z).<c>.M d)[b\<turnstile>c>a]" using neq fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6359
  qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6360
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6361
  case (ImpL c M u N v b a x)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6362
  have fs: "c\<sharp>b" "c\<sharp>a" "c\<sharp>x" "u\<sharp>b" "u\<sharp>a" "u\<sharp>x" "c\<sharp>N" "c\<sharp>v" "u\<sharp>M" "u\<sharp>v" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6363
  have ih1: "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6364
  have ih2: "N{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* N[b\<turnstile>c>a]" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6365
  have "(ImpL <c>.M (u).N v){b:=(x).Ax x a} = ImpL <c>.(M{b:=(x).Ax x a}) (u).(N{b:=(x).Ax x a}) v" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6366
    using fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6367
  also have "\<dots> \<longrightarrow>\<^isub>a* ImpL <c>.(M[b\<turnstile>c>a]) (u).(N[b\<turnstile>c>a]) v" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6368
    using ih1 ih2 by (auto intro: a_star_congs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6369
  finally show "(ImpL <c>.M (u).N v){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (ImpL <c>.M (u).N v)[b\<turnstile>c>a]" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6370
    using fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6371
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6372
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6373
text {* substitution lemmas *}
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6374
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6375
lemma not_Ax1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6376
  shows "\<not>(b\<sharp>M) \<Longrightarrow> M{b:=(y).Q} \<noteq> Ax x a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6377
apply(nominal_induct M avoiding: b y Q x a rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6378
apply(auto simp add: fresh_atm abs_fresh abs_supp fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6379
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname:=(y).Q},Q)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6380
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6381
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6382
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6383
apply(simp add: fresh_fun_simp_NotR abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6384
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6385
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname:=(y).Q},Q)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6386
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6387
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6388
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6389
apply(simp add: fresh_fun_simp_NotR abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6390
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6391
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm1{coname3:=(y).Q},Q,trm2{coname3:=(y).Q},coname1,coname2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6392
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6393
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6394
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6395
apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6396
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6397
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm1{coname3:=(y).Q},Q,trm2{coname3:=(y).Q},coname1,coname2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6398
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6399
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6400
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6401
apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6402
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6403
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm1{coname3:=(y).Q},Q,trm2{coname3:=(y).Q},coname1,coname2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6404
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6405
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6406
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6407
apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6408
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6409
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(y).Q},Q,coname1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6410
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6411
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6412
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6413
apply(simp add: fresh_fun_simp_OrR1 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6414
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6415
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(y).Q},Q,coname1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6416
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6417
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6418
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6419
apply(simp add: fresh_fun_simp_OrR1 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6420
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6421
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(y).Q},Q,coname1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6422
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6423
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6424
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6425
apply(simp add: fresh_fun_simp_OrR2 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6426
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6427
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(y).Q},Q,coname1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6428
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6429
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6430
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6431
apply(simp add: fresh_fun_simp_OrR2 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6432
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6433
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(y).Q},Q,coname1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6434
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6435
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6436
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6437
apply(simp add: fresh_fun_simp_ImpR abs_fresh abs_supp fin_supp fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6438
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6439
apply(subgoal_tac "\<exists>x'::coname. x'\<sharp>(trm{coname2:=(y).Q},Q,coname1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6440
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6441
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6442
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6443
apply(simp add: fresh_fun_simp_ImpR abs_fresh abs_supp fin_supp fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6444
apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6445
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6446
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6447
lemma not_Ax2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6448
  shows "\<not>(x\<sharp>M) \<Longrightarrow> M{x:=<b>.Q} \<noteq> Ax y a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6449
apply(nominal_induct M avoiding: b y Q x a rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6450
apply(auto simp add: fresh_atm abs_fresh abs_supp fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6451
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{x:=<b>.Q},Q)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6452
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6453
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6454
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6455
apply(simp add: fresh_fun_simp_NotL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6456
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6457
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{x:=<b>.Q},Q)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6458
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6459
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6460
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6461
apply(simp add: fresh_fun_simp_NotL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6462
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6463
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{x:=<b>.Q},Q,name1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6464
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6465
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6466
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6467
apply(simp add: fresh_fun_simp_AndL1 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6468
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6469
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{x:=<b>.Q},Q,name1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6470
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6471
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6472
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6473
apply(simp add: fresh_fun_simp_AndL1 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6474
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6475
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{x:=<b>.Q},Q,name1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6476
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6477
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6478
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6479
apply(simp add: fresh_fun_simp_AndL2 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6480
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6481
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm{x:=<b>.Q},Q,name1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6482
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6483
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6484
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6485
apply(simp add: fresh_fun_simp_AndL2 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6486
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6487
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{x:=<b>.Q},Q,trm2{x:=<b>.Q},name1,name2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6488
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6489
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6490
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6491
apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6492
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6493
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{x:=<b>.Q},Q,trm2{x:=<b>.Q},name1,name2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6494
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6495
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6496
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6497
apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6498
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6499
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{x:=<b>.Q},Q,trm2{x:=<b>.Q},name1,name2)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6500
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6501
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6502
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6503
apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6504
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6505
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{name2:=<b>.Q},Q,trm2{name2:=<b>.Q},name1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6506
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6507
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6508
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6509
apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6510
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6511
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{name2:=<b>.Q},Q,trm2{name2:=<b>.Q},name1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6512
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6513
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6514
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6515
apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6516
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6517
apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(trm1{name2:=<b>.Q},Q,trm2{name2:=<b>.Q},name1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6518
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6519
apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6520
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6521
apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6522
apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6523
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6524
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6525
lemma interesting_subst1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6526
  assumes a: "x\<noteq>y" "x\<sharp>P" "y\<sharp>P" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6527
  shows "N{y:=<c>.P}{x:=<c>.P} = N{x:=<c>.Ax y c}{y:=<c>.P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6528
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6529
proof(nominal_induct N avoiding: x y c P rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6530
  case Ax
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6531
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6532
    by (auto simp add: abs_fresh fresh_atm forget trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6533
next 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6534
  case (Cut d M u M' x' y' c P)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6535
  from prems show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6536
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6537
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6538
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6539
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6540
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6541
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6542
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6543
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6544
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6545
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6546
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6547
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6548
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6549
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6550
    apply(simp add: trm.inject alpha forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6551
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6552
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6553
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6554
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6555
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6556
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6557
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6558
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6559
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6560
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6561
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6562
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6563
    apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6564
    apply(simp add: trm.inject alpha forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6565
    apply(simp add: trm.inject alpha forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6566
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6567
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6568
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6569
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6570
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6571
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6572
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6573
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6574
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6575
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6576
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6577
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6578
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6579
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6580
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6581
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6582
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6583
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6584
    apply(simp add: trm.inject alpha forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6585
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6586
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6587
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6588
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6589
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6590
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6591
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6592
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6593
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6594
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6595
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6596
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6597
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6598
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6599
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6600
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6601
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6602
    apply(case_tac "y'\<sharp>M")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6603
    apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6604
    apply(simp add: not_Ax2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6605
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6606
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6607
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6608
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6609
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6610
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6611
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6612
    apply(case_tac "x'\<sharp>M")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6613
    apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6614
    apply(simp add: not_Ax2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6615
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6616
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6617
  case NotR
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6618
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6619
    by (auto simp add: abs_fresh fresh_atm forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6620
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6621
  case (NotL d M u)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6622
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6623
    apply (auto simp add: abs_fresh fresh_atm forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6624
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(P,M{y:=<c>.P},M{x:=<c>.Ax y c}{y:=<c>.P},y,x)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6625
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6626
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6627
    apply(simp add: fresh_fun_simp_NotL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6628
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6629
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6630
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6631
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6632
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6633
    apply(auto simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6634
    apply(simp add: trm.inject alpha forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6635
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6636
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(P,M{x:=<c>.Ax y c},M{x:=<c>.Ax y c}{y:=<c>.P},Ax y c,y,x)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6637
    apply(erule exE, simp only: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6638
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6639
    apply(simp only: fresh_fun_simp_NotL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6640
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6641
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6642
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6643
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6644
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6645
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6646
    apply(simp add: trm.inject alpha forget subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6647
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6648
    apply(rule substn.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6649
    apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6650
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6651
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6652
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6653
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6654
  case (AndR d1 M d2 M' d3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6655
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6656
    by (auto simp add: abs_fresh fresh_atm forget trm.inject subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6657
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6658
  case (AndL1 u M d)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6659
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6660
    apply(auto simp add: abs_fresh fresh_atm forget trm.inject subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6661
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(P,M{y:=<c>.P},M{x:=<c>.Ax y c}{y:=<c>.P},u,y,x)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6662
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6663
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6664
    apply(simp add: fresh_fun_simp_AndL1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6665
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6666
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6667
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6668
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6669
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6670
    apply(auto simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6671
    apply(simp add: trm.inject alpha forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6672
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6673
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(P,Ax y c,M{x:=<c>.Ax y c},M{x:=<c>.Ax y c}{y:=<c>.P},u,y,x)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6674
    apply(erule exE, simp only: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6675
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6676
    apply(simp only: fresh_fun_simp_AndL1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6677
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6678
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6679
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6680
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6681
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6682
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6683
    apply(auto simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6684
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6685
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6686
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6687
  case (AndL2 u M d)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6688
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6689
    apply(auto simp add: abs_fresh fresh_atm forget trm.inject subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6690
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(P,M{y:=<c>.P},M{x:=<c>.Ax y c}{y:=<c>.P},u,y,x)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6691
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6692
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6693
    apply(simp add: fresh_fun_simp_AndL2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6694
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6695
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6696
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6697
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6698
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6699
    apply(auto simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6700
    apply(simp add: trm.inject alpha forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6701
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6702
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(P,Ax y c,M{x:=<c>.Ax y c},M{x:=<c>.Ax y c}{y:=<c>.P},u,y,x)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6703
    apply(erule exE, simp only: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6704
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6705
    apply(simp only: fresh_fun_simp_AndL2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6706
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6707
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6708
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6709
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6710
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6711
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6712
    apply(auto simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6713
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6714
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6715
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6716
  case OrR1
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6717
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6718
    by (auto simp add: abs_fresh fresh_atm forget trm.inject subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6719
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6720
  case OrR2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6721
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6722
    by (auto simp add: abs_fresh fresh_atm forget trm.inject subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6723
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6724
  case (OrL x1 M x2 M' x3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6725
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6726
    apply(auto simp add: abs_fresh fresh_atm forget trm.inject subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6727
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(P,M{y:=<c>.P},M{x:=<c>.Ax y c}{y:=<c>.P},
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6728
                                        M'{y:=<c>.P},M'{x:=<c>.Ax y c}{y:=<c>.P},x1,x2,x3,y,x)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6729
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6730
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6731
    apply(simp add: fresh_fun_simp_OrL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6732
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6733
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6734
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6735
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6736
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6737
    apply(auto simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6738
    apply(simp add: trm.inject alpha forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6739
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6740
    apply(rule substn.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6741
    apply(simp add: abs_fresh subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6742
    apply(simp add: abs_fresh subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6743
    apply(force)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6744
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6745
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6746
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(P,Ax y c,M{x:=<c>.Ax y c},M{x:=<c>.Ax y c}{y:=<c>.P},
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6747
                                        M'{x:=<c>.Ax y c},M'{x:=<c>.Ax y c}{y:=<c>.P},x1,x2,x3,y,x)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6748
    apply(erule exE, simp only: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6749
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6750
    apply(simp only: fresh_fun_simp_OrL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6751
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6752
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6753
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6754
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6755
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6756
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6757
    apply(simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6758
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6759
    apply(rule substn.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6760
    apply(simp add: abs_fresh subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6761
    apply(simp add: abs_fresh subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6762
    apply(force)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6763
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6764
    apply(auto simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6765
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6766
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6767
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6768
  case ImpR
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6769
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6770
    by (auto simp add: abs_fresh fresh_atm forget trm.inject subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6771
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6772
  case (ImpL a M x1 M' x2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6773
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6774
    apply(auto simp add: abs_fresh fresh_atm forget trm.inject subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6775
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(P,M{x2:=<c>.P},M{x:=<c>.Ax x2 c}{x2:=<c>.P},
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6776
                                        M'{x2:=<c>.P},M'{x:=<c>.Ax x2 c}{x2:=<c>.P},x1,y,x)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6777
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6778
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6779
    apply(simp add: fresh_fun_simp_ImpL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6780
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6781
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6782
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6783
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6784
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6785
    apply(auto simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6786
    apply(simp add: trm.inject alpha forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6787
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6788
    apply(rule substn.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6789
    apply(simp add: abs_fresh subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6790
    apply(simp add: abs_fresh subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6791
    apply(force)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6792
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6793
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(P,Ax y c,M{x2:=<c>.Ax y c},M{x2:=<c>.Ax y c}{y:=<c>.P},
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6794
                                        M'{x2:=<c>.Ax y c},M'{x2:=<c>.Ax y c}{y:=<c>.P},x1,x2,x3,y,x)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6795
    apply(erule exE, simp only: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6796
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6797
    apply(simp only: fresh_fun_simp_ImpL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6798
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6799
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6800
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6801
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6802
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6803
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6804
    apply(simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6805
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6806
    apply(rule substn.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6807
    apply(simp add: abs_fresh subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6808
    apply(simp add: abs_fresh subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6809
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6810
    apply(auto simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6811
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6812
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6813
qed 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6814
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6815
lemma interesting_subst1':
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6816
  assumes a: "x\<noteq>y" "x\<sharp>P" "y\<sharp>P" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6817
  shows "N{y:=<c>.P}{x:=<c>.P} = N{x:=<a>.Ax y a}{y:=<c>.P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6818
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6819
  show ?thesis
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6820
  proof (cases "c=a")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6821
    case True then show ?thesis using a by (simp add: interesting_subst1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6822
  next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6823
    case False then show ?thesis using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6824
      apply - 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6825
      apply(subgoal_tac "N{x:=<a>.Ax y a} = N{x:=<c>.([(c,a)]\<bullet>Ax y a)}") 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6826
      apply(simp add: interesting_subst1 calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6827
      apply(rule subst_rename)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6828
      apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6829
      done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6830
  qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6831
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6832
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6833
lemma interesting_subst2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6834
  assumes a: "a\<noteq>b" "a\<sharp>P" "b\<sharp>P" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6835
  shows "N{a:=(y).P}{b:=(y).P} = N{b:=(y).Ax y a}{a:=(y).P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6836
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6837
proof(nominal_induct N avoiding: a b y P rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6838
  case Ax
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6839
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6840
    by (auto simp add: abs_fresh fresh_atm forget trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6841
next 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6842
  case (Cut d M u M' x' y' c P)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6843
  from prems show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6844
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6845
    apply(auto simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6846
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6847
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6848
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6849
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6850
    apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6851
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6852
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6853
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6854
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6855
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6856
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6857
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6858
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6859
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6860
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6861
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6862
    apply(simp add: abs_fresh) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6863
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6864
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6865
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6866
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6867
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6868
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6869
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6870
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6871
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6872
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6873
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6874
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6875
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6876
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6877
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6878
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6879
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6880
    apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6881
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6882
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6883
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6884
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6885
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6886
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6887
    apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6888
    apply(simp add: fresh_atm trm.inject alpha forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6889
    apply(case_tac "x'\<sharp>M'")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6890
    apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6891
    apply(simp add: not_Ax1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6892
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6893
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6894
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6895
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6896
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6897
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6898
    apply(case_tac "y'\<sharp>M'")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6899
    apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6900
    apply(simp add: not_Ax1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6901
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6902
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6903
  case NotL
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6904
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6905
    by (auto simp add: abs_fresh fresh_atm forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6906
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6907
  case (NotR u M d)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6908
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6909
    apply (auto simp add: abs_fresh fresh_atm forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6910
    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(b,P,M{d:=(y).P},M{b:=(y).Ax y d}{d:=(y).P},u,y)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6911
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6912
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6913
    apply(simp add: fresh_fun_simp_NotR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6914
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6915
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6916
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6917
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6918
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6919
    apply(auto simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6920
    apply(simp add: trm.inject alpha forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6921
    apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6922
    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(P,M{d:=(y).Ax y a},M{d:=(y).Ax y a}{a:=(y).P},Ax y a,y,d)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6923
    apply(erule exE, simp only: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6924
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6925
    apply(simp only: fresh_fun_simp_NotR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6926
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6927
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6928
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6929
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6930
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6931
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6932
    apply(simp add: trm.inject alpha forget subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6933
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6934
    apply(rule substc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6935
    apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6936
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6937
    apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6938
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6939
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6940
  case (AndR d1 M d2 M' d3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6941
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6942
    apply(auto simp add: abs_fresh fresh_atm forget trm.inject subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6943
    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(P,M{d3:=(y).P},M{b:=(y).Ax y d3}{d3:=(y).P},
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6944
                                        M'{d3:=(y).P},M'{b:=(y).Ax y d3}{d3:=(y).P},d1,d2,d3,b,y)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6945
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6946
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6947
    apply(simp add: fresh_fun_simp_AndR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6948
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6949
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6950
    apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6951
    apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6952
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6953
    apply(auto simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6954
    apply(simp add: trm.inject alpha forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6955
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6956
    apply(rule substc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6957
    apply(simp add: abs_fresh subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6958
    apply(simp add: abs_fresh subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6959
    apply(force)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6960
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6961
    apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6962
    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(P,Ax y a,M{d3:=(y).Ax y a},M{d3:=(y).Ax y a}{a:=(y).P},
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6963
                                        M'{d3:=(y).Ax y a},M'{d3:=(y).Ax y a}{a:=(y).P},d1,d2,d3,y,b)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6964
    apply(erule exE, simp only: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6965
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6966
    apply(simp only: fresh_fun_simp_AndR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6967
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6968
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6969
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6970
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6971
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6972
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6973
    apply(simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6974
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6975
    apply(rule substc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6976
    apply(simp add: abs_fresh subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6977
    apply(simp add: abs_fresh subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6978
    apply(force)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6979
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6980
    apply(auto simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6981
    apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6982
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6983
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6984
  case (AndL1 u M d)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6985
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6986
    by (auto simp add: abs_fresh fresh_atm forget trm.inject subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6987
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6988
  case (AndL2 u M d)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6989
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6990
    by (auto simp add: abs_fresh fresh_atm forget trm.inject subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6991
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6992
  case (OrR1 d M e)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6993
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6994
    apply (auto simp add: abs_fresh fresh_atm forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6995
    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(b,P,M{e:=(y).P},M{b:=(y).Ax y e}{e:=(y).P},d,e)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6996
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6997
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6998
    apply(simp add: fresh_fun_simp_OrR1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  6999
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7000
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7001
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7002
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7003
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7004
    apply(auto simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7005
    apply(simp add: trm.inject alpha forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7006
    apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7007
    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(b,P,Ax y a,M{e:=(y).Ax y a},M{e:=(y).Ax y a}{a:=(y).P},d,e)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7008
    apply(erule exE, simp only: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7009
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7010
    apply(simp only: fresh_fun_simp_OrR1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7011
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7012
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7013
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7014
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7015
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7016
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7017
    apply(simp add: trm.inject alpha forget subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7018
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7019
    apply(rule substc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7020
    apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7021
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7022
    apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7023
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7024
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7025
  case (OrR2 d M e)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7026
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7027
    apply (auto simp add: abs_fresh fresh_atm forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7028
    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(b,P,M{e:=(y).P},M{b:=(y).Ax y e}{e:=(y).P},d,e)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7029
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7030
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7031
    apply(simp add: fresh_fun_simp_OrR2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7032
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7033
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7034
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7035
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7036
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7037
    apply(auto simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7038
    apply(simp add: trm.inject alpha forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7039
    apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7040
    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(b,P,Ax y a,M{e:=(y).Ax y a},M{e:=(y).Ax y a}{a:=(y).P},d,e)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7041
    apply(erule exE, simp only: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7042
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7043
    apply(simp only: fresh_fun_simp_OrR2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7044
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7045
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7046
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7047
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7048
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7049
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7050
    apply(simp add: trm.inject alpha forget subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7051
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7052
    apply(rule substc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7053
    apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7054
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7055
    apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7056
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7057
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7058
  case (OrL x1 M x2 M' x3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7059
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7060
    by(auto simp add: abs_fresh fresh_atm forget trm.inject subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7061
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7062
  case ImpL
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7063
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7064
    by (auto simp add: abs_fresh fresh_atm forget trm.inject subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7065
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7066
  case (ImpR u e M d)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7067
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7068
    apply(auto simp add: abs_fresh fresh_atm forget trm.inject subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7069
    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(b,e,d,P,M{d:=(y).P},M{b:=(y).Ax y d}{d:=(y).P})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7070
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7071
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7072
    apply(simp add: fresh_fun_simp_ImpR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7073
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7074
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7075
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7076
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7077
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7078
    apply(auto simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7079
    apply(simp add: trm.inject alpha forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7080
    apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7081
    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(e,d,P,Ax y a,M{d:=(y).Ax y a},M{d:=(y).Ax y a}{a:=(y).P})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7082
    apply(erule exE, simp only: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7083
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7084
    apply(simp only: fresh_fun_simp_ImpR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7085
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7086
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7087
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7088
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7089
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7090
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7091
    apply(simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7092
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7093
    apply(rule substc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7094
    apply(simp add: abs_fresh subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7095
    apply(simp add: abs_fresh subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7096
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7097
    apply(auto simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7098
    apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7099
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7100
qed 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7101
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7102
lemma interesting_subst2':
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7103
  assumes a: "a\<noteq>b" "a\<sharp>P" "b\<sharp>P" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7104
  shows "N{a:=(y).P}{b:=(y).P} = N{b:=(z).Ax z a}{a:=(y).P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7105
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7106
  show ?thesis
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7107
  proof (cases "z=y")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7108
    case True then show ?thesis using a by (simp add: interesting_subst2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7109
  next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7110
    case False then show ?thesis using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7111
      apply - 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7112
      apply(subgoal_tac "N{b:=(z).Ax z a} = N{b:=(y).([(y,z)]\<bullet>Ax z a)}") 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7113
      apply(simp add: interesting_subst2 calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7114
      apply(rule subst_rename)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7115
      apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7116
      done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7117
  qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7118
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7119
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7120
lemma subst_subst1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7121
  assumes a: "a\<sharp>(Q,b)" "x\<sharp>(y,P,Q)" "b\<sharp>Q" "y\<sharp>P" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7122
  shows "M{x:=<a>.P}{b:=(y).Q} = M{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7123
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7124
proof(nominal_induct M avoiding: x a P b y Q rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7125
  case (Ax z c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7126
  have fs: "a\<sharp>(Q,b)" "x\<sharp>(y,P,Q)" "b\<sharp>Q" "y\<sharp>P" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7127
  { assume asm: "z=x \<and> c=b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7128
    have "(Ax x b){x:=<a>.P}{b:=(y).Q} = (Cut <a>.P (x).Ax x b){b:=(y).Q}" using fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7129
    also have "\<dots> = Cut <a>.(P{b:=(y).Q}) (y).Q"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7130
      using fs by (simp_all add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7131
    also have "\<dots> = Cut <a>.(P{b:=(y).Q}) (y).(Q{x:=<a>.(P{b:=(y).Q})})" using fs by (simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7132
    also have "\<dots> = (Cut <b>.Ax x b (y).Q){x:=<a>.(P{b:=(y).Q})}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7133
      using fs asm by (auto simp add: fresh_prod fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7134
    also have "\<dots> = (Ax x b){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}" using fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7135
    finally have "(Ax z c){x:=<a>.P}{b:=(y).Q} = (Ax z c){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7136
      using asm by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7137
  }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7138
  moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7139
  { assume asm: "z\<noteq>x \<and> c=b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7140
    have "(Ax z c){x:=<a>.P}{b:=(y).Q} = (Ax z c){b:=(y).Q}" using asm by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7141
    also have "\<dots> = Cut <b>.Ax z c (y).Q" using fs asm by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7142
    also have "\<dots> = Cut <b>.(Ax z c{x:=<a>.(P{b:=(y).Q})}) (y).(Q{x:=<a>.(P{b:=(y).Q})})" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7143
      using fs asm by (simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7144
    also have "\<dots> = (Cut <b>.Ax z c (y).Q){x:=<a>.(P{b:=(y).Q})}" using asm fs
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7145
      by (auto simp add: trm.inject subst_fresh fresh_prod fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7146
    also have "\<dots> = (Ax z c){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}" using asm fs by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7147
    finally have "(Ax z c){x:=<a>.P}{b:=(y).Q} = (Ax z c){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7148
  }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7149
  moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7150
  { assume asm: "z=x \<and> c\<noteq>b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7151
    have "(Ax z c){x:=<a>.P}{b:=(y).Q} = (Cut <a>.P (x).Ax z c){b:=(y).Q}" using fs asm by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7152
    also have "\<dots> = Cut <a>.(P{b:=(y).Q}) (x).Ax z c" using fs asm by (auto simp add: trm.inject abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7153
    also have "\<dots> = (Ax z c){x:=<a>.(P{b:=(y).Q})}" using fs asm by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7154
    also have "\<dots> = (Ax z c){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}" using asm by auto
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7155
    finally have "(Ax z c){x:=<a>.P}{b:=(y).Q} = (Ax z c){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7156
  }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7157
  moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7158
  { assume asm: "z\<noteq>x \<and> c\<noteq>b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7159
    have "(Ax z c){x:=<a>.P}{b:=(y).Q} = (Ax z c){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}" using asm by auto
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7160
  }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7161
  ultimately show ?case by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7162
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7163
  case (Cut c M z N)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7164
  { assume asm: "M = Ax x c \<and> N = Ax z b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7165
    have "(Cut <c>.M (z).N){x:=<a>.P}{b:=(y).Q} = (Cut <a>.P (z).(N{x:=<a>.P})){b:=(y).Q}" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7166
      using asm prems by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7167
    also have "\<dots> = (Cut <a>.P (z).N){b:=(y).Q}" using asm prems by (simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7168
    also have "\<dots> = (Cut <a>.(P{b:=(y).Q}) (y).Q)" using asm prems by (auto simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7169
    finally have eq1: "(Cut <c>.M (z).N){x:=<a>.P}{b:=(y).Q} = (Cut <a>.(P{b:=(y).Q}) (y).Q)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7170
    have "(Cut <c>.M (z).N){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})} = (Cut <c>.M (y).Q){x:=<a>.(P{b:=(y).Q})}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7171
      using prems asm by (simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7172
    also have "\<dots> = Cut <a>.(P{b:=(y).Q}) (y).(Q{x:=<a>.(P{b:=(y).Q})})" using asm prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7173
      by (auto simp add: fresh_prod fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7174
    also have "\<dots> = Cut <a>.(P{b:=(y).Q}) (y).Q" using asm prems by (simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7175
    finally have eq2: "(Cut <c>.M (z).N){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})} = Cut <a>.(P{b:=(y).Q}) (y).Q"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7176
      by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7177
    have "(Cut <c>.M (z).N){x:=<a>.P}{b:=(y).Q} = (Cut <c>.M (z).N){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7178
      using eq1 eq2 by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7179
  }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7180
  moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7181
  { assume asm: "M \<noteq> Ax x c \<and> N = Ax z b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7182
    have neq: "M{b:=(y).Q} \<noteq> Ax x c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7183
    proof (cases "b\<sharp>M")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7184
      case True then show ?thesis using asm by (simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7185
    next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7186
      case False then show ?thesis by (simp add: not_Ax1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7187
    qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7188
    have "(Cut <c>.M (z).N){x:=<a>.P}{b:=(y).Q} = (Cut <c>.(M{x:=<a>.P}) (z).(N{x:=<a>.P})){b:=(y).Q}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7189
      using asm prems by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7190
    also have "\<dots> = (Cut <c>.(M{x:=<a>.P}) (z).N){b:=(y).Q}" using asm prems by (simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7191
    also have "\<dots> = Cut <c>.(M{x:=<a>.P}{b:=(y).Q}) (y).Q" using asm prems by (simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7192
    also have "\<dots> = Cut <c>.(M{b:=(y).Q}{x:=<a>.P{b:=(y).Q}}) (y).Q" using asm prems by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7193
    finally 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7194
    have eq1: "(Cut <c>.M (z).N){x:=<a>.P}{b:=(y).Q} = Cut <c>.(M{b:=(y).Q}{x:=<a>.P{b:=(y).Q}}) (y).Q" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7195
      by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7196
    have "(Cut <c>.M (z).N){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7197
               (Cut <c>.(M{b:=(y).Q}) (y).Q){x:=<a>.(P{b:=(y).Q})}" using asm prems by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7198
    also have "\<dots> = Cut <c>.(M{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}) (y).(Q{x:=<a>.(P{b:=(y).Q})})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7199
      using asm prems neq by (auto simp add: fresh_prod fresh_atm subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7200
    also have "\<dots> = Cut <c>.(M{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}) (y).Q" using asm prems by (simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7201
    finally have eq2: "(Cut <c>.M (z).N){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})} 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7202
                                       = Cut <c>.(M{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}) (y).Q" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7203
    have "(Cut <c>.M (z).N){x:=<a>.P}{b:=(y).Q} = (Cut <c>.M (z).N){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7204
      using eq1 eq2 by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7205
  }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7206
  moreover 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7207
  { assume asm: "M = Ax x c \<and> N \<noteq> Ax z b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7208
    have neq: "N{x:=<a>.P} \<noteq> Ax z b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7209
    proof (cases "x\<sharp>N")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7210
      case True then show ?thesis using asm by (simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7211
    next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7212
      case False then show ?thesis by (simp add: not_Ax2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7213
    qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7214
    have "(Cut <c>.M (z).N){x:=<a>.P}{b:=(y).Q} = (Cut <a>.P (z).(N{x:=<a>.P})){b:=(y).Q}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7215
      using asm prems by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7216
    also have "\<dots> = Cut <a>.(P{b:=(y).Q}) (z).(N{x:=<a>.P}{b:=(y).Q})" using asm prems neq 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7217
      by (simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7218
    also have "\<dots> = Cut <a>.(P{b:=(y).Q}) (z).(N{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})})" using asm prems by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7219
    finally have eq1: "(Cut <c>.M (z).N){x:=<a>.P}{b:=(y).Q} 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7220
                    = Cut <a>.(P{b:=(y).Q}) (z).(N{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})})" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7221
    have "(Cut <c>.M (z).N){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})} 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7222
                    = (Cut <c>.(M{b:=(y).Q}) (z).(N{b:=(y).Q})){x:=<a>.(P{b:=(y).Q})}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7223
      using asm prems by auto
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7224
    also have "\<dots> = (Cut <c>.M (z).(N{b:=(y).Q})){x:=<a>.(P{b:=(y).Q})}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7225
      using asm prems by (auto simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7226
    also have "\<dots> = Cut <a>.(P{b:=(y).Q}) (z).(N{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})})" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7227
      using asm prems by (simp add: fresh_prod fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7228
    finally 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7229
    have eq2: "(Cut <c>.M (z).N){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})} 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7230
         = Cut <a>.(P{b:=(y).Q}) (z).(N{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})})" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7231
    have "(Cut <c>.M (z).N){x:=<a>.P}{b:=(y).Q} = (Cut <c>.M (z).N){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7232
      using eq1 eq2 by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7233
  }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7234
  moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7235
  { assume asm: "M \<noteq> Ax x c \<and> N \<noteq> Ax z b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7236
    have neq1: "N{x:=<a>.P} \<noteq> Ax z b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7237
    proof (cases "x\<sharp>N")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7238
      case True then show ?thesis using asm by (simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7239
    next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7240
      case False then show ?thesis by (simp add: not_Ax2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7241
    qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7242
    have neq2: "M{b:=(y).Q} \<noteq> Ax x c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7243
    proof (cases "b\<sharp>M")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7244
      case True then show ?thesis using asm by (simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7245
    next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7246
      case False then show ?thesis by (simp add: not_Ax1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7247
    qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7248
    have "(Cut <c>.M (z).N){x:=<a>.P}{b:=(y).Q} = (Cut <c>.(M{x:=<a>.P}) (z).(N{x:=<a>.P})){b:=(y).Q}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7249
      using asm prems by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7250
    also have "\<dots> = Cut <c>.(M{x:=<a>.P}{b:=(y).Q}) (z).(N{x:=<a>.P}{b:=(y).Q})" using asm prems neq1
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7251
      by (simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7252
    also have "\<dots> = Cut <c>.(M{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}) (z).(N{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7253
      using asm prems by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7254
    finally have eq1: "(Cut <c>.M (z).N){x:=<a>.P}{b:=(y).Q}
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7255
             = Cut <c>.(M{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}) (z).(N{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})})" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7256
    have "(Cut <c>.M (z).N){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7257
                (Cut <c>.(M{b:=(y).Q}) (z).(N{b:=(y).Q})){x:=<a>.(P{b:=(y).Q})}" using asm neq1 prems by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7258
    also have "\<dots> = Cut <c>.(M{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}) (z).(N{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7259
      using asm neq2 prems by (simp add: fresh_prod fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7260
    finally have eq2: "(Cut <c>.M (z).N){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7261
           Cut <c>.(M{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}) (z).(N{b:=(y).Q}{x:=<a>.(P{b:=(y).Q})})" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7262
    have "(Cut <c>.M (z).N){x:=<a>.P}{b:=(y).Q} = (Cut <c>.M (z).N){b:=(y).Q}{x:=<a>.(P{b:=(y).Q})}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7263
      using eq1 eq2 by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7264
  }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7265
  ultimately show ?case by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7266
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7267
  case (NotR z M c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7268
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7269
    apply(auto simp add: fresh_prod fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7270
    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(M{c:=(y).Q},M{c:=(y).Q}{x:=<a>.P{c:=(y).Q}},Q,a,P,c,y)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7271
    apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7272
    apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7273
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7274
    apply(simp add: fresh_fun_simp_NotR abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7275
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7276
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7277
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7278
    apply(simp add: fresh_prod fresh_atm subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7279
    apply(simp add: fresh_prod fresh_atm subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7280
    apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7281
    apply(simp add: fresh_prod fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7282
    apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7283
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7284
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7285
  case (NotL c M z)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7286
  then show ?case  
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7287
    apply(auto simp add: fresh_prod fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7288
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(P,M{x:=<a>.P},P{b:=(y).Q},M{b:=(y).Q}{x:=<a>.P{b:=(y).Q}},y,Q)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7289
    apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7290
    apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7291
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7292
    apply(simp add: fresh_fun_simp_NotL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7293
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7294
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7295
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7296
  case (AndR c1 M c2 N c3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7297
  then show ?case  
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7298
    apply(auto simp add: fresh_prod fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7299
    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(Q,M{c3:=(y).Q},M{c3:=(y).Q}{x:=<a>.P{c3:=(y).Q}},c2,c3,a,
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7300
                                     P{c3:=(y).Q},N{c3:=(y).Q},N{c3:=(y).Q}{x:=<a>.P{c3:=(y).Q}},c1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7301
    apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7302
    apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7303
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7304
    apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7305
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7306
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7307
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7308
    apply(simp_all add: fresh_atm abs_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7309
    apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7310
    apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7311
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7312
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7313
  case (AndL1 z1 M z2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7314
  then show ?case  
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7315
    apply(auto simp add: fresh_prod fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7316
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(P,M{x:=<a>.P},P{b:=(y).Q},z1,y,Q,M{b:=(y).Q}{x:=<a>.P{b:=(y).Q}})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7317
    apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7318
    apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7319
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7320
    apply(simp add: fresh_fun_simp_AndL1 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7321
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7322
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7323
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7324
  case (AndL2 z1 M z2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7325
  then show ?case  
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7326
    apply(auto simp add: fresh_prod fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7327
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(P,M{x:=<a>.P},P{b:=(y).Q},z1,y,Q,M{b:=(y).Q}{x:=<a>.P{b:=(y).Q}})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7328
    apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7329
    apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7330
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7331
    apply(simp add: fresh_fun_simp_AndL2 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7332
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7333
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7334
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7335
  case (OrL z1 M z2 N z3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7336
  then show ?case  
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7337
    apply(auto simp add: fresh_prod fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7338
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(P,M{x:=<a>.P},M{b:=(y).Q}{x:=<a>.P{b:=(y).Q}},z2,z3,a,y,Q,
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7339
                                     P{b:=(y).Q},N{x:=<a>.P},N{b:=(y).Q}{x:=<a>.P{b:=(y).Q}},z1)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7340
    apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7341
    apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7342
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7343
    apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7344
    apply(simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7345
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7346
    apply(rule substc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7347
    apply(simp_all add: fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7348
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7349
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7350
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7351
  case (OrR1 c1 M c2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7352
  then show ?case  
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7353
    apply(auto simp add: fresh_prod fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7354
    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(Q,M{c2:=(y).Q},a,P{c2:=(y).Q},c1,
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7355
                                                     M{c2:=(y).Q}{x:=<a>.P{c2:=(y).Q}})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7356
    apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7357
    apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7358
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7359
    apply(simp add: fresh_fun_simp_OrR1 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7360
    apply(simp_all add: fresh_atm subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7361
    apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7362
    apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7363
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7364
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7365
  case (OrR2 c1 M c2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7366
  then show ?case  
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7367
    apply(auto simp add: fresh_prod fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7368
    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(Q,M{c2:=(y).Q},a,P{c2:=(y).Q},c1,
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7369
                                                     M{c2:=(y).Q}{x:=<a>.P{c2:=(y).Q}})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7370
    apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7371
    apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7372
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7373
    apply(simp add: fresh_fun_simp_OrR2 abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7374
    apply(simp_all add: fresh_atm subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7375
    apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7376
    apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7377
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7378
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7379
  case (ImpR z c M d)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7380
  then show ?case  
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7381
    apply(auto simp add: fresh_prod fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7382
    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(Q,M{d:=(y).Q},a,P{d:=(y).Q},c,
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7383
                                                     M{d:=(y).Q}{x:=<a>.P{d:=(y).Q}})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7384
    apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7385
    apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7386
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7387
    apply(simp add: fresh_fun_simp_ImpR abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7388
    apply(simp_all add: fresh_atm subst_fresh forget abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7389
    apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7390
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7391
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7392
  case (ImpL c M z N u)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7393
  then show ?case  
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7394
    apply(auto simp add: fresh_prod fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7395
    apply(subgoal_tac "\<exists>z'::name. z'\<sharp>(P,P{b:=(y).Q},M{u:=<a>.P},N{u:=<a>.P},y,Q,
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7396
                        M{b:=(y).Q}{u:=<a>.P{b:=(y).Q}},N{b:=(y).Q}{u:=<a>.P{b:=(y).Q}},z)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7397
    apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7398
    apply(simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7399
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7400
    apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7401
    apply(simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7402
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7403
    apply(rule substc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7404
    apply(simp_all add: fresh_atm subst_fresh forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7405
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7406
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7407
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7408
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7409
lemma subst_subst2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7410
  assumes a: "a\<sharp>(b,P,N)" "x\<sharp>(y,P,M)" "b\<sharp>(M,N)" "y\<sharp>P"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7411
  shows "M{a:=(x).N}{y:=<b>.P} = M{y:=<b>.P}{a:=(x).N{y:=<b>.P}}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7412
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7413
proof(nominal_induct M avoiding: a x N y b P rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7414
  case (Ax z c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7415
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7416
    by (auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7417
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7418
  case (Cut d M' u M'')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7419
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7420
    apply(simp add: fresh_atm fresh_prod trm.inject abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7421
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7422
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7423
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7424
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7425
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7426
    apply(simp add: abs_fresh subst_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7427
    apply(simp add: fresh_prod subst_fresh fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7428
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7429
    apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7430
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7431
    apply(case_tac "a\<sharp>M'")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7432
    apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7433
    apply(simp add: not_Ax1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7434
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7435
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7436
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7437
    apply(simp add: abs_fresh subst_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7438
    apply(simp add: fresh_prod subst_fresh fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7439
    apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7440
    apply(case_tac "y\<sharp>M''")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7441
    apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7442
    apply(simp add: not_Ax2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7443
    apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7444
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7445
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7446
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7447
    apply(simp add: subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7448
    apply(simp add: abs_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7449
    apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7450
    apply(case_tac "y\<sharp>M''")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7451
    apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7452
    apply(simp add: not_Ax2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7453
    apply(case_tac "a\<sharp>M'")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7454
    apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7455
    apply(simp add: not_Ax1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7456
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7457
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7458
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7459
    apply(simp add: subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7460
    apply(simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7461
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7462
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7463
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7464
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7465
    apply(simp add: subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7466
    apply(simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7467
    apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7468
    apply(case_tac "y\<sharp>M''")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7469
    apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7470
    apply(simp add: not_Ax2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7471
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7472
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7473
  case (NotR z M' d) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7474
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7475
    apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7476
    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(y,P,N,N{y:=<b>.P},M'{d:=(x).N},M'{y:=<b>.P}{d:=(x).N{y:=<b>.P}})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7477
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7478
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7479
    apply(simp add: fresh_fun_simp_NotR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7480
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7481
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7482
    apply(simp add: abs_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7483
    apply(simp add: fresh_prod subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7484
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7485
    apply(simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7486
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7487
    apply(rule substn.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7488
    apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7489
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7490
    apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7491
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7492
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7493
  case (NotL d M' z) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7494
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7495
    apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7496
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(z,y,P,N,N{y:=<b>.P},M'{y:=<b>.P},M'{y:=<b>.P}{a:=(x).N{y:=<b>.P}})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7497
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7498
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7499
    apply(simp add: fresh_fun_simp_NotL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7500
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7501
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7502
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7503
    apply(simp add: abs_fresh subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7504
    apply(simp add: fresh_prod subst_fresh fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7505
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7506
    apply(simp add: trm.inject alpha forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7507
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7508
    apply(rule substc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7509
    apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7510
    apply(simp add: fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7511
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7512
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7513
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7514
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7515
  case (AndR d M' e M'' f) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7516
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7517
    apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7518
    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(P,b,d,e,N,N{y:=<b>.P},M'{f:=(x).N},M''{f:=(x).N},
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7519
                  M'{y:=<b>.P}{f:=(x).N{y:=<b>.P}},M''{y:=<b>.P}{f:=(x).N{y:=<b>.P}})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7520
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7521
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7522
    apply(simp add: fresh_fun_simp_AndR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7523
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7524
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7525
    apply(simp add: abs_fresh subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7526
    apply(simp add: fresh_prod subst_fresh fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7527
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7528
    apply(simp add: trm.inject alpha forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7529
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7530
    apply(rule substn.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7531
    apply(simp add: fresh_prod fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7532
    apply(simp add: fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7533
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7534
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7535
    apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7536
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7537
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7538
  case (AndL1 z M' u) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7539
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7540
    apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7541
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(P,b,z,u,x,N,M'{y:=<b>.P},M'{y:=<b>.P}{a:=(x).N{y:=<b>.P}})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7542
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7543
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7544
    apply(simp add: fresh_fun_simp_AndL1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7545
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7546
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7547
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7548
    apply(simp add: abs_fresh subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7549
    apply(simp add: fresh_prod subst_fresh fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7550
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7551
    apply(simp add: trm.inject alpha forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7552
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7553
    apply(rule substc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7554
    apply(simp add: fresh_prod fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7555
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7556
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7557
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7558
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7559
  case (AndL2 z M' u) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7560
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7561
    apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7562
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(P,b,z,u,x,N,M'{y:=<b>.P},M'{y:=<b>.P}{a:=(x).N{y:=<b>.P}})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7563
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7564
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7565
    apply(simp add: fresh_fun_simp_AndL2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7566
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7567
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7568
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7569
    apply(simp add: abs_fresh subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7570
    apply(simp add: fresh_prod subst_fresh fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7571
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7572
    apply(simp add: trm.inject alpha forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7573
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7574
    apply(rule substc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7575
    apply(simp add: fresh_prod fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7576
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7577
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7578
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7579
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7580
  case (OrL u M' v M'' w) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7581
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7582
    apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7583
    apply(subgoal_tac "\<exists>z'::name. z'\<sharp>(P,b,u,w,v,N,N{y:=<b>.P},M'{y:=<b>.P},M''{y:=<b>.P},
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7584
                  M'{y:=<b>.P}{a:=(x).N{y:=<b>.P}},M''{y:=<b>.P}{a:=(x).N{y:=<b>.P}})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7585
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7586
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7587
    apply(simp add: fresh_fun_simp_OrL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7588
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7589
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7590
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7591
    apply(simp add: abs_fresh subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7592
    apply(simp add: fresh_prod subst_fresh fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7593
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7594
    apply(simp add: trm.inject alpha forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7595
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7596
    apply(rule substc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7597
    apply(simp add: fresh_prod fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7598
    apply(simp add: fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7599
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7600
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7601
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7602
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7603
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7604
  case (OrR1 e M' f) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7605
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7606
    apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7607
    apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(P,b,e,f,x,N,N{y:=<b>.P},
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7608
                                        M'{f:=(x).N},M'{y:=<b>.P}{f:=(x).N{y:=<b>.P}})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7609
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7610
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7611
    apply(simp add: fresh_fun_simp_OrR1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7612
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7613
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7614
    apply(simp add: abs_fresh subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7615
    apply(simp add: fresh_prod subst_fresh fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7616
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7617
    apply(simp add: trm.inject alpha forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7618
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7619
    apply(rule substn.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7620
    apply(simp add: fresh_prod fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7621
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7622
    apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7623
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7624
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7625
  case (OrR2 e M' f) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7626
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7627
    apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7628
    apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(P,b,e,f,x,N,N{y:=<b>.P},
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7629
                                        M'{f:=(x).N},M'{y:=<b>.P}{f:=(x).N{y:=<b>.P}})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7630
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7631
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7632
    apply(simp add: fresh_fun_simp_OrR2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7633
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7634
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7635
    apply(simp add: abs_fresh subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7636
    apply(simp add: fresh_prod subst_fresh fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7637
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7638
    apply(simp add: trm.inject alpha forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7639
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7640
    apply(rule substn.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7641
    apply(simp add: fresh_prod fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7642
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7643
    apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7644
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7645
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7646
  case (ImpR x e M' f) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7647
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7648
    apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7649
    apply(subgoal_tac "\<exists>c'::coname. c'\<sharp>(P,b,e,f,x,N,N{y:=<b>.P},
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7650
                                        M'{f:=(x).N},M'{y:=<b>.P}{f:=(x).N{y:=<b>.P}})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7651
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7652
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7653
    apply(simp add: fresh_fun_simp_ImpR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7654
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7655
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7656
    apply(simp add: abs_fresh subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7657
    apply(simp add: fresh_prod subst_fresh fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7658
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7659
    apply(simp add: trm.inject alpha forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7660
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7661
    apply(rule substn.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7662
    apply(simp add: fresh_prod fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7663
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7664
    apply(simp add: fresh_atm trm.inject alpha abs_fresh fin_supp abs_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7665
    apply(rule exists_fresh'(2)[OF fs_coname1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7666
    apply(simp add: fresh_atm trm.inject alpha abs_fresh fin_supp abs_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7667
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7668
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7669
  case (ImpL e M' v M'' w) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7670
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7671
    apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7672
    apply(subgoal_tac "\<exists>z'::name. z'\<sharp>(P,b,e,w,v,N,N{y:=<b>.P},M'{w:=<b>.P},M''{w:=<b>.P},
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7673
                  M'{w:=<b>.P}{a:=(x).N{w:=<b>.P}},M''{w:=<b>.P}{a:=(x).N{w:=<b>.P}})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7674
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7675
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7676
    apply(simp add: fresh_fun_simp_ImpL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7677
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7678
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7679
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7680
    apply(simp add: abs_fresh subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7681
    apply(simp add: fresh_prod subst_fresh fresh_atm abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7682
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7683
    apply(simp add: trm.inject alpha forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7684
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7685
    apply(rule substc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7686
    apply(simp add: fresh_prod fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7687
    apply(simp add: fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7688
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7689
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7690
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7691
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7692
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7693
lemma subst_subst3:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7694
  assumes a: "a\<sharp>(P,N,c)" "c\<sharp>(M,N)" "x\<sharp>(y,P,M)" "y\<sharp>(P,x)" "M\<noteq>Ax y a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7695
  shows "N{x:=<a>.M}{y:=<c>.P} = N{y:=<c>.P}{x:=<a>.(M{y:=<c>.P})}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7696
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7697
proof(nominal_induct N avoiding: x y a c M P rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7698
  case (Ax z c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7699
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7700
    by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7701
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7702
  case (Cut d M' u M'')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7703
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7704
    apply(simp add: fresh_atm fresh_prod trm.inject abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7705
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7706
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7707
    apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7708
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7709
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7710
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7711
    apply(simp add: abs_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7712
    apply(simp add: fresh_prod subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7713
    apply(subgoal_tac "P \<noteq> Ax x c")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7714
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7715
    apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7716
    apply(clarify)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7717
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7718
    apply(case_tac "x\<sharp>M'")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7719
    apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7720
    apply(simp add: not_Ax2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7721
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7722
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7723
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7724
    apply(simp add: abs_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7725
    apply(simp add: fresh_prod subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7726
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7727
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7728
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7729
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7730
    apply(simp add: abs_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7731
    apply(simp add: fresh_prod subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7732
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7733
    apply(case_tac "y\<sharp>M'")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7734
    apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7735
    apply(simp add: not_Ax2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7736
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7737
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7738
  case NotR
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7739
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7740
    by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7741
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7742
  case (NotL d M' u)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7743
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7744
    apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7745
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(y,P,M,M{y:=<c>.P},M'{x:=<a>.M},M'{y:=<c>.P}{x:=<a>.M{y:=<c>.P}})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7746
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7747
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7748
    apply(simp add: fresh_fun_simp_NotL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7749
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7750
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7751
    apply(simp add: abs_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7752
    apply(simp add: fresh_prod subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7753
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7754
    apply(simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7755
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7756
    apply(rule substn.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7757
    apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7758
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7759
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7760
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(x,y,P,M,M'{y:=<c>.P},M'{y:=<c>.P}{x:=<a>.M{y:=<c>.P}})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7761
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7762
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7763
    apply(simp add: fresh_fun_simp_NotL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7764
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7765
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7766
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7767
    apply(simp add: abs_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7768
    apply(simp add: fresh_atm subst_fresh fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7769
    apply(subgoal_tac "P \<noteq> Ax x c")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7770
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7771
    apply(simp add: forget trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7772
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7773
    apply(rule substn.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7774
    apply(simp add: fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7775
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7776
    apply(clarify)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7777
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7778
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7779
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7780
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7781
  case AndR
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7782
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7783
    by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7784
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7785
  case (AndL1 u M' v)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7786
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7787
    apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7788
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(u,y,v,P,M,M{y:=<c>.P},M'{x:=<a>.M},M'{y:=<c>.P}{x:=<a>.M{y:=<c>.P}})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7789
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7790
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7791
    apply(simp add: fresh_fun_simp_AndL1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7792
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7793
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7794
    apply(simp add: abs_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7795
    apply(simp add: fresh_prod subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7796
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7797
    apply(simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7798
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7799
    apply(rule substn.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7800
    apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7801
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7802
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7803
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(x,y,u,v,P,M,M'{y:=<c>.P},M'{y:=<c>.P}{x:=<a>.M{y:=<c>.P}})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7804
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7805
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7806
    apply(simp add: fresh_fun_simp_AndL1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7807
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7808
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7809
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7810
    apply(simp add: abs_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7811
    apply(simp add: fresh_atm subst_fresh fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7812
    apply(subgoal_tac "P \<noteq> Ax x c")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7813
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7814
    apply(simp add: forget trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7815
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7816
    apply(rule substn.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7817
    apply(simp add: fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7818
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7819
    apply(clarify)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7820
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7821
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7822
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7823
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7824
  case (AndL2 u M' v)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7825
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7826
    apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7827
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(u,y,v,P,M,M{y:=<c>.P},M'{x:=<a>.M},M'{y:=<c>.P}{x:=<a>.M{y:=<c>.P}})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7828
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7829
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7830
    apply(simp add: fresh_fun_simp_AndL2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7831
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7832
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7833
    apply(simp add: abs_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7834
    apply(simp add: fresh_prod subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7835
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7836
    apply(simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7837
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7838
    apply(rule substn.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7839
    apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7840
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7841
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7842
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(x,y,u,v,P,M,M'{y:=<c>.P},M'{y:=<c>.P}{x:=<a>.M{y:=<c>.P}})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7843
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7844
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7845
    apply(simp add: fresh_fun_simp_AndL2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7846
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7847
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7848
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7849
    apply(simp add: abs_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7850
    apply(simp add: fresh_atm subst_fresh fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7851
    apply(subgoal_tac "P \<noteq> Ax x c")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7852
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7853
    apply(simp add: forget trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7854
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7855
    apply(rule substn.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7856
    apply(simp add: fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7857
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7858
    apply(clarify)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7859
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7860
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7861
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7862
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7863
  case OrR1
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7864
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7865
    by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7866
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7867
  case OrR2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7868
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7869
    by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7870
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7871
  case (OrL x1 M' x2 M'' x3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7872
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7873
    apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7874
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(y,P,M,M{y:=<c>.P},M'{x:=<a>.M},M'{y:=<c>.P}{x:=<a>.M{y:=<c>.P}},
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7875
                                      x1,x2,x3,M''{x:=<a>.M},M''{y:=<c>.P}{x:=<a>.M{y:=<c>.P}})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7876
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7877
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7878
    apply(simp add: fresh_fun_simp_OrL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7879
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7880
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7881
    apply(simp add: abs_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7882
    apply(simp add: fresh_prod subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7883
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7884
    apply(simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7885
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7886
    apply(rule substn.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7887
    apply(simp add: fresh_prod fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7888
    apply(simp add: fresh_prod fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7889
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7890
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7891
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7892
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(x,y,P,M,M'{y:=<c>.P},M'{y:=<c>.P}{x:=<a>.M{y:=<c>.P}},
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7893
                                      x1,x2,x3,M''{y:=<c>.P},M''{y:=<c>.P}{x:=<a>.M{y:=<c>.P}})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7894
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7895
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7896
    apply(simp add: fresh_fun_simp_OrL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7897
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7898
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7899
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7900
    apply(simp add: abs_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7901
    apply(simp add: fresh_atm subst_fresh fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7902
    apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7903
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7904
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7905
    apply(simp add: forget trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7906
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7907
    apply(rule substn.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7908
    apply(simp add: fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7909
    apply(simp add: fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7910
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7911
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7912
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7913
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7914
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7915
  case ImpR
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7916
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7917
    by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7918
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7919
  case (ImpL d M' x1 M'' x2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7920
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7921
    apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7922
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(y,P,M,M{y:=<c>.P},M'{x2:=<a>.M},M'{y:=<c>.P}{x2:=<a>.M{y:=<c>.P}},
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7923
                                      x1,x2,M''{x2:=<a>.M},M''{y:=<c>.P}{x2:=<a>.M{y:=<c>.P}})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7924
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7925
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7926
    apply(simp add: fresh_fun_simp_ImpL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7927
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7928
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7929
    apply(simp add: abs_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7930
    apply(simp add: fresh_prod subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7931
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7932
    apply(simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7933
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7934
    apply(rule substn.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7935
    apply(simp add: fresh_prod fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7936
    apply(simp add: fresh_prod fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7937
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7938
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7939
    apply(subgoal_tac "\<exists>x'::name. x'\<sharp>(x,y,P,M,M'{x2:=<c>.P},M'{x2:=<c>.P}{x:=<a>.M{x2:=<c>.P}},
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7940
                                      x1,x2,M''{x2:=<c>.P},M''{x2:=<c>.P}{x:=<a>.M{x2:=<c>.P}})")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7941
    apply(erule exE, simp add: fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7942
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7943
    apply(simp add: fresh_fun_simp_ImpL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7944
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7945
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7946
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7947
    apply(simp add: abs_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7948
    apply(simp add: fresh_atm subst_fresh fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7949
    apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7950
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7951
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7952
    apply(simp add: forget trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7953
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7954
    apply(rule substn.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7955
    apply(simp add: fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7956
    apply(simp add: fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7957
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7958
    apply(rule exists_fresh'(1)[OF fs_name1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7959
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7960
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7961
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7962
lemma subst_subst4:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7963
  assumes a: "x\<sharp>(P,N,y)" "y\<sharp>(M,N)" "a\<sharp>(c,P,M)" "c\<sharp>(P,a)" "M\<noteq>Ax x c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7964
  shows "N{a:=(x).M}{c:=(y).P} = N{c:=(y).P}{a:=(x).(M{c:=(y).P})}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7965
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7966
proof(nominal_induct N avoiding: x y a c M P rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7967
  case (Ax z c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7968
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7969
    by (auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7970
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7971
  case (Cut d M' u M'')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7972
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7973
    apply(simp add: fresh_atm fresh_prod trm.inject abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7974
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7975
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7976
    apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7977
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7978
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7979
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7980
    apply(simp add: abs_fresh subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7981
    apply(simp add: fresh_prod subst_fresh abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7982
    apply(subgoal_tac "P \<noteq> Ax y a")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7983
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7984
    apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7985
    apply(clarify)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7986
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7987
    apply(case_tac "a\<sharp>M''")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7988
    apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7989
    apply(simp add: not_Ax1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7990
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7991
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7992
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7993
    apply(simp add: fresh_prod subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7994
    apply(simp add: abs_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7995
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7996
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7997
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7998
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  7999
    apply(simp add: fresh_prod subst_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8000
    apply(simp add: abs_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8001
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8002
    apply(case_tac "c\<sharp>M''")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8003
    apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8004
    apply(simp add: not_Ax1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8005
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8006
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8007
  case NotL
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8008
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8009
    by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8010
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8011
  case (NotR u M' d)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8012
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8013
    apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8014
    apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8015
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8016
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8017
    apply(simp add: abs_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8018
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8019
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8020
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8021
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8022
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8023
    apply(simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8024
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8025
    apply(rule substc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8026
    apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8027
    apply(auto simp add: fresh_atm fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8028
    apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8029
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8030
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8031
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8032
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8033
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8034
    apply(simp add: fresh_prod fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8035
    apply(simp add: abs_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8036
    apply(auto simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8037
    apply(simp add: trm.inject alpha forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8038
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8039
    apply(rule substc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8040
    apply(simp add: fresh_atm subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8041
    apply(auto simp add: fresh_prod fresh_atm) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8042
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8043
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8044
  case AndL1
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8045
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8046
    by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8047
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8048
  case AndL2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8049
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8050
    by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8051
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8052
  case (AndR d M e M' f)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8053
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8054
    apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8055
    apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8056
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8057
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8058
    apply(simp add: abs_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8059
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8060
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8061
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8062
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8063
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8064
    apply(simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8065
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8066
    apply(rule substc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8067
    apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8068
    apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8069
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8070
    apply(auto simp add: fresh_atm fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8071
    apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8072
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8073
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8074
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8075
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8076
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8077
    apply(simp add: subst_fresh fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8078
    apply(simp add: abs_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8079
    apply(auto simp add: fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8080
    apply(simp add: trm.inject alpha forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8081
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8082
    apply(rule substc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8083
    apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8084
    apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8085
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8086
    apply(auto simp add: fresh_atm fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8087
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8088
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8089
  case OrL
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8090
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8091
    by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8092
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8093
  case (OrR1 d M' e)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8094
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8095
    apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8096
    apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8097
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8098
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8099
    apply(simp add: abs_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8100
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8101
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8102
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8103
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8104
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8105
    apply(simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8106
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8107
    apply(rule substc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8108
    apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8109
    apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8110
    apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8111
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8112
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8113
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8114
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8115
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8116
    apply(simp add: subst_fresh fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8117
    apply(simp add: abs_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8118
    apply(auto simp add: fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8119
    apply(simp add: trm.inject alpha forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8120
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8121
    apply(rule substc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8122
    apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8123
    apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8124
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8125
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8126
  case (OrR2 d M' e)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8127
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8128
    apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8129
    apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8130
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8131
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8132
    apply(simp add: abs_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8133
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8134
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8135
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8136
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8137
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8138
    apply(simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8139
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8140
    apply(rule substc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8141
    apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8142
    apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8143
    apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8144
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8145
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8146
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8147
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8148
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8149
    apply(simp add: subst_fresh fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8150
    apply(simp add: abs_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8151
    apply(auto simp add: fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8152
    apply(simp add: trm.inject alpha forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8153
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8154
    apply(rule substc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8155
    apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8156
    apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8157
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8158
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8159
  case ImpL
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8160
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8161
    by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8162
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8163
  case (ImpR u d M' e)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8164
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8165
    apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8166
    apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8167
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8168
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8169
    apply(simp add: abs_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8170
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8171
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8172
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8173
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8174
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8175
    apply(simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8176
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8177
    apply(rule substc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8178
    apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8179
    apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8180
    apply(auto simp add: fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8181
    apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8182
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8183
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8184
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8185
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8186
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8187
    apply(simp add: subst_fresh fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8188
    apply(simp add: abs_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8189
    apply(auto simp add: fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8190
    apply(simp add: trm.inject alpha forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8191
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8192
    apply(rule substc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8193
    apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8194
    apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8195
    apply(auto simp add: fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8196
    apply(auto simp add: fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8197
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8198
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8199
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8200
text {* Reduction *}
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8201
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8202
lemma fin_not_Cut:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8203
  assumes a: "fin M x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8204
  shows "\<not>(\<exists>a M' x N'. M = Cut <a>.M' (x).N')"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8205
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8206
by (induct) (auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8207
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8208
lemma fresh_not_fin:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8209
  assumes a: "x\<sharp>M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8210
  shows "\<not>fin M x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8211
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8212
  have "fin M x \<Longrightarrow> x\<sharp>M \<Longrightarrow> False" by (induct rule: fin.induct) (auto simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8213
  with a show "\<not>fin M x" by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8214
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8215
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8216
lemma fresh_not_fic:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8217
  assumes a: "a\<sharp>M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8218
  shows "\<not>fic M a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8219
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8220
  have "fic M a \<Longrightarrow> a\<sharp>M \<Longrightarrow> False" by (induct rule: fic.induct) (auto simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8221
  with a show "\<not>fic M a" by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8222
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8223
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8224
lemma c_redu_subst1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8225
  assumes a: "M \<longrightarrow>\<^isub>c M'" "c\<sharp>M" "y\<sharp>P"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8226
  shows "M{y:=<c>.P} \<longrightarrow>\<^isub>c M'{y:=<c>.P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8227
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8228
proof(nominal_induct avoiding: y c P rule: c_redu.strong_induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8229
  case (left M a N x)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8230
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8231
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8232
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8233
    apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8234
    apply(force)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8235
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8236
    apply(subgoal_tac "M{a:=(x).N}{y:=<c>.P} = M{y:=<c>.P}{a:=(x).(N{y:=<c>.P})}")(*A*)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8237
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8238
    apply(rule c_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8239
    apply(rule not_fic_subst1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8240
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8241
    apply(simp add: subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8242
    apply(simp add: subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8243
    apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8244
    apply(rule subst_subst2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8245
    apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8246
    apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8247
    apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8248
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8249
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8250
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8251
  case (right N x a M)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8252
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8253
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8254
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8255
    apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8256
    (* case M = Ax y a *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8257
    apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8258
    apply(subgoal_tac "N{x:=<a>.Ax y a}{y:=<c>.P} = N{y:=<c>.P}{x:=<c>.P}")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8259
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8260
    apply(rule c_redu.right)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8261
    apply(rule not_fin_subst2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8262
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8263
    apply(rule subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8264
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8265
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8266
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8267
    apply(rule interesting_subst1')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8268
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8269
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8270
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8271
    (* case M \<noteq> Ax y a*)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8272
    apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8273
    apply(subgoal_tac "N{x:=<a>.M}{y:=<c>.P} = N{y:=<c>.P}{x:=<a>.(M{y:=<c>.P})}")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8274
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8275
    apply(rule c_redu.right)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8276
    apply(rule not_fin_subst2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8277
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8278
    apply(simp add: subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8279
    apply(simp add: subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8280
    apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8281
    apply(rule subst_subst3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8282
    apply(simp_all add: fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8283
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8284
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8285
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8286
lemma c_redu_subst2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8287
  assumes a: "M \<longrightarrow>\<^isub>c M'" "c\<sharp>P" "y\<sharp>M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8288
  shows "M{c:=(y).P} \<longrightarrow>\<^isub>c M'{c:=(y).P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8289
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8290
proof(nominal_induct avoiding: y c P rule: c_redu.strong_induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8291
  case (right N x a M)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8292
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8293
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8294
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8295
    apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8296
    apply(force)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8297
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8298
    apply(subgoal_tac "N{x:=<a>.M}{c:=(y).P} = N{c:=(y).P}{x:=<a>.(M{c:=(y).P})}")(*A*)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8299
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8300
    apply(rule c_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8301
    apply(rule not_fin_subst1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8302
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8303
    apply(simp add: subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8304
    apply(simp add: subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8305
    apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8306
    apply(rule subst_subst1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8307
    apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8308
    apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8309
    apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8310
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8311
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8312
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8313
  case (left M a N x)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8314
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8315
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8316
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8317
    apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8318
    (* case N = Ax x c *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8319
    apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8320
    apply(subgoal_tac "M{a:=(x).Ax x c}{c:=(y).P} = M{c:=(y).P}{a:=(y).P}")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8321
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8322
    apply(rule c_redu.left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8323
    apply(rule not_fic_subst2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8324
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8325
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8326
    apply(rule subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8327
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8328
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8329
    apply(rule interesting_subst2')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8330
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8331
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8332
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8333
    (* case M \<noteq> Ax y a*)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8334
    apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8335
    apply(subgoal_tac "M{a:=(x).N}{c:=(y).P} = M{c:=(y).P}{a:=(x).(N{c:=(y).P})}")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8336
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8337
    apply(rule c_redu.left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8338
    apply(rule not_fic_subst2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8339
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8340
    apply(simp add: subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8341
    apply(simp add: subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8342
    apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8343
    apply(rule subst_subst4)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8344
    apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8345
    apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8346
    apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8347
    apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8348
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8349
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8350
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8351
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8352
lemma c_redu_subst1':
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8353
  assumes a: "M \<longrightarrow>\<^isub>c M'" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8354
  shows "M{y:=<c>.P} \<longrightarrow>\<^isub>c M'{y:=<c>.P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8355
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8356
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8357
  obtain y'::"name"   where fs1: "y'\<sharp>(M,M',P,P,y)" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8358
  obtain c'::"coname" where fs2: "c'\<sharp>(M,M',P,P,c)" by (rule exists_fresh(2), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8359
  have "M{y:=<c>.P} = ([(y',y)]\<bullet>M){y':=<c'>.([(c',c)]\<bullet>P)}" using fs1 fs2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8360
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8361
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8362
    apply(rule_tac y="y'" in subst_rename(3))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8363
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8364
    apply(rule subst_rename(4))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8365
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8366
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8367
  also have "\<dots> \<longrightarrow>\<^isub>c ([(y',y)]\<bullet>M'){y':=<c'>.([(c',c)]\<bullet>P)}" using fs1 fs2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8368
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8369
    apply(rule c_redu_subst1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8370
    apply(simp add: c_redu.eqvt a)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8371
    apply(simp_all add: fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8372
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8373
  also have "\<dots> = M'{y:=<c>.P}" using fs1 fs2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8374
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8375
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8376
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8377
    apply(rule_tac y="y'" in subst_rename(3))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8378
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8379
    apply(rule subst_rename(4))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8380
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8381
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8382
  finally show ?thesis by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8383
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8384
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8385
lemma c_redu_subst2':
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8386
  assumes a: "M \<longrightarrow>\<^isub>c M'" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8387
  shows "M{c:=(y).P} \<longrightarrow>\<^isub>c M'{c:=(y).P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8388
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8389
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8390
  obtain y'::"name"   where fs1: "y'\<sharp>(M,M',P,P,y)" by (rule exists_fresh(1), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8391
  obtain c'::"coname" where fs2: "c'\<sharp>(M,M',P,P,c)" by (rule exists_fresh(2), rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8392
  have "M{c:=(y).P} = ([(c',c)]\<bullet>M){c':=(y').([(y',y)]\<bullet>P)}" using fs1 fs2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8393
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8394
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8395
    apply(rule_tac c="c'" in subst_rename(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8396
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8397
    apply(rule subst_rename(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8398
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8399
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8400
  also have "\<dots> \<longrightarrow>\<^isub>c ([(c',c)]\<bullet>M'){c':=(y').([(y',y)]\<bullet>P)}" using fs1 fs2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8401
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8402
    apply(rule c_redu_subst2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8403
    apply(simp add: c_redu.eqvt a)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8404
    apply(simp_all add: fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8405
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8406
  also have "\<dots> = M'{c:=(y).P}" using fs1 fs2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8407
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8408
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8409
    apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8410
    apply(rule_tac c="c'" in subst_rename(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8411
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8412
    apply(rule subst_rename(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8413
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8414
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8415
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8416
  finally show ?thesis by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8417
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8418
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8419
lemma aux1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8420
  assumes a: "M = M'" "M' \<longrightarrow>\<^isub>l M''"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8421
  shows "M \<longrightarrow>\<^isub>l M''"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8422
using a by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8423
  
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8424
lemma aux2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8425
  assumes a: "M \<longrightarrow>\<^isub>l M'" "M' = M''"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8426
  shows "M \<longrightarrow>\<^isub>l M''"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8427
using a by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8428
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8429
lemma aux3:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8430
  assumes a: "M = M'" "M' \<longrightarrow>\<^isub>a* M''"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8431
  shows "M \<longrightarrow>\<^isub>a* M''"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8432
using a by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8433
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8434
lemma aux4:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8435
  assumes a: "M = M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8436
  shows "M \<longrightarrow>\<^isub>a* M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8437
using a by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8438
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8439
lemma l_redu_subst1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8440
  assumes a: "M \<longrightarrow>\<^isub>l M'" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8441
  shows "M{y:=<c>.P} \<longrightarrow>\<^isub>a* M'{y:=<c>.P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8442
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8443
proof(nominal_induct M M' avoiding: y c P rule: l_redu.strong_induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8444
  case LAxR
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8445
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8446
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8447
    apply(rule aux3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8448
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8449
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8450
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8451
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8452
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8453
    apply(rule aux4)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8454
    apply(simp add: trm.inject alpha calc_atm fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8455
    apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8456
    apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8457
    apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8458
    apply(rule l_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8459
    apply(simp add: subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8460
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8461
    apply(rule fic_subst2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8462
    apply(simp_all)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8463
    apply(rule aux4)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8464
    apply(rule subst_comm')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8465
    apply(simp_all)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8466
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8467
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8468
  case LAxL
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8469
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8470
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8471
    apply(rule aux3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8472
    apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8473
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8474
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8475
    apply(simp add: trm.inject fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8476
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8477
    apply(rule aux4)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8478
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8479
    apply(rule fin_substn_nrename)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8480
    apply(simp_all)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8481
    apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8482
    apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8483
    apply(rule aux2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8484
    apply(rule l_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8485
    apply(simp add: subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8486
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8487
    apply(rule fin_subst1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8488
    apply(simp_all)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8489
    apply(rule subst_comm')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8490
    apply(simp_all)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8491
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8492
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8493
  case (LNot v M N u a b)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8494
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8495
  proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8496
    { assume asm: "N\<noteq>Ax y b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8497
      have "(Cut <a>.NotR (u).M a (v).NotL <b>.N v){y:=<c>.P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8498
        (Cut <a>.NotR (u).(M{y:=<c>.P}) a (v).NotL <b>.(N{y:=<c>.P}) v)" using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8499
	by (simp add: subst_fresh abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8500
      also have "\<dots> \<longrightarrow>\<^isub>l (Cut <b>.(N{y:=<c>.P}) (u).(M{y:=<c>.P}))" using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8501
	by (auto intro: l_redu.intros simp add: subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8502
      also have "\<dots> = (Cut <b>.N (u).M){y:=<c>.P}" using prems 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8503
	by (simp add: subst_fresh abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8504
      finally have ?thesis by auto
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8505
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8506
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8507
    { assume asm: "N=Ax y b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8508
      have "(Cut <a>.NotR (u).M a (v).NotL <b>.N v){y:=<c>.P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8509
        (Cut <a>.NotR (u).(M{y:=<c>.P}) a (v).NotL <b>.(N{y:=<c>.P}) v)" using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8510
	by (simp add: subst_fresh abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8511
      also have "\<dots> \<longrightarrow>\<^isub>a* (Cut <b>.(N{y:=<c>.P}) (u).(M{y:=<c>.P}))" using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8512
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8513
	apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8514
	apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8515
	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8516
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8517
      also have "\<dots> = (Cut <b>.(Cut <c>.P (y).Ax y b) (u).(M{y:=<c>.P}))" using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8518
	by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8519
      also have "\<dots> \<longrightarrow>\<^isub>a* (Cut <b>.(P[c\<turnstile>c>b]) (u).(M{y:=<c>.P}))" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8520
      proof (cases "fic P c")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8521
	case True 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8522
	assume "fic P c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8523
	then show ?thesis using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8524
	  apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8525
	  apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8526
	  apply(rule better_CutL_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8527
	  apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8528
	  apply(rule better_LAxR_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8529
	  apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8530
	  done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8531
      next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8532
	case False 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8533
	assume "\<not>fic P c" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8534
	then show ?thesis
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8535
	  apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8536
	  apply(rule a_star_CutL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8537
	  apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8538
	  apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8539
	  apply(rule ac_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8540
	  apply(rule better_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8541
	  apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8542
	  apply(simp add: subst_with_ax2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8543
	  done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8544
      qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8545
      also have "\<dots> = (Cut <b>.N (u).M){y:=<c>.P}" using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8546
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8547
	apply(auto simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8548
	apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8549
	apply(simp add: alpha fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8550
	apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8551
	apply(rule crename_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8552
	apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8553
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8554
      finally have "(Cut <a>.NotR (u).M a (v).NotL <b>.N v){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <b>.N (u).M){y:=<c>.P}" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8555
	by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8556
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8557
    ultimately show ?thesis by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8558
  qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8559
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8560
  case (LAnd1 b a1 M1 a2 M2 N z u)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8561
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8562
  proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8563
    { assume asm: "M1\<noteq>Ax y a1"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8564
      have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){y:=<c>.P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8565
        Cut <b>.AndR <a1>.(M1{y:=<c>.P}) <a2>.(M2{y:=<c>.P}) b (z).AndL1 (u).(N{y:=<c>.P}) z" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8566
	using prems by (simp add: subst_fresh abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8567
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a1>.(M1{y:=<c>.P}) (u).(N{y:=<c>.P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8568
	using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8569
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8570
	apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8571
	apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8572
	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8573
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8574
      also have "\<dots> = (Cut <a1>.M1 (u).N){y:=<c>.P}" using prems 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8575
	by (simp add: subst_fresh abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8576
      finally 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8577
      have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <a1>.M1 (u).N){y:=<c>.P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8578
	by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8579
    } 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8580
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8581
    { assume asm: "M1=Ax y a1"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8582
      have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){y:=<c>.P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8583
        Cut <b>.AndR <a1>.(M1{y:=<c>.P}) <a2>.(M2{y:=<c>.P}) b (z).AndL1 (u).(N{y:=<c>.P}) z" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8584
	using prems by (simp add: subst_fresh abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8585
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a1>.(M1{y:=<c>.P}) (u).(N{y:=<c>.P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8586
	using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8587
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8588
	apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8589
	apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8590
	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8591
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8592
      also have "\<dots> = Cut <a1>.(Cut <c>.P (y). Ax y a1) (u).(N{y:=<c>.P})" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8593
	using prems by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8594
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a1>.P[c\<turnstile>c>a1] (u).(N{y:=<c>.P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8595
      proof (cases "fic P c")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8596
	case True 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8597
	assume "fic P c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8598
	then show ?thesis using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8599
	  apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8600
	  apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8601
	  apply(rule better_CutL_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8602
	  apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8603
	  apply(rule better_LAxR_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8604
	  apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8605
	  done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8606
      next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8607
	case False 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8608
	assume "\<not>fic P c" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8609
	then show ?thesis
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8610
	  apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8611
	  apply(rule a_star_CutL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8612
	  apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8613
	  apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8614
	  apply(rule ac_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8615
	  apply(rule better_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8616
	  apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8617
	  apply(simp add: subst_with_ax2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8618
	  done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8619
      qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8620
      also have "\<dots> = (Cut <a1>.M1 (u).N){y:=<c>.P}" using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8621
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8622
	apply(auto simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8623
	apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8624
	apply(simp add: alpha fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8625
	apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8626
	apply(rule crename_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8627
	apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8628
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8629
      finally 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8630
      have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <a1>.M1 (u).N){y:=<c>.P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8631
	by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8632
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8633
    ultimately show ?thesis by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8634
  qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8635
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8636
  case (LAnd2 b a1 M1 a2 M2 N z u)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8637
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8638
  proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8639
    { assume asm: "M2\<noteq>Ax y a2"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8640
      have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){y:=<c>.P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8641
        Cut <b>.AndR <a1>.(M1{y:=<c>.P}) <a2>.(M2{y:=<c>.P}) b (z).AndL2 (u).(N{y:=<c>.P}) z" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8642
	using prems by (simp add: subst_fresh abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8643
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a2>.(M2{y:=<c>.P}) (u).(N{y:=<c>.P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8644
	using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8645
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8646
	apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8647
	apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8648
	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8649
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8650
      also have "\<dots> = (Cut <a2>.M2 (u).N){y:=<c>.P}" using prems 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8651
	by (simp add: subst_fresh abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8652
      finally 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8653
      have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <a2>.M2 (u).N){y:=<c>.P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8654
	by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8655
    } 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8656
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8657
    { assume asm: "M2=Ax y a2"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8658
      have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){y:=<c>.P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8659
        Cut <b>.AndR <a1>.(M1{y:=<c>.P}) <a2>.(M2{y:=<c>.P}) b (z).AndL2 (u).(N{y:=<c>.P}) z" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8660
	using prems by (simp add: subst_fresh abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8661
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a2>.(M2{y:=<c>.P}) (u).(N{y:=<c>.P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8662
	using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8663
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8664
	apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8665
	apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8666
	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8667
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8668
      also have "\<dots> = Cut <a2>.(Cut <c>.P (y). Ax y a2) (u).(N{y:=<c>.P})" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8669
	using prems by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8670
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a2>.P[c\<turnstile>c>a2] (u).(N{y:=<c>.P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8671
      proof (cases "fic P c")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8672
	case True 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8673
	assume "fic P c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8674
	then show ?thesis using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8675
	  apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8676
	  apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8677
	  apply(rule better_CutL_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8678
	  apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8679
	  apply(rule better_LAxR_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8680
	  apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8681
	  done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8682
      next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8683
	case False 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8684
	assume "\<not>fic P c" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8685
	then show ?thesis
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8686
	  apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8687
	  apply(rule a_star_CutL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8688
	  apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8689
	  apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8690
	  apply(rule ac_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8691
	  apply(rule better_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8692
	  apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8693
	  apply(simp add: subst_with_ax2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8694
	  done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8695
      qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8696
      also have "\<dots> = (Cut <a2>.M2 (u).N){y:=<c>.P}" using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8697
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8698
	apply(auto simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8699
	apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8700
	apply(simp add: alpha fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8701
	apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8702
	apply(rule crename_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8703
	apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8704
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8705
      finally 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8706
      have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <a2>.M2 (u).N){y:=<c>.P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8707
	by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8708
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8709
    ultimately show ?thesis by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8710
  qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8711
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8712
  case (LOr1 b a M N1 N2 z x1 x2 y c P)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8713
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8714
  proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8715
    { assume asm: "M\<noteq>Ax y a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8716
      have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8717
        Cut <b>.OrR1 <a>.(M{y:=<c>.P}) b (z).OrL (x1).(N1{y:=<c>.P}) (x2).(N2{y:=<c>.P}) z" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8718
	using prems by (simp add: subst_fresh abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8719
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{y:=<c>.P}) (x1).(N1{y:=<c>.P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8720
	using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8721
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8722
	apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8723
	apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8724
	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8725
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8726
      also have "\<dots> = (Cut <a>.M (x1).N1){y:=<c>.P}" using prems 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8727
	by (simp add: subst_fresh abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8728
      finally 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8729
      have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <a>.M (x1).N1){y:=<c>.P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8730
	by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8731
    } 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8732
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8733
    { assume asm: "M=Ax y a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8734
      have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8735
        Cut <b>.OrR1 <a>.(M{y:=<c>.P}) b (z).OrL (x1).(N1{y:=<c>.P}) (x2).(N2{y:=<c>.P}) z" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8736
	using prems by (simp add: subst_fresh abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8737
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{y:=<c>.P}) (x1).(N1{y:=<c>.P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8738
	using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8739
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8740
	apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8741
	apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8742
	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8743
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8744
      also have "\<dots> = Cut <a>.(Cut <c>.P (y). Ax y a) (x1).(N1{y:=<c>.P})" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8745
	using prems by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8746
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.P[c\<turnstile>c>a] (x1).(N1{y:=<c>.P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8747
      proof (cases "fic P c")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8748
	case True 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8749
	assume "fic P c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8750
	then show ?thesis using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8751
	  apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8752
	  apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8753
	  apply(rule better_CutL_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8754
	  apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8755
	  apply(rule better_LAxR_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8756
	  apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8757
	  done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8758
      next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8759
	case False 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8760
	assume "\<not>fic P c" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8761
	then show ?thesis
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8762
	  apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8763
	  apply(rule a_star_CutL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8764
	  apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8765
	  apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8766
	  apply(rule ac_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8767
	  apply(rule better_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8768
	  apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8769
	  apply(simp add: subst_with_ax2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8770
	  done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8771
      qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8772
      also have "\<dots> = (Cut <a>.M (x1).N1){y:=<c>.P}" using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8773
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8774
	apply(auto simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8775
	apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8776
	apply(simp add: alpha fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8777
	apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8778
	apply(rule crename_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8779
	apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8780
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8781
      finally 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8782
      have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <a>.M (x1).N1){y:=<c>.P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8783
	by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8784
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8785
    ultimately show ?thesis by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8786
  qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8787
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8788
  case (LOr2 b a M N1 N2 z x1 x2 y c P)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8789
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8790
  proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8791
    { assume asm: "M\<noteq>Ax y a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8792
      have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8793
        Cut <b>.OrR2 <a>.(M{y:=<c>.P}) b (z).OrL (x1).(N1{y:=<c>.P}) (x2).(N2{y:=<c>.P}) z" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8794
	using prems by (simp add: subst_fresh abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8795
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{y:=<c>.P}) (x2).(N2{y:=<c>.P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8796
	using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8797
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8798
	apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8799
	apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8800
	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8801
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8802
      also have "\<dots> = (Cut <a>.M (x2).N2){y:=<c>.P}" using prems 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8803
	by (simp add: subst_fresh abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8804
      finally 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8805
      have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <a>.M (x2).N2){y:=<c>.P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8806
	by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8807
    } 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8808
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8809
    { assume asm: "M=Ax y a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8810
      have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8811
        Cut <b>.OrR2 <a>.(M{y:=<c>.P}) b (z).OrL (x1).(N1{y:=<c>.P}) (x2).(N2{y:=<c>.P}) z" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8812
	using prems by (simp add: subst_fresh abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8813
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{y:=<c>.P}) (x2).(N2{y:=<c>.P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8814
	using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8815
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8816
	apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8817
	apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8818
	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8819
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8820
      also have "\<dots> = Cut <a>.(Cut <c>.P (y). Ax y a) (x2).(N2{y:=<c>.P})" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8821
	using prems by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8822
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.P[c\<turnstile>c>a] (x2).(N2{y:=<c>.P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8823
      proof (cases "fic P c")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8824
	case True 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8825
	assume "fic P c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8826
	then show ?thesis using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8827
	  apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8828
	  apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8829
	  apply(rule better_CutL_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8830
	  apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8831
	  apply(rule better_LAxR_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8832
	  apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8833
	  done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8834
      next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8835
	case False 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8836
	assume "\<not>fic P c" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8837
	then show ?thesis
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8838
	  apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8839
	  apply(rule a_star_CutL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8840
	  apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8841
	  apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8842
	  apply(rule ac_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8843
	  apply(rule better_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8844
	  apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8845
	  apply(simp add: subst_with_ax2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8846
	  done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8847
      qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8848
      also have "\<dots> = (Cut <a>.M (x2).N2){y:=<c>.P}" using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8849
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8850
	apply(auto simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8851
	apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8852
	apply(simp add: alpha fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8853
	apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8854
	apply(rule crename_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8855
	apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8856
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8857
      finally 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8858
      have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <a>.M (x2).N2){y:=<c>.P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8859
	by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8860
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8861
    ultimately show ?thesis by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8862
  qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8863
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8864
  case (LImp z N u Q x M b a d y c P)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8865
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8866
  proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8867
    { assume asm: "N\<noteq>Ax y d"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8868
      have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){y:=<c>.P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8869
        Cut <b>.ImpR (x).<a>.(M{y:=<c>.P}) b (z).ImpL <d>.(N{y:=<c>.P}) (u).(Q{y:=<c>.P}) z" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8870
	using prems by (simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8871
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{y:=<c>.P})  (x).(M{y:=<c>.P})) (u).(Q{y:=<c>.P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8872
	using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8873
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8874
	apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8875
	apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8876
	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8877
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8878
      also have "\<dots> = (Cut <a>.(Cut <d>.N  (x).M) (u).Q){y:=<c>.P}" using prems 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8879
	by (simp add: subst_fresh abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8880
      finally 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8881
      have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){y:=<c>.P} \<longrightarrow>\<^isub>a* 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8882
                     (Cut <a>.(Cut <d>.N  (x).M) (u).Q){y:=<c>.P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8883
	by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8884
    } 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8885
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8886
    { assume asm: "N=Ax y d"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8887
      have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){y:=<c>.P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8888
        Cut <b>.ImpR (x).<a>.(M{y:=<c>.P}) b (z).ImpL <d>.(N{y:=<c>.P}) (u).(Q{y:=<c>.P}) z" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8889
	using prems by (simp add: subst_fresh abs_fresh fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8890
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{y:=<c>.P})  (x).(M{y:=<c>.P})) (u).(Q{y:=<c>.P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8891
	using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8892
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8893
	apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8894
	apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8895
	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8896
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8897
      also have "\<dots> = Cut <a>.(Cut <d>.(Cut <c>.P (y).Ax y d)  (x).(M{y:=<c>.P})) (u).(Q{y:=<c>.P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8898
	using prems by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8899
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(P[c\<turnstile>c>d]) (x).(M{y:=<c>.P})) (u).(Q{y:=<c>.P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8900
      proof (cases "fic P c")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8901
	case True 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8902
	assume "fic P c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8903
	then show ?thesis using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8904
	  apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8905
	  apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8906
	  apply(rule better_CutL_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8907
	  apply(rule a_Cut_l)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8908
	  apply(simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8909
	  apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8910
	  apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8911
	  apply(rule better_LAxR_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8912
	  apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8913
	  done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8914
      next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8915
	case False 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8916
	assume "\<not>fic P c" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8917
	then show ?thesis using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8918
	  apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8919
	  apply(rule a_star_CutL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8920
	  apply(rule a_star_CutL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8921
	  apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8922
	  apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8923
	  apply(rule ac_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8924
	  apply(rule better_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8925
	  apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8926
	  apply(simp add: subst_with_ax2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8927
	  done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8928
      qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8929
      also have "\<dots> = (Cut <a>.(Cut <d>.N (x).M) (u).Q){y:=<c>.P}" using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8930
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8931
	apply(auto simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8932
	apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8933
	apply(simp add: alpha fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8934
	apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8935
	apply(simp add: alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8936
	apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8937
	apply(rule crename_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8938
	apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8939
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8940
      finally 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8941
      have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){y:=<c>.P} \<longrightarrow>\<^isub>a* 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8942
	       (Cut <a>.(Cut <d>.N (x).M) (u).Q){y:=<c>.P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8943
	by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8944
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8945
    ultimately show ?thesis by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8946
  qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8947
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8948
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8949
lemma l_redu_subst2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8950
  assumes a: "M \<longrightarrow>\<^isub>l M'" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8951
  shows "M{c:=(y).P} \<longrightarrow>\<^isub>a* M'{c:=(y).P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8952
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8953
proof(nominal_induct M M' avoiding: y c P rule: l_redu.strong_induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8954
  case LAxR
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8955
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8956
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8957
    apply(rule aux3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8958
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8959
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8960
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8961
    apply(simp add: trm.inject fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8962
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8963
    apply(rule aux4)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8964
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8965
    apply(rule fic_substc_crename)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8966
    apply(simp_all)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8967
    apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8968
    apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8969
    apply(rule aux2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8970
    apply(rule l_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8971
    apply(simp add: subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8972
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8973
    apply(rule fic_subst1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8974
    apply(simp_all)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8975
    apply(rule subst_comm')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8976
    apply(simp_all)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8977
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8978
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8979
  case LAxL
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8980
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8981
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8982
    apply(rule aux3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8983
    apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8984
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8985
    apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8986
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8987
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8988
    apply(rule aux4)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8989
    apply(simp add: trm.inject alpha calc_atm fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8990
    apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8991
    apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8992
    apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8993
    apply(rule l_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8994
    apply(simp add: subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8995
    apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8996
    apply(rule fin_subst2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8997
    apply(simp_all)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8998
    apply(rule aux4)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  8999
    apply(rule subst_comm')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9000
    apply(simp_all)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9001
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9002
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9003
  case (LNot v M N u a b)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9004
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9005
  proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9006
    { assume asm: "M\<noteq>Ax u c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9007
      have "(Cut <a>.NotR (u).M a (v).NotL <b>.N v){c:=(y).P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9008
        (Cut <a>.NotR (u).(M{c:=(y).P}) a (v).NotL <b>.(N{c:=(y).P}) v)" using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9009
	by (simp add: subst_fresh abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9010
      also have "\<dots> \<longrightarrow>\<^isub>l (Cut <b>.(N{c:=(y).P}) (u).(M{c:=(y).P}))" using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9011
	by (auto intro: l_redu.intros simp add: subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9012
      also have "\<dots> = (Cut <b>.N (u).M){c:=(y).P}" using prems 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9013
	by (simp add: subst_fresh abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9014
      finally have ?thesis by auto
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9015
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9016
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9017
    { assume asm: "M=Ax u c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9018
      have "(Cut <a>.NotR (u).M a (v).NotL <b>.N v){c:=(y).P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9019
        (Cut <a>.NotR (u).(M{c:=(y).P}) a (v).NotL <b>.(N{c:=(y).P}) v)" using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9020
	by (simp add: subst_fresh abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9021
      also have "\<dots> \<longrightarrow>\<^isub>a* (Cut <b>.(N{c:=(y).P}) (u).(M{c:=(y).P}))" using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9022
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9023
	apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9024
	apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9025
	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9026
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9027
      also have "\<dots> = (Cut <b>.(N{c:=(y).P}) (u).(Cut <c>.(Ax u c) (y).P))" using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9028
	by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9029
      also have "\<dots> \<longrightarrow>\<^isub>a* (Cut <b>.(N{c:=(y).P})  (u).(P[y\<turnstile>n>u]))" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9030
      proof (cases "fin P y")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9031
	case True 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9032
	assume "fin P y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9033
	then show ?thesis using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9034
	  apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9035
	  apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9036
	  apply(rule better_CutR_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9037
	  apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9038
	  apply(rule better_LAxL_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9039
	  apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9040
	  done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9041
      next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9042
	case False 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9043
	assume "\<not>fin P y" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9044
	then show ?thesis
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9045
	  apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9046
	  apply(rule a_star_CutR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9047
	  apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9048
	  apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9049
	  apply(rule ac_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9050
	  apply(rule better_right)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9051
	  apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9052
	  apply(simp add: subst_with_ax1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9053
	  done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9054
      qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9055
      also have "\<dots> = (Cut <b>.N (u).M){c:=(y).P}" using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9056
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9057
	apply(auto simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9058
	apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9059
	apply(simp add: alpha fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9060
	apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9061
	apply(rule nrename_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9062
	apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9063
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9064
      finally have "(Cut <a>.NotR (u).M a (v).NotL <b>.N v){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <b>.N (u).M){c:=(y).P}" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9065
	by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9066
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9067
    ultimately show ?thesis by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9068
  qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9069
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9070
  case (LAnd1 b a1 M1 a2 M2 N z u)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9071
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9072
  proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9073
    { assume asm: "N\<noteq>Ax u c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9074
      have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){c:=(y).P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9075
        Cut <b>.AndR <a1>.(M1{c:=(y).P}) <a2>.(M2{c:=(y).P}) b (z).AndL1 (u).(N{c:=(y).P}) z" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9076
	using prems by (simp add: subst_fresh abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9077
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a1>.(M1{c:=(y).P}) (u).(N{c:=(y).P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9078
	using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9079
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9080
	apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9081
	apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9082
	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9083
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9084
      also have "\<dots> = (Cut <a1>.M1 (u).N){c:=(y).P}" using prems 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9085
	by (simp add: subst_fresh abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9086
      finally 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9087
      have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <a1>.M1 (u).N){c:=(y).P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9088
	by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9089
    } 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9090
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9091
    { assume asm: "N=Ax u c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9092
      have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){c:=(y).P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9093
        Cut <b>.AndR <a1>.(M1{c:=(y).P}) <a2>.(M2{c:=(y).P}) b (z).AndL1 (u).(N{c:=(y).P}) z" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9094
	using prems by (simp add: subst_fresh abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9095
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a1>.(M1{c:=(y).P}) (u).(N{c:=(y).P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9096
	using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9097
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9098
	apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9099
	apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9100
	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9101
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9102
      also have "\<dots> = Cut <a1>.(M1{c:=(y).P}) (u).(Cut <c>.(Ax u c) (y).P)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9103
	using prems by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9104
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a1>.(M1{c:=(y).P}) (u).(P[y\<turnstile>n>u])"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9105
      proof (cases "fin P y")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9106
	case True 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9107
	assume "fin P y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9108
	then show ?thesis using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9109
	  apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9110
	  apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9111
	  apply(rule better_CutR_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9112
	  apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9113
	  apply(rule better_LAxL_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9114
	  apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9115
	  done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9116
      next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9117
	case False 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9118
	assume "\<not>fin P y" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9119
	then show ?thesis
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9120
	  apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9121
	  apply(rule a_star_CutR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9122
	  apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9123
	  apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9124
	  apply(rule ac_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9125
	  apply(rule better_right)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9126
	  apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9127
	  apply(simp add: subst_with_ax1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9128
	  done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9129
      qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9130
      also have "\<dots> = (Cut <a1>.M1 (u).N){c:=(y).P}" using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9131
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9132
	apply(auto simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9133
	apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9134
	apply(simp add: alpha fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9135
	apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9136
	apply(rule nrename_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9137
	apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9138
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9139
      finally 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9140
      have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <a1>.M1 (u).N){c:=(y).P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9141
	by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9142
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9143
    ultimately show ?thesis by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9144
  qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9145
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9146
  case (LAnd2 b a1 M1 a2 M2 N z u)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9147
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9148
  proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9149
    { assume asm: "N\<noteq>Ax u c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9150
      have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){c:=(y).P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9151
        Cut <b>.AndR <a1>.(M1{c:=(y).P}) <a2>.(M2{c:=(y).P}) b (z).AndL2 (u).(N{c:=(y).P}) z" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9152
	using prems by (simp add: subst_fresh abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9153
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a2>.(M2{c:=(y).P}) (u).(N{c:=(y).P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9154
	using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9155
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9156
	apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9157
	apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9158
	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9159
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9160
      also have "\<dots> = (Cut <a2>.M2 (u).N){c:=(y).P}" using prems 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9161
	by (simp add: subst_fresh abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9162
      finally 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9163
      have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <a2>.M2 (u).N){c:=(y).P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9164
	by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9165
    } 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9166
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9167
    { assume asm: "N=Ax u c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9168
      have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){c:=(y).P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9169
        Cut <b>.AndR <a1>.(M1{c:=(y).P}) <a2>.(M2{c:=(y).P}) b (z).AndL2 (u).(N{c:=(y).P}) z" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9170
	using prems by (simp add: subst_fresh abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9171
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a2>.(M2{c:=(y).P}) (u).(N{c:=(y).P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9172
	using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9173
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9174
	apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9175
	apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9176
	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9177
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9178
      also have "\<dots> = Cut <a2>.(M2{c:=(y).P}) (u).(Cut <c>.(Ax u c) (y).P)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9179
	using prems by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9180
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a2>.(M2{c:=(y).P}) (u).(P[y\<turnstile>n>u])"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9181
      proof (cases "fin P y")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9182
	case True 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9183
	assume "fin P y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9184
	then show ?thesis using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9185
	  apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9186
	  apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9187
	  apply(rule better_CutR_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9188
	  apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9189
	  apply(rule better_LAxL_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9190
	  apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9191
	  done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9192
      next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9193
	case False 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9194
	assume "\<not>fin P y" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9195
	then show ?thesis
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9196
	  apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9197
	  apply(rule a_star_CutR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9198
	  apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9199
	  apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9200
	  apply(rule ac_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9201
	  apply(rule better_right)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9202
	  apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9203
	  apply(simp add: subst_with_ax1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9204
	  done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9205
      qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9206
      also have "\<dots> = (Cut <a2>.M2 (u).N){c:=(y).P}" using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9207
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9208
	apply(auto simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9209
	apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9210
	apply(simp add: alpha fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9211
	apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9212
	apply(rule nrename_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9213
	apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9214
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9215
      finally 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9216
      have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <a2>.M2 (u).N){c:=(y).P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9217
	by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9218
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9219
    ultimately show ?thesis by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9220
  qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9221
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9222
  case (LOr1 b a M N1 N2 z x1 x2 y c P)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9223
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9224
  proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9225
    { assume asm: "N1\<noteq>Ax x1 c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9226
      have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9227
        Cut <b>.OrR1 <a>.(M{c:=(y).P}) b (z).OrL (x1).(N1{c:=(y).P}) (x2).(N2{c:=(y).P}) z" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9228
	using prems by (simp add: subst_fresh abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9229
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{c:=(y).P}) (x1).(N1{c:=(y).P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9230
	using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9231
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9232
	apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9233
	apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9234
	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9235
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9236
      also have "\<dots> = (Cut <a>.M (x1).N1){c:=(y).P}" using prems 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9237
	by (simp add: subst_fresh abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9238
      finally 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9239
      have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <a>.M (x1).N1){c:=(y).P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9240
	by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9241
    } 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9242
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9243
    { assume asm: "N1=Ax x1 c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9244
      have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9245
        Cut <b>.OrR1 <a>.(M{c:=(y).P}) b (z).OrL (x1).(N1{c:=(y).P}) (x2).(N2{c:=(y).P}) z" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9246
	using prems by (simp add: subst_fresh abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9247
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{c:=(y).P}) (x1).(N1{c:=(y).P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9248
	using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9249
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9250
	apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9251
	apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9252
	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9253
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9254
      also have "\<dots> = Cut <a>.(M{c:=(y).P}) (x1).(Cut <c>.(Ax x1 c) (y).P)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9255
	using prems by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9256
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{c:=(y).P})   (x1).(P[y\<turnstile>n>x1])"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9257
      proof (cases "fin P y")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9258
	case True 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9259
	assume "fin P y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9260
	then show ?thesis using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9261
	  apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9262
	  apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9263
	  apply(rule better_CutR_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9264
	  apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9265
	  apply(rule better_LAxL_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9266
	  apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9267
	  done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9268
      next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9269
	case False 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9270
	assume "\<not>fin P y" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9271
	then show ?thesis
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9272
	  apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9273
	  apply(rule a_star_CutR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9274
	  apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9275
	  apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9276
	  apply(rule ac_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9277
	  apply(rule better_right)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9278
	  apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9279
	  apply(simp add: subst_with_ax1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9280
	  done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9281
      qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9282
      also have "\<dots> = (Cut <a>.M (x1).N1){c:=(y).P}" using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9283
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9284
	apply(auto simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9285
	apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9286
	apply(simp add: alpha fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9287
	apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9288
	apply(rule nrename_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9289
	apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9290
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9291
      finally 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9292
      have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <a>.M (x1).N1){c:=(y).P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9293
	by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9294
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9295
    ultimately show ?thesis by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9296
  qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9297
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9298
  case (LOr2 b a M N1 N2 z x1 x2 y c P)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9299
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9300
  proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9301
    { assume asm: "N2\<noteq>Ax x2 c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9302
      have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9303
        Cut <b>.OrR2 <a>.(M{c:=(y).P}) b (z).OrL (x1).(N1{c:=(y).P}) (x2).(N2{c:=(y).P}) z" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9304
	using prems by (simp add: subst_fresh abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9305
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{c:=(y).P}) (x2).(N2{c:=(y).P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9306
	using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9307
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9308
	apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9309
	apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9310
	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9311
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9312
      also have "\<dots> = (Cut <a>.M (x2).N2){c:=(y).P}" using prems 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9313
	by (simp add: subst_fresh abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9314
      finally 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9315
      have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <a>.M (x2).N2){c:=(y).P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9316
	by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9317
    } 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9318
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9319
    { assume asm: "N2=Ax x2 c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9320
      have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9321
        Cut <b>.OrR2 <a>.(M{c:=(y).P}) b (z).OrL (x1).(N1{c:=(y).P}) (x2).(N2{c:=(y).P}) z" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9322
	using prems by (simp add: subst_fresh abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9323
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{c:=(y).P}) (x2).(N2{c:=(y).P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9324
	using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9325
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9326
	apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9327
	apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9328
	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9329
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9330
      also have "\<dots> = Cut <a>.(M{c:=(y).P}) (x2).(Cut <c>.(Ax x2 c) (y).P)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9331
	using prems by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9332
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{c:=(y).P}) (x2).(P[y\<turnstile>n>x2])"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9333
      proof (cases "fin P y")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9334
	case True 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9335
	assume "fin P y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9336
	then show ?thesis using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9337
	  apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9338
	  apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9339
	  apply(rule better_CutR_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9340
	  apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9341
	  apply(rule better_LAxL_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9342
	  apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9343
	  done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9344
      next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9345
	case False 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9346
	assume "\<not>fin P y" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9347
	then show ?thesis
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9348
	  apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9349
	  apply(rule a_star_CutR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9350
	  apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9351
	  apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9352
	  apply(rule ac_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9353
	  apply(rule better_right)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9354
	  apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9355
	  apply(simp add: subst_with_ax1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9356
	  done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9357
      qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9358
      also have "\<dots> = (Cut <a>.M (x2).N2){c:=(y).P}" using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9359
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9360
	apply(auto simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9361
	apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9362
	apply(simp add: alpha fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9363
	apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9364
	apply(rule nrename_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9365
	apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9366
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9367
      finally 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9368
      have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <a>.M (x2).N2){c:=(y).P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9369
	by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9370
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9371
    ultimately show ?thesis by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9372
  qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9373
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9374
  case (LImp z N u Q x M b a d y c P)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9375
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9376
  proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9377
    { assume asm: "M\<noteq>Ax x c \<and> Q\<noteq>Ax u c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9378
      have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9379
        Cut <b>.ImpR (x).<a>.(M{c:=(y).P}) b (z).ImpL <d>.(N{c:=(y).P}) (u).(Q{c:=(y).P}) z" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9380
	using prems by (simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9381
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{c:=(y).P})  (x).(M{c:=(y).P})) (u).(Q{c:=(y).P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9382
	using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9383
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9384
	apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9385
	apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9386
	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9387
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9388
      also have "\<dots> = (Cut <a>.(Cut <d>.N  (x).M) (u).Q){c:=(y).P}" using prems 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9389
	by (simp add: subst_fresh abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9390
      finally 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9391
      have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} \<longrightarrow>\<^isub>a* 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9392
                     (Cut <a>.(Cut <d>.N  (x).M) (u).Q){c:=(y).P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9393
	by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9394
    } 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9395
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9396
    { assume asm: "M=Ax x c \<and> Q\<noteq>Ax u c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9397
      have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9398
        Cut <b>.ImpR (x).<a>.(M{c:=(y).P}) b (z).ImpL <d>.(N{c:=(y).P}) (u).(Q{c:=(y).P}) z" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9399
	using prems by (simp add: subst_fresh abs_fresh fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9400
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{c:=(y).P})  (x).(M{c:=(y).P})) (u).(Q{c:=(y).P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9401
	using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9402
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9403
	apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9404
	apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9405
	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9406
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9407
      also have "\<dots> = Cut <a>.(Cut <d>.(N{c:=(y).P})  (x).(Cut <c>.Ax x c (y).P)) (u).(Q{c:=(y).P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9408
	using prems by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9409
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{c:=(y).P})  (x).(P[y\<turnstile>n>x])) (u).(Q{c:=(y).P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9410
      proof (cases "fin P y")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9411
	case True 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9412
	assume "fin P y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9413
	then show ?thesis using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9414
	  apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9415
	  apply(rule a_star_CutL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9416
	  apply(rule a_star_CutR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9417
	  apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9418
	  apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9419
	  apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9420
	  apply(rule better_LAxL_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9421
	  apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9422
	  apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9423
	  done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9424
      next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9425
	case False 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9426
	assume "\<not>fin P y" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9427
	then show ?thesis using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9428
	  apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9429
	  apply(rule a_star_CutL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9430
	  apply(rule a_star_CutR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9431
	  apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9432
	  apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9433
	  apply(rule ac_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9434
	  apply(rule better_right)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9435
	  apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9436
	  apply(simp add: subst_with_ax1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9437
	  done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9438
      qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9439
      also have "\<dots> = (Cut <a>.(Cut <d>.N (x).M) (u).Q){c:=(y).P}" using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9440
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9441
	apply(auto simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9442
	apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9443
	apply(simp add: alpha fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9444
	apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9445
	apply(simp add: alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9446
	apply(simp add: nrename_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9447
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9448
      finally 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9449
      have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} \<longrightarrow>\<^isub>a* 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9450
	       (Cut <a>.(Cut <d>.N (x).M) (u).Q){c:=(y).P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9451
	by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9452
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9453
     moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9454
    { assume asm: "M\<noteq>Ax x c \<and> Q=Ax u c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9455
      have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9456
        Cut <b>.ImpR (x).<a>.(M{c:=(y).P}) b (z).ImpL <d>.(N{c:=(y).P}) (u).(Q{c:=(y).P}) z" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9457
	using prems by (simp add: subst_fresh abs_fresh fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9458
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{c:=(y).P})  (x).(M{c:=(y).P})) (u).(Q{c:=(y).P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9459
	using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9460
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9461
	apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9462
	apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9463
	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9464
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9465
      also have "\<dots> = Cut <a>.(Cut <d>.(N{c:=(y).P})  (x).(M{c:=(y).P})) (u).(Cut <c>.Ax u c (y).P)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9466
	using prems by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9467
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{c:=(y).P})  (x).(M{c:=(y).P})) (u).(P[y\<turnstile>n>u])"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9468
      proof (cases "fin P y")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9469
	case True 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9470
	assume "fin P y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9471
	then show ?thesis using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9472
	  apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9473
	  apply(rule a_star_CutR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9474
	  apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9475
	  apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9476
	  apply(rule better_LAxL_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9477
	  apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9478
	  done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9479
      next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9480
	case False 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9481
	assume "\<not>fin P y" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9482
	then show ?thesis using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9483
	  apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9484
	  apply(rule a_star_CutR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9485
	  apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9486
	  apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9487
	  apply(rule ac_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9488
	  apply(rule better_right)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9489
	  apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9490
	  apply(simp add: subst_with_ax1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9491
	  done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9492
      qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9493
      also have "\<dots> = (Cut <a>.(Cut <d>.N (x).M) (u).Q){c:=(y).P}" using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9494
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9495
	apply(auto simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9496
	apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9497
	apply(simp add: alpha fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9498
	apply(simp add: nrename_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9499
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9500
      finally 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9501
      have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} \<longrightarrow>\<^isub>a* 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9502
	       (Cut <a>.(Cut <d>.N (x).M) (u).Q){c:=(y).P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9503
	by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9504
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9505
     moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9506
    { assume asm: "M=Ax x c \<and> Q=Ax u c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9507
      have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9508
        Cut <b>.ImpR (x).<a>.(M{c:=(y).P}) b (z).ImpL <d>.(N{c:=(y).P}) (u).(Q{c:=(y).P}) z" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9509
	using prems by (simp add: subst_fresh abs_fresh fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9510
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{c:=(y).P})  (x).(M{c:=(y).P})) (u).(Q{c:=(y).P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9511
	using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9512
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9513
	apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9514
	apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9515
	apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9516
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9517
      also have "\<dots> = Cut <a>.(Cut <d>.(N{c:=(y).P})  (x).(Cut <c>.Ax x c (y).P)) (u).(Cut <c>.Ax u c (y).P)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9518
	using prems by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9519
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{c:=(y).P})  (x).(Cut <c>.Ax x c (y).P)) (u).(P[y\<turnstile>n>u])"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9520
      proof (cases "fin P y")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9521
	case True 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9522
	assume "fin P y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9523
	then show ?thesis using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9524
	  apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9525
	  apply(rule a_star_CutR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9526
	  apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9527
	  apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9528
	  apply(rule better_LAxL_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9529
	  apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9530
	  done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9531
      next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9532
	case False 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9533
	assume "\<not>fin P y" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9534
	then show ?thesis using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9535
	  apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9536
	  apply(rule a_star_CutR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9537
	  apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9538
	  apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9539
	  apply(rule ac_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9540
	  apply(rule better_right)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9541
	  apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9542
	  apply(simp add: subst_with_ax1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9543
	  done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9544
      qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9545
      also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{c:=(y).P})  (x).(P[y\<turnstile>n>x])) (u).(P[y\<turnstile>n>u])"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9546
      proof (cases "fin P y")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9547
	case True 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9548
	assume "fin P y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9549
	then show ?thesis using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9550
	  apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9551
	  apply(rule a_star_CutL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9552
	  apply(rule a_star_CutR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9553
	  apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9554
	  apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9555
	  apply(rule better_LAxL_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9556
	  apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9557
	  done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9558
      next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9559
	case False 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9560
	assume "\<not>fin P y" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9561
	then show ?thesis using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9562
	  apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9563
	  apply(rule a_star_CutL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9564
	  apply(rule a_star_CutR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9565
	  apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9566
	  apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9567
	  apply(rule ac_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9568
	  apply(rule better_right)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9569
	  apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9570
	  apply(simp add: subst_with_ax1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9571
	  done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9572
      qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9573
      also have "\<dots> = (Cut <a>.(Cut <d>.N (x).M) (u).Q){c:=(y).P}" using prems
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9574
	apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9575
	apply(auto simp add: subst_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9576
	apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9577
	apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9578
	apply(simp add: alpha fresh_atm trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9579
	apply(simp add: nrename_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9580
	apply(simp add: alpha fresh_atm trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9581
	apply(simp add: nrename_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9582
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9583
      finally 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9584
      have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} \<longrightarrow>\<^isub>a* 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9585
	       (Cut <a>.(Cut <d>.N (x).M) (u).Q){c:=(y).P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9586
	by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9587
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9588
    ultimately show ?thesis by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9589
  qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9590
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9591
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9592
lemma a_redu_subst1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9593
  assumes a: "M \<longrightarrow>\<^isub>a M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9594
  shows "M{y:=<c>.P} \<longrightarrow>\<^isub>a* M'{y:=<c>.P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9595
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9596
proof(nominal_induct avoiding: y c P rule: a_redu.strong_induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9597
  case al_redu
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9598
  then show ?case by (simp only: l_redu_subst1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9599
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9600
  case ac_redu
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9601
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9602
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9603
    apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9604
    apply(rule a_redu.ac_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9605
    apply(simp only: c_redu_subst1')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9606
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9607
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9608
  case (a_Cut_l a N x M M' y c P)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9609
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9610
    apply(simp add: subst_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9611
    apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9612
    apply(rule impI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9613
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9614
    apply(drule ax_do_not_a_reduce)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9615
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9616
    apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9617
    apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9618
    apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9619
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9620
    apply(drule_tac x="y" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9621
    apply(drule_tac x="c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9622
    apply(drule_tac x="P" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9623
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9624
    apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9625
    apply(rule a_star_CutL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9626
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9627
    apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9628
    apply(rule_tac M'="P[c\<turnstile>c>a]" in a_star_CutL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9629
    apply(case_tac "fic P c")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9630
    apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9631
    apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9632
    apply(rule better_LAxR_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9633
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9634
    apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9635
    apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9636
    apply(rule ac_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9637
    apply(rule better_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9638
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9639
    apply(rule subst_with_ax2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9640
    apply(rule aux4)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9641
    apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9642
    apply(simp add: alpha fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9643
    apply(simp add: crename_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9644
    apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9645
    apply(rule a_star_CutL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9646
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9647
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9648
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9649
  case (a_Cut_r a N x M M' y c P)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9650
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9651
    apply(auto simp add: subst_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9652
    apply(rule a_star_CutR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9653
    apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9654
    apply(rule a_star_CutR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9655
    apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9656
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9657
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9658
  case a_NotL
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9659
  then show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9660
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9661
    apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9662
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9663
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9664
    apply(simp add: subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9665
    apply(rule a_star_CutR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9666
    apply(rule a_star_NotL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9667
    apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9668
    apply(rule a_star_NotL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9669
    apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9670
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9671
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9672
  case a_NotR
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9673
  then show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9674
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9675
    apply(rule a_star_NotR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9676
    apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9677
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9678
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9679
  case a_AndR_l
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9680
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9681
    apply(auto simp add: subst_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9682
    apply(rule a_star_AndR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9683
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9684
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9685
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9686
  case a_AndR_r
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9687
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9688
    apply(auto simp add: subst_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9689
    apply(rule a_star_AndR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9690
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9691
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9692
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9693
  case a_AndL1
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9694
  then show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9695
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9696
    apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9697
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9698
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9699
    apply(simp add: subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9700
    apply(rule a_star_CutR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9701
    apply(rule a_star_AndL1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9702
    apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9703
    apply(rule a_star_AndL1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9704
    apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9705
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9706
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9707
  case a_AndL2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9708
  then show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9709
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9710
    apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9711
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9712
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9713
    apply(simp add: subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9714
    apply(rule a_star_CutR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9715
    apply(rule a_star_AndL2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9716
    apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9717
    apply(rule a_star_AndL2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9718
    apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9719
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9720
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9721
  case a_OrR1
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9722
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9723
    apply(auto simp add: subst_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9724
    apply(rule a_star_OrR1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9725
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9726
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9727
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9728
  case a_OrR2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9729
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9730
    apply(auto simp add: subst_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9731
    apply(rule a_star_OrR2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9732
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9733
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9734
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9735
  case a_OrL_l
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9736
  then show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9737
    apply(auto simp add: subst_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9738
    apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9739
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9740
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9741
    apply(simp add: subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9742
    apply(rule a_star_CutR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9743
    apply(rule a_star_OrL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9744
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9745
    apply(rule a_star_OrL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9746
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9747
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9748
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9749
  case a_OrL_r
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9750
  then show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9751
    apply(auto simp add: subst_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9752
    apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9753
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9754
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9755
    apply(simp add: subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9756
    apply(rule a_star_CutR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9757
    apply(rule a_star_OrL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9758
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9759
    apply(rule a_star_OrL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9760
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9761
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9762
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9763
  case a_ImpR
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9764
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9765
    apply(auto simp add: subst_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9766
    apply(rule a_star_ImpR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9767
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9768
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9769
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9770
  case a_ImpL_r
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9771
  then show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9772
    apply(auto simp add: subst_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9773
    apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9774
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9775
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9776
    apply(simp add: subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9777
    apply(rule a_star_CutR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9778
    apply(rule a_star_ImpL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9779
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9780
    apply(rule a_star_ImpL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9781
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9782
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9783
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9784
  case a_ImpL_l
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9785
  then show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9786
    apply(auto simp add: subst_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9787
    apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9788
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9789
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9790
    apply(simp add: subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9791
    apply(rule a_star_CutR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9792
    apply(rule a_star_ImpL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9793
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9794
    apply(rule a_star_ImpL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9795
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9796
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9797
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9798
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9799
lemma a_redu_subst2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9800
  assumes a: "M \<longrightarrow>\<^isub>a M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9801
  shows "M{c:=(y).P} \<longrightarrow>\<^isub>a* M'{c:=(y).P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9802
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9803
proof(nominal_induct avoiding: y c P rule: a_redu.strong_induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9804
  case al_redu
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9805
  then show ?case by (simp only: l_redu_subst2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9806
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9807
  case ac_redu
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9808
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9809
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9810
    apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9811
    apply(rule a_redu.ac_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9812
    apply(simp only: c_redu_subst2')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9813
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9814
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9815
  case (a_Cut_r a N x M M' y c P)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9816
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9817
    apply(simp add: subst_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9818
    apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9819
    apply(rule impI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9820
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9821
    apply(drule ax_do_not_a_reduce)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9822
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9823
    apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9824
    apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9825
    apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9826
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9827
    apply(drule_tac x="c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9828
    apply(drule_tac x="y" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9829
    apply(drule_tac x="P" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9830
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9831
    apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9832
    apply(rule a_star_CutR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9833
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9834
    apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9835
    apply(rule_tac N'="P[y\<turnstile>n>x]" in a_star_CutR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9836
    apply(case_tac "fin P y")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9837
    apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9838
    apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9839
    apply(rule better_LAxL_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9840
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9841
    apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9842
    apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9843
    apply(rule ac_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9844
    apply(rule better_right)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9845
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9846
    apply(rule subst_with_ax1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9847
    apply(rule aux4)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9848
    apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9849
    apply(simp add: alpha fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9850
    apply(simp add: nrename_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9851
    apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9852
    apply(rule a_star_CutR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9853
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9854
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9855
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9856
  case (a_Cut_l a N x M M' y c P)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9857
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9858
    apply(auto simp add: subst_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9859
    apply(rule a_star_CutL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9860
    apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9861
    apply(rule a_star_CutL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9862
    apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9863
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9864
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9865
  case a_NotR
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9866
  then show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9867
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9868
    apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9869
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9870
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9871
    apply(simp add: subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9872
    apply(rule a_star_CutL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9873
    apply(rule a_star_NotR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9874
    apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9875
    apply(rule a_star_NotR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9876
    apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9877
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9878
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9879
  case a_NotL
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9880
  then show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9881
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9882
    apply(rule a_star_NotL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9883
    apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9884
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9885
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9886
  case a_AndR_l
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9887
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9888
    apply(auto simp add: subst_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9889
    apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9890
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9891
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9892
    apply(simp add: subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9893
    apply(rule a_star_CutL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9894
    apply(rule a_star_AndR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9895
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9896
    apply(rule a_star_AndR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9897
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9898
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9899
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9900
  case a_AndR_r
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9901
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9902
    apply(auto simp add: subst_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9903
    apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9904
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9905
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9906
    apply(simp add: subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9907
    apply(rule a_star_CutL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9908
    apply(rule a_star_AndR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9909
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9910
    apply(rule a_star_AndR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9911
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9912
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9913
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9914
  case a_AndL1
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9915
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9916
    apply(auto simp add: subst_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9917
    apply(rule a_star_AndL1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9918
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9919
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9920
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9921
  case a_AndL2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9922
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9923
    apply(auto simp add: subst_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9924
    apply(rule a_star_AndL2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9925
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9926
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9927
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9928
  case a_OrR1
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9929
  then show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9930
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9931
    apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9932
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9933
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9934
    apply(simp add: subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9935
    apply(rule a_star_CutL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9936
    apply(rule a_star_OrR1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9937
    apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9938
    apply(rule a_star_OrR1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9939
    apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9940
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9941
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9942
  case a_OrR2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9943
  then show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9944
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9945
    apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9946
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9947
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9948
    apply(simp add: subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9949
    apply(rule a_star_CutL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9950
    apply(rule a_star_OrR2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9951
    apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9952
    apply(rule a_star_OrR2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9953
    apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9954
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9955
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9956
  case a_OrL_l
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9957
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9958
    apply(auto simp add: subst_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9959
    apply(rule a_star_OrL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9960
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9961
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9962
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9963
  case a_OrL_r
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9964
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9965
    apply(auto simp add: subst_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9966
    apply(rule a_star_OrL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9967
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9968
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9969
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9970
  case a_ImpR
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9971
  then show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9972
    apply(auto simp add: subst_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9973
    apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9974
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9975
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9976
    apply(simp add: subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9977
    apply(rule a_star_CutL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9978
    apply(rule a_star_ImpR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9979
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9980
    apply(rule a_star_ImpR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9981
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9982
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9983
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9984
  case a_ImpL_l
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9985
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9986
    apply(auto simp add: subst_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9987
    apply(rule a_star_ImpL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9988
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9989
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9990
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9991
  case a_ImpL_r
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9992
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9993
    apply(auto simp add: subst_fresh fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9994
    apply(rule a_star_ImpL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9995
    apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9996
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9997
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9998
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
  9999
lemma a_star_subst1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10000
  assumes a: "M \<longrightarrow>\<^isub>a* M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10001
  shows "M{y:=<c>.P} \<longrightarrow>\<^isub>a* M'{y:=<c>.P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10002
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10003
apply(induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10004
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10005
apply(drule_tac y="y" and c="c" and P="P" in a_redu_subst1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10006
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10007
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10008
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10009
lemma a_star_subst2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10010
  assumes a: "M \<longrightarrow>\<^isub>a* M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10011
  shows "M{c:=(y).P} \<longrightarrow>\<^isub>a* M'{c:=(y).P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10012
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10013
apply(induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10014
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10015
apply(drule_tac y="y" and c="c" and P="P" in a_redu_subst2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10016
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10017
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10018
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10019
text {* Candidates and SN *}
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10020
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10021
text {* SNa *}
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10022
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10023
inductive2 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10024
  SNa :: "trm \<Rightarrow> bool"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10025
where
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10026
  SNaI: "(\<And>M'. M \<longrightarrow>\<^isub>a M' \<Longrightarrow> SNa M') \<Longrightarrow> SNa M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10027
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10028
lemma SNa_induct[consumes 1]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10029
  assumes major: "SNa M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10030
  assumes hyp: "\<And>M'. SNa M' \<Longrightarrow> (\<forall>M''. M'\<longrightarrow>\<^isub>a M'' \<longrightarrow> P M'' \<Longrightarrow> P M')"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10031
  shows "P M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10032
  apply (rule major[THEN SNa.induct])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10033
  apply (rule hyp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10034
  apply (rule SNaI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10035
  apply (blast)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10036
  done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10037
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10038
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10039
lemma double_SNa_aux:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10040
  assumes a_SNa: "SNa a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10041
  and b_SNa: "SNa b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10042
  and hyp: "\<And>x z.
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10043
    (\<And>y. x\<longrightarrow>\<^isub>a y \<Longrightarrow> SNa y) \<Longrightarrow>
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10044
    (\<And>y. x\<longrightarrow>\<^isub>a y \<Longrightarrow> P y z) \<Longrightarrow>
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10045
    (\<And>u. z\<longrightarrow>\<^isub>a u \<Longrightarrow> SNa u) \<Longrightarrow>
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10046
    (\<And>u. z\<longrightarrow>\<^isub>a u \<Longrightarrow> P x u) \<Longrightarrow> P x z"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10047
  shows "P a b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10048
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10049
  from a_SNa
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10050
  have r: "\<And>b. SNa b \<Longrightarrow> P a b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10051
  proof (induct a rule: SNa.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10052
    case (SNaI x)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10053
    note SNa' = this
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10054
    have "SNa b" .
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10055
    thus ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10056
    proof (induct b rule: SNa.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10057
      case (SNaI y)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10058
      show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10059
	apply (rule hyp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10060
	apply (erule SNa')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10061
	apply (erule SNa')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10062
	apply (rule SNa.SNaI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10063
	apply (erule SNaI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10064
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10065
    qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10066
  qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10067
  from b_SNa show ?thesis by (rule r)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10068
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10069
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10070
lemma double_SNa:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10071
  "\<lbrakk>SNa a; SNa b; \<forall>x z. ((\<forall>y. x\<longrightarrow>\<^isub>ay \<longrightarrow> P y z) \<and> (\<forall>u. z\<longrightarrow>\<^isub>a u \<longrightarrow> P x u)) \<longrightarrow> P x z\<rbrakk> \<Longrightarrow> P a b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10072
apply(rule_tac double_SNa_aux)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10073
apply(assumption)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10074
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10075
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10076
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10077
lemma a_preserves_SNa:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10078
  assumes a: "SNa M" "M\<longrightarrow>\<^isub>a M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10079
  shows "SNa M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10080
using a 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10081
by (erule_tac SNa.cases) (simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10082
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10083
lemma a_star_preserves_SNa:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10084
  assumes a: "SNa M" and b: "M\<longrightarrow>\<^isub>a* M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10085
  shows "SNa M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10086
using b a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10087
by (induct) (auto simp add: a_preserves_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10088
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10089
lemma Ax_in_SNa:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10090
  shows "SNa (Ax x a)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10091
apply(rule SNaI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10092
apply(erule a_redu.cases, auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10093
apply(erule l_redu.cases, auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10094
apply(erule c_redu.cases, auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10095
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10096
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10097
lemma NotL_in_SNa:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10098
  assumes a: "SNa M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10099
  shows "SNa (NotL <a>.M x)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10100
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10101
apply(induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10102
apply(rule SNaI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10103
apply(erule a_redu.cases, auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10104
apply(erule l_redu.cases, auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10105
apply(erule c_redu.cases, auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10106
apply(auto simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10107
apply(rotate_tac 1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10108
apply(drule_tac x="[(a,aa)]\<bullet>M'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10109
apply(simp add: a_redu.eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10110
apply(subgoal_tac "NotL <a>.([(a,aa)]\<bullet>M'a) x = NotL <aa>.M'a x")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10111
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10112
apply(simp add: trm.inject alpha fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10113
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10114
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10115
lemma NotR_in_SNa:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10116
  assumes a: "SNa M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10117
  shows "SNa (NotR (x).M a)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10118
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10119
apply(induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10120
apply(rule SNaI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10121
apply(erule a_redu.cases, auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10122
apply(erule l_redu.cases, auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10123
apply(erule c_redu.cases, auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10124
apply(auto simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10125
apply(rotate_tac 1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10126
apply(drule_tac x="[(x,xa)]\<bullet>M'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10127
apply(simp add: a_redu.eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10128
apply(rule_tac s="NotR (x).([(x,xa)]\<bullet>M'a) a" in subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10129
apply(simp add: trm.inject alpha fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10130
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10131
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10132
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10133
lemma AndL1_in_SNa:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10134
  assumes a: "SNa M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10135
  shows "SNa (AndL1 (x).M y)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10136
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10137
apply(induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10138
apply(rule SNaI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10139
apply(erule a_redu.cases, auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10140
apply(erule l_redu.cases, auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10141
apply(erule c_redu.cases, auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10142
apply(auto simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10143
apply(rotate_tac 1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10144
apply(drule_tac x="[(x,xa)]\<bullet>M'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10145
apply(simp add: a_redu.eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10146
apply(rule_tac s="AndL1 x.([(x,xa)]\<bullet>M'a) y" in subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10147
apply(simp add: trm.inject alpha fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10148
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10149
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10150
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10151
lemma AndL2_in_SNa:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10152
  assumes a: "SNa M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10153
  shows "SNa (AndL2 (x).M y)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10154
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10155
apply(induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10156
apply(rule SNaI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10157
apply(erule a_redu.cases, auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10158
apply(erule l_redu.cases, auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10159
apply(erule c_redu.cases, auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10160
apply(auto simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10161
apply(rotate_tac 1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10162
apply(drule_tac x="[(x,xa)]\<bullet>M'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10163
apply(simp add: a_redu.eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10164
apply(rule_tac s="AndL2 x.([(x,xa)]\<bullet>M'a) y" in subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10165
apply(simp add: trm.inject alpha fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10166
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10167
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10168
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10169
lemma OrR1_in_SNa:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10170
  assumes a: "SNa M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10171
  shows "SNa (OrR1 <a>.M b)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10172
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10173
apply(induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10174
apply(rule SNaI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10175
apply(erule a_redu.cases, auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10176
apply(erule l_redu.cases, auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10177
apply(erule c_redu.cases, auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10178
apply(auto simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10179
apply(rotate_tac 1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10180
apply(drule_tac x="[(a,aa)]\<bullet>M'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10181
apply(simp add: a_redu.eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10182
apply(rule_tac s="OrR1 <a>.([(a,aa)]\<bullet>M'a) b" in subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10183
apply(simp add: trm.inject alpha fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10184
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10185
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10186
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10187
lemma OrR2_in_SNa:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10188
  assumes a: "SNa M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10189
  shows "SNa (OrR2 <a>.M b)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10190
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10191
apply(induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10192
apply(rule SNaI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10193
apply(erule a_redu.cases, auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10194
apply(erule l_redu.cases, auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10195
apply(erule c_redu.cases, auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10196
apply(auto simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10197
apply(rotate_tac 1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10198
apply(drule_tac x="[(a,aa)]\<bullet>M'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10199
apply(simp add: a_redu.eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10200
apply(rule_tac s="OrR2 <a>.([(a,aa)]\<bullet>M'a) b" in subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10201
apply(simp add: trm.inject alpha fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10202
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10203
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10204
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10205
lemma ImpR_in_SNa:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10206
  assumes a: "SNa M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10207
  shows "SNa (ImpR (x).<a>.M b)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10208
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10209
apply(induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10210
apply(rule SNaI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10211
apply(erule a_redu.cases, auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10212
apply(erule l_redu.cases, auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10213
apply(erule c_redu.cases, auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10214
apply(auto simp add: trm.inject alpha abs_fresh abs_perm calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10215
apply(rotate_tac 1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10216
apply(drule_tac x="[(a,aa)]\<bullet>M'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10217
apply(simp add: a_redu.eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10218
apply(rule_tac s="ImpR (x).<a>.([(a,aa)]\<bullet>M'a) b" in subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10219
apply(simp add: trm.inject alpha fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10220
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10221
apply(rotate_tac 1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10222
apply(drule_tac x="[(x,xa)]\<bullet>M'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10223
apply(simp add: a_redu.eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10224
apply(rule_tac s="ImpR (x).<a>.([(x,xa)]\<bullet>M'a) b" in subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10225
apply(simp add: trm.inject alpha fresh_a_redu abs_fresh abs_perm calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10226
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10227
apply(rotate_tac 1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10228
apply(drule_tac x="[(a,aa)]\<bullet>[(x,xa)]\<bullet>M'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10229
apply(simp add: a_redu.eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10230
apply(rule_tac s="ImpR (x).<a>.([(a,aa)]\<bullet>[(x,xa)]\<bullet>M'a) b" in subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10231
apply(simp add: trm.inject alpha fresh_a_redu abs_fresh abs_perm calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10232
apply(simp add: fresh_left calc_atm fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10233
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10234
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10235
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10236
lemma AndR_in_SNa:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10237
  assumes a: "SNa M" "SNa N"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10238
  shows "SNa (AndR <a>.M <b>.N c)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10239
apply(rule_tac a="M" and b="N" in double_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10240
apply(rule a)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10241
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10242
apply(rule SNaI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10243
apply(drule a_redu_AndR_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10244
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10245
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10246
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10247
lemma OrL_in_SNa:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10248
  assumes a: "SNa M" "SNa N"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10249
  shows "SNa (OrL (x).M (y).N z)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10250
apply(rule_tac a="M" and b="N" in double_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10251
apply(rule a)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10252
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10253
apply(rule SNaI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10254
apply(drule a_redu_OrL_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10255
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10256
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10257
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10258
lemma ImpL_in_SNa:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10259
  assumes a: "SNa M" "SNa N"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10260
  shows "SNa (ImpL <a>.M (y).N z)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10261
apply(rule_tac a="M" and b="N" in double_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10262
apply(rule a)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10263
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10264
apply(rule SNaI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10265
apply(drule a_redu_ImpL_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10266
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10267
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10268
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10269
lemma SNa_eqvt:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10270
  fixes pi1::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10271
  and   pi2::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10272
  shows "SNa M \<Longrightarrow> SNa (pi1\<bullet>M)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10273
  and   "SNa M \<Longrightarrow> SNa (pi2\<bullet>M)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10274
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10275
apply(induct rule: SNa.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10276
apply(rule SNaI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10277
apply(drule_tac pi="(rev pi1)" in a_redu.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10278
apply(rotate_tac 1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10279
apply(drule_tac x="(rev pi1)\<bullet>M'" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10280
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10281
apply(induct rule: SNa.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10282
apply(rule SNaI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10283
apply(drule_tac pi="(rev pi2)" in a_redu.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10284
apply(rotate_tac 1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10285
apply(drule_tac x="(rev pi2)\<bullet>M'" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10286
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10287
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10288
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10289
text {* set operators *}
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10290
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10291
constdefs
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10292
  AXIOMSn::"ty \<Rightarrow> ntrm set"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10293
  "AXIOMSn B \<equiv> { (x):(Ax y b) | x y b. True }"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10294
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10295
  AXIOMSc::"ty \<Rightarrow> ctrm set"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10296
  "AXIOMSc B \<equiv> { <a>:(Ax y b) | a y b. True }"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10297
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10298
  BINDINGn::"ty \<Rightarrow> ctrm set \<Rightarrow> ntrm set"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10299
  "BINDINGn B X \<equiv> { (x):M | x M. \<forall>a P. <a>:P\<in>X \<longrightarrow> SNa (M{x:=<a>.P})}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10300
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10301
  BINDINGc::"ty \<Rightarrow> ntrm set \<Rightarrow> ctrm set"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10302
  "BINDINGc B X \<equiv> { <a>:M | a M. \<forall>x P. (x):P\<in>X \<longrightarrow> SNa (M{a:=(x).P})}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10303
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10304
lemma BINDINGn_decreasing:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10305
  shows "X\<subseteq>Y \<Longrightarrow> BINDINGn B Y \<subseteq> BINDINGn B X"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10306
by (simp add: BINDINGn_def) (blast) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10307
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10308
lemma BINDINGc_decreasing:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10309
  shows "X\<subseteq>Y \<Longrightarrow> BINDINGc B Y \<subseteq> BINDINGc B X"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10310
by (simp add: BINDINGc_def) (blast) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10311
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10312
consts
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10313
  NOTRIGHT::"ty \<Rightarrow> ntrm set \<Rightarrow> ctrm set"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10314
  
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10315
nominal_primrec
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10316
 "NOTRIGHT (NOT B) X = { <a>:NotR (x).M a | a x M. fic (NotR (x).M a) a \<and> (x):M \<in> X }"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10317
apply(rule TrueI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10318
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10319
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10320
lemma NOTRIGHT_eqvt_name:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10321
  fixes pi::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10322
  shows "(pi\<bullet>(NOTRIGHT (NOT B) X)) = NOTRIGHT (NOT B) (pi\<bullet>X)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10323
apply(auto simp add: perm_set_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10324
apply(rule_tac x="pi\<bullet>a" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10325
apply(rule_tac x="pi\<bullet>xb" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10326
apply(rule_tac x="pi\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10327
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10328
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10329
apply(drule_tac pi="pi" in fic.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10330
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10331
apply(rule_tac x="(xb):M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10332
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10333
apply(rule_tac x="(rev pi)\<bullet>(<a>:NotR (xa).M a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10334
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10335
apply(rule_tac x="(rev pi)\<bullet>a" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10336
apply(rule_tac x="(rev pi)\<bullet>xa" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10337
apply(rule_tac x="(rev pi)\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10338
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10339
apply(drule_tac pi="rev pi" in fic.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10340
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10341
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10342
apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10343
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10344
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10345
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10346
lemma NOTRIGHT_eqvt_coname:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10347
  fixes pi::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10348
  shows "(pi\<bullet>(NOTRIGHT (NOT B) X)) = NOTRIGHT (NOT B) (pi\<bullet>X)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10349
apply(auto simp add: perm_set_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10350
apply(rule_tac x="pi\<bullet>a" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10351
apply(rule_tac x="pi\<bullet>xb" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10352
apply(rule_tac x="pi\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10353
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10354
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10355
apply(drule_tac pi="pi" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10356
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10357
apply(rule_tac x="(xb):M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10358
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10359
apply(rule_tac x="<((rev pi)\<bullet>a)>:NotR ((rev pi)\<bullet>xa).((rev pi)\<bullet>M) ((rev pi)\<bullet>a)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10360
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10361
apply(rule_tac x="(rev pi)\<bullet>a" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10362
apply(rule_tac x="(rev pi)\<bullet>xa" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10363
apply(rule_tac x="(rev pi)\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10364
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10365
apply(drule_tac pi="rev pi" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10366
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10367
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10368
apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10369
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10370
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10371
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10372
consts
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10373
  NOTLEFT::"ty \<Rightarrow> ctrm set \<Rightarrow> ntrm set"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10374
  
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10375
nominal_primrec
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10376
 "NOTLEFT (NOT B) X = { (x):NotL <a>.M x | a x M. fin (NotL <a>.M x) x \<and> <a>:M \<in> X }"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10377
apply(rule TrueI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10378
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10379
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10380
lemma NOTLEFT_eqvt_name:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10381
  fixes pi::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10382
  shows "(pi\<bullet>(NOTLEFT (NOT B) X)) = NOTLEFT (NOT B) (pi\<bullet>X)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10383
apply(auto simp add: perm_set_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10384
apply(rule_tac x="pi\<bullet>a" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10385
apply(rule_tac x="pi\<bullet>xb" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10386
apply(rule_tac x="pi\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10387
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10388
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10389
apply(drule_tac pi="pi" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10390
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10391
apply(rule_tac x="<a>:M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10392
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10393
apply(rule_tac x="(((rev pi)\<bullet>xa)):NotL <((rev pi)\<bullet>a)>.((rev pi)\<bullet>M) ((rev pi)\<bullet>xa)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10394
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10395
apply(rule_tac x="(rev pi)\<bullet>a" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10396
apply(rule_tac x="(rev pi)\<bullet>xa" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10397
apply(rule_tac x="(rev pi)\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10398
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10399
apply(drule_tac pi="rev pi" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10400
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10401
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10402
apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10403
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10404
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10405
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10406
lemma NOTLEFT_eqvt_coname:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10407
  fixes pi::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10408
  shows "(pi\<bullet>(NOTLEFT (NOT B) X)) = NOTLEFT (NOT B) (pi\<bullet>X)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10409
apply(auto simp add: perm_set_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10410
apply(rule_tac x="pi\<bullet>a" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10411
apply(rule_tac x="pi\<bullet>xb" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10412
apply(rule_tac x="pi\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10413
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10414
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10415
apply(drule_tac pi="pi" in fin.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10416
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10417
apply(rule_tac x="<a>:M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10418
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10419
apply(rule_tac x="(((rev pi)\<bullet>xa)):NotL <((rev pi)\<bullet>a)>.((rev pi)\<bullet>M) ((rev pi)\<bullet>xa)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10420
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10421
apply(rule_tac x="(rev pi)\<bullet>a" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10422
apply(rule_tac x="(rev pi)\<bullet>xa" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10423
apply(rule_tac x="(rev pi)\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10424
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10425
apply(drule_tac pi="rev pi" in fin.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10426
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10427
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10428
apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10429
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10430
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10431
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10432
consts
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10433
  ANDRIGHT::"ty \<Rightarrow> ctrm set \<Rightarrow> ctrm set \<Rightarrow> ctrm set"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10434
  
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10435
nominal_primrec
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10436
 "ANDRIGHT (B AND C) X Y = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10437
            { <c>:AndR <a>.M <b>.N c | c a b M N. fic (AndR <a>.M <b>.N c) c \<and> <a>:M \<in> X \<and> <b>:N \<in> Y }"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10438
apply(rule TrueI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10439
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10440
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10441
lemma ANDRIGHT_eqvt_name:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10442
  fixes pi::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10443
  shows "(pi\<bullet>(ANDRIGHT (A AND B) X Y)) = ANDRIGHT (A AND B) (pi\<bullet>X) (pi\<bullet>Y)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10444
apply(auto simp add: perm_set_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10445
apply(rule_tac x="pi\<bullet>c" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10446
apply(rule_tac x="pi\<bullet>a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10447
apply(rule_tac x="pi\<bullet>b" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10448
apply(rule_tac x="pi\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10449
apply(rule_tac x="pi\<bullet>N" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10450
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10451
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10452
apply(drule_tac pi="pi" in fic.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10453
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10454
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10455
apply(rule_tac x="<a>:M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10456
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10457
apply(rule_tac x="<b>:N" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10458
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10459
apply(rule_tac x="(rev pi)\<bullet>(<c>:AndR <a>.M <b>.N c)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10460
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10461
apply(rule_tac x="(rev pi)\<bullet>c" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10462
apply(rule_tac x="(rev pi)\<bullet>a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10463
apply(rule_tac x="(rev pi)\<bullet>b" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10464
apply(rule_tac x="(rev pi)\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10465
apply(rule_tac x="(rev pi)\<bullet>N" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10466
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10467
apply(drule_tac pi="rev pi" in fic.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10468
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10469
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10470
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10471
apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10472
apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10473
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10474
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10475
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10476
lemma ANDRIGHT_eqvt_coname:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10477
  fixes pi::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10478
  shows "(pi\<bullet>(ANDRIGHT (A AND B) X Y)) = ANDRIGHT (A AND B) (pi\<bullet>X) (pi\<bullet>Y)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10479
apply(auto simp add: perm_set_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10480
apply(rule_tac x="pi\<bullet>c" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10481
apply(rule_tac x="pi\<bullet>a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10482
apply(rule_tac x="pi\<bullet>b" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10483
apply(rule_tac x="pi\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10484
apply(rule_tac x="pi\<bullet>N" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10485
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10486
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10487
apply(drule_tac pi="pi" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10488
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10489
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10490
apply(rule_tac x="<a>:M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10491
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10492
apply(rule_tac x="<b>:N" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10493
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10494
apply(rule_tac x="(rev pi)\<bullet>(<c>:AndR <a>.M <b>.N c)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10495
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10496
apply(rule_tac x="(rev pi)\<bullet>c" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10497
apply(rule_tac x="(rev pi)\<bullet>a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10498
apply(rule_tac x="(rev pi)\<bullet>b" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10499
apply(rule_tac x="(rev pi)\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10500
apply(rule_tac x="(rev pi)\<bullet>N" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10501
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10502
apply(drule_tac pi="rev pi" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10503
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10504
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10505
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10506
apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10507
apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10508
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10509
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10510
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10511
consts
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10512
  ANDLEFT1::"ty \<Rightarrow> ntrm set \<Rightarrow> ntrm set"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10513
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10514
nominal_primrec
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10515
 "ANDLEFT1 (B AND C) X = { (y):AndL1 (x).M y | x y M. fin (AndL1 (x).M y) y \<and> (x):M \<in> X }"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10516
apply(rule TrueI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10517
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10518
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10519
lemma ANDLEFT1_eqvt_name:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10520
  fixes pi::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10521
  shows "(pi\<bullet>(ANDLEFT1 (A AND B) X)) = ANDLEFT1 (A AND B) (pi\<bullet>X)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10522
apply(auto simp add: perm_set_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10523
apply(rule_tac x="pi\<bullet>xb" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10524
apply(rule_tac x="pi\<bullet>y" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10525
apply(rule_tac x="pi\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10526
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10527
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10528
apply(drule_tac pi="pi" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10529
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10530
apply(rule_tac x="(xb):M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10531
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10532
apply(rule_tac x="(rev pi)\<bullet>((y):AndL1 (xa).M y)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10533
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10534
apply(rule_tac x="(rev pi)\<bullet>xa" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10535
apply(rule_tac x="(rev pi)\<bullet>y" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10536
apply(rule_tac x="(rev pi)\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10537
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10538
apply(drule_tac pi="rev pi" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10539
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10540
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10541
apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10542
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10543
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10544
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10545
lemma ANDLEFT1_eqvt_coname:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10546
  fixes pi::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10547
  shows "(pi\<bullet>(ANDLEFT1 (A AND B) X)) = ANDLEFT1 (A AND B) (pi\<bullet>X)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10548
apply(auto simp add: perm_set_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10549
apply(rule_tac x="pi\<bullet>xb" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10550
apply(rule_tac x="pi\<bullet>y" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10551
apply(rule_tac x="pi\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10552
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10553
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10554
apply(drule_tac pi="pi" in fin.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10555
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10556
apply(rule_tac x="(xb):M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10557
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10558
apply(rule_tac x="(rev pi)\<bullet>((y):AndL1 (xa).M y)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10559
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10560
apply(rule_tac x="(rev pi)\<bullet>xa" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10561
apply(rule_tac x="(rev pi)\<bullet>y" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10562
apply(rule_tac x="(rev pi)\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10563
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10564
apply(drule_tac pi="rev pi" in fin.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10565
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10566
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10567
apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10568
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10569
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10570
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10571
consts
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10572
  ANDLEFT2::"ty \<Rightarrow> ntrm set \<Rightarrow> ntrm set"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10573
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10574
nominal_primrec
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10575
 "ANDLEFT2 (B AND C) X = { (y):AndL2 (x).M y | x y M. fin (AndL2 (x).M y) y \<and> (x):M \<in> X }"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10576
apply(rule TrueI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10577
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10578
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10579
lemma ANDLEFT2_eqvt_name:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10580
  fixes pi::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10581
  shows "(pi\<bullet>(ANDLEFT2 (A AND B) X)) = ANDLEFT2 (A AND B) (pi\<bullet>X)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10582
apply(auto simp add: perm_set_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10583
apply(rule_tac x="pi\<bullet>xb" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10584
apply(rule_tac x="pi\<bullet>y" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10585
apply(rule_tac x="pi\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10586
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10587
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10588
apply(drule_tac pi="pi" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10589
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10590
apply(rule_tac x="(xb):M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10591
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10592
apply(rule_tac x="(rev pi)\<bullet>((y):AndL2 (xa).M y)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10593
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10594
apply(rule_tac x="(rev pi)\<bullet>xa" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10595
apply(rule_tac x="(rev pi)\<bullet>y" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10596
apply(rule_tac x="(rev pi)\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10597
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10598
apply(drule_tac pi="rev pi" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10599
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10600
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10601
apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10602
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10603
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10604
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10605
lemma ANDLEFT2_eqvt_coname:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10606
  fixes pi::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10607
  shows "(pi\<bullet>(ANDLEFT2 (A AND B) X)) = ANDLEFT2 (A AND B) (pi\<bullet>X)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10608
apply(auto simp add: perm_set_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10609
apply(rule_tac x="pi\<bullet>xb" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10610
apply(rule_tac x="pi\<bullet>y" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10611
apply(rule_tac x="pi\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10612
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10613
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10614
apply(drule_tac pi="pi" in fin.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10615
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10616
apply(rule_tac x="(xb):M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10617
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10618
apply(rule_tac x="(rev pi)\<bullet>((y):AndL2 (xa).M y)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10619
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10620
apply(rule_tac x="(rev pi)\<bullet>xa" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10621
apply(rule_tac x="(rev pi)\<bullet>y" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10622
apply(rule_tac x="(rev pi)\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10623
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10624
apply(drule_tac pi="rev pi" in fin.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10625
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10626
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10627
apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10628
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10629
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10630
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10631
consts
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10632
  ORLEFT::"ty \<Rightarrow> ntrm set \<Rightarrow> ntrm set \<Rightarrow> ntrm set"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10633
  
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10634
nominal_primrec
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10635
 "ORLEFT (B OR C) X Y = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10636
            { (z):OrL (x).M (y).N z | x y z M N. fin (OrL (x).M (y).N z) z \<and> (x):M \<in> X \<and> (y):N \<in> Y }"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10637
apply(rule TrueI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10638
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10639
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10640
lemma ORLEFT_eqvt_name:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10641
  fixes pi::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10642
  shows "(pi\<bullet>(ORLEFT (A OR B) X Y)) = ORLEFT (A OR B) (pi\<bullet>X) (pi\<bullet>Y)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10643
apply(auto simp add: perm_set_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10644
apply(rule_tac x="pi\<bullet>xb" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10645
apply(rule_tac x="pi\<bullet>y" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10646
apply(rule_tac x="pi\<bullet>z" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10647
apply(rule_tac x="pi\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10648
apply(rule_tac x="pi\<bullet>N" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10649
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10650
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10651
apply(drule_tac pi="pi" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10652
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10653
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10654
apply(rule_tac x="(xb):M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10655
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10656
apply(rule_tac x="(y):N" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10657
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10658
apply(rule_tac x="(rev pi)\<bullet>((z):OrL (xa).M (y).N z)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10659
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10660
apply(rule_tac x="(rev pi)\<bullet>xa" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10661
apply(rule_tac x="(rev pi)\<bullet>y" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10662
apply(rule_tac x="(rev pi)\<bullet>z" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10663
apply(rule_tac x="(rev pi)\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10664
apply(rule_tac x="(rev pi)\<bullet>N" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10665
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10666
apply(drule_tac pi="rev pi" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10667
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10668
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10669
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10670
apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10671
apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10672
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10673
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10674
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10675
lemma ORLEFT_eqvt_coname:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10676
  fixes pi::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10677
  shows "(pi\<bullet>(ORLEFT (A OR B) X Y)) = ORLEFT (A OR B) (pi\<bullet>X) (pi\<bullet>Y)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10678
apply(auto simp add: perm_set_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10679
apply(rule_tac x="pi\<bullet>xb" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10680
apply(rule_tac x="pi\<bullet>y" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10681
apply(rule_tac x="pi\<bullet>z" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10682
apply(rule_tac x="pi\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10683
apply(rule_tac x="pi\<bullet>N" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10684
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10685
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10686
apply(drule_tac pi="pi" in fin.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10687
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10688
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10689
apply(rule_tac x="(xb):M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10690
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10691
apply(rule_tac x="(y):N" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10692
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10693
apply(rule_tac x="(rev pi)\<bullet>((z):OrL (xa).M (y).N z)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10694
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10695
apply(rule_tac x="(rev pi)\<bullet>xa" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10696
apply(rule_tac x="(rev pi)\<bullet>y" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10697
apply(rule_tac x="(rev pi)\<bullet>z" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10698
apply(rule_tac x="(rev pi)\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10699
apply(rule_tac x="(rev pi)\<bullet>N" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10700
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10701
apply(drule_tac pi="rev pi" in fin.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10702
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10703
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10704
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10705
apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10706
apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10707
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10708
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10709
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10710
consts
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10711
  ORRIGHT1::"ty \<Rightarrow> ctrm set \<Rightarrow> ctrm set"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10712
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10713
nominal_primrec
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10714
 "ORRIGHT1 (B OR C) X = { <b>:OrR1 <a>.M b | a b M. fic (OrR1 <a>.M b) b \<and> <a>:M \<in> X }"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10715
apply(rule TrueI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10716
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10717
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10718
lemma ORRIGHT1_eqvt_name:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10719
  fixes pi::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10720
  shows "(pi\<bullet>(ORRIGHT1 (A OR B) X)) = ORRIGHT1 (A OR B) (pi\<bullet>X)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10721
apply(auto simp add: perm_set_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10722
apply(rule_tac x="pi\<bullet>a" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10723
apply(rule_tac x="pi\<bullet>b" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10724
apply(rule_tac x="pi\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10725
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10726
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10727
apply(drule_tac pi="pi" in fic.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10728
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10729
apply(rule_tac x="<a>:M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10730
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10731
apply(rule_tac x="(rev pi)\<bullet>(<b>:OrR1 <a>.M b)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10732
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10733
apply(rule_tac x="(rev pi)\<bullet>a" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10734
apply(rule_tac x="(rev pi)\<bullet>b" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10735
apply(rule_tac x="(rev pi)\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10736
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10737
apply(drule_tac pi="rev pi" in fic.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10738
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10739
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10740
apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10741
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10742
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10743
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10744
lemma ORRIGHT1_eqvt_coname:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10745
  fixes pi::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10746
  shows "(pi\<bullet>(ORRIGHT1 (A OR B) X)) = ORRIGHT1 (A OR B) (pi\<bullet>X)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10747
apply(auto simp add: perm_set_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10748
apply(rule_tac x="pi\<bullet>a" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10749
apply(rule_tac x="pi\<bullet>b" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10750
apply(rule_tac x="pi\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10751
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10752
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10753
apply(drule_tac pi="pi" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10754
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10755
apply(rule_tac x="<a>:M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10756
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10757
apply(rule_tac x="(rev pi)\<bullet>(<b>:OrR1 <a>.M b)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10758
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10759
apply(rule_tac x="(rev pi)\<bullet>a" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10760
apply(rule_tac x="(rev pi)\<bullet>b" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10761
apply(rule_tac x="(rev pi)\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10762
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10763
apply(drule_tac pi="rev pi" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10764
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10765
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10766
apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10767
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10768
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10769
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10770
consts
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10771
  ORRIGHT2::"ty \<Rightarrow> ctrm set \<Rightarrow> ctrm set"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10772
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10773
nominal_primrec
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10774
 "ORRIGHT2 (B OR C) X = { <b>:OrR2 <a>.M b | a b M. fic (OrR2 <a>.M b) b \<and> <a>:M \<in> X }"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10775
apply(rule TrueI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10776
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10777
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10778
lemma ORRIGHT2_eqvt_name:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10779
  fixes pi::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10780
  shows "(pi\<bullet>(ORRIGHT2 (A OR B) X)) = ORRIGHT2 (A OR B) (pi\<bullet>X)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10781
apply(auto simp add: perm_set_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10782
apply(rule_tac x="pi\<bullet>a" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10783
apply(rule_tac x="pi\<bullet>b" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10784
apply(rule_tac x="pi\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10785
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10786
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10787
apply(drule_tac pi="pi" in fic.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10788
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10789
apply(rule_tac x="<a>:M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10790
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10791
apply(rule_tac x="(rev pi)\<bullet>(<b>:OrR2 <a>.M b)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10792
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10793
apply(rule_tac x="(rev pi)\<bullet>a" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10794
apply(rule_tac x="(rev pi)\<bullet>b" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10795
apply(rule_tac x="(rev pi)\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10796
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10797
apply(drule_tac pi="rev pi" in fic.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10798
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10799
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10800
apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10801
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10802
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10803
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10804
lemma ORRIGHT2_eqvt_coname:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10805
  fixes pi::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10806
  shows "(pi\<bullet>(ORRIGHT2 (A OR B) X)) = ORRIGHT2 (A OR B) (pi\<bullet>X)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10807
apply(auto simp add: perm_set_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10808
apply(rule_tac x="pi\<bullet>a" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10809
apply(rule_tac x="pi\<bullet>b" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10810
apply(rule_tac x="pi\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10811
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10812
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10813
apply(drule_tac pi="pi" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10814
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10815
apply(rule_tac x="<a>:M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10816
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10817
apply(rule_tac x="(rev pi)\<bullet>(<b>:OrR2 <a>.M b)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10818
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10819
apply(rule_tac x="(rev pi)\<bullet>a" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10820
apply(rule_tac x="(rev pi)\<bullet>b" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10821
apply(rule_tac x="(rev pi)\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10822
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10823
apply(drule_tac pi="rev pi" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10824
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10825
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10826
apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10827
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10828
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10829
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10830
consts
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10831
  IMPRIGHT::"ty \<Rightarrow> ntrm set \<Rightarrow> ctrm set \<Rightarrow> ntrm set \<Rightarrow> ctrm set \<Rightarrow> ctrm set"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10832
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10833
nominal_primrec
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10834
 "IMPRIGHT (B IMP C) X Y Z U= 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10835
        { <b>:ImpR (x).<a>.M b | x a b M. fic (ImpR (x).<a>.M b) b 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10836
                                        \<and> (\<forall>z P. x\<sharp>(z,P) \<and> (z):P \<in> Z \<longrightarrow> (x):(M{a:=(z).P}) \<in> X)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10837
                                        \<and> (\<forall>c Q. a\<sharp>(c,Q) \<and> <c>:Q \<in> U \<longrightarrow> <a>:(M{x:=<c>.Q}) \<in> Y)}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10838
apply(rule TrueI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10839
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10840
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10841
lemma IMPRIGHT_eqvt_name:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10842
  fixes pi::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10843
  shows "(pi\<bullet>(IMPRIGHT (A IMP B) X Y Z U)) = IMPRIGHT (A IMP B) (pi\<bullet>X) (pi\<bullet>Y) (pi\<bullet>Z) (pi\<bullet>U)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10844
apply(auto simp add: perm_set_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10845
apply(rule_tac x="pi\<bullet>xb" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10846
apply(rule_tac x="pi\<bullet>a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10847
apply(rule_tac x="pi\<bullet>b" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10848
apply(rule_tac x="pi\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10849
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10850
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10851
apply(drule_tac pi="pi" in fic.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10852
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10853
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10854
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10855
apply(rule_tac x="(xb):(M{a:=((rev pi)\<bullet>z).((rev pi)\<bullet>P)})" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10856
apply(perm_simp add: csubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10857
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10858
apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10859
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10860
apply(simp add: fresh_right)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10861
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10862
apply(rule_tac x="<a>:(M{xb:=<((rev pi)\<bullet>c)>.((rev pi)\<bullet>Q)})" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10863
apply(perm_simp add: nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10864
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10865
apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10866
apply(simp add: fresh_right)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10867
apply(rule_tac x="(rev pi)\<bullet>(<b>:ImpR xa.<a>.M b)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10868
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10869
apply(rule_tac x="(rev pi)\<bullet>xa" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10870
apply(rule_tac x="(rev pi)\<bullet>a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10871
apply(rule_tac x="(rev pi)\<bullet>b" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10872
apply(rule_tac x="(rev pi)\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10873
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10874
apply(drule_tac pi="rev pi" in fic.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10875
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10876
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10877
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10878
apply(drule_tac x="pi\<bullet>z" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10879
apply(drule_tac x="pi\<bullet>P" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10880
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10881
apply(simp add: fresh_right)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10882
apply(rule_tac x="(z):P" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10883
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10884
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10885
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10886
apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10887
apply(perm_simp add: csubst_eqvt fresh_right)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10888
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10889
apply(drule_tac x="pi\<bullet>c" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10890
apply(drule_tac x="pi\<bullet>Q" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10891
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10892
apply(simp add: fresh_right)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10893
apply(rule_tac x="<c>:Q" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10894
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10895
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10896
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10897
apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10898
apply(perm_simp add: nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10899
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10900
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10901
lemma IMPRIGHT_eqvt_coname:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10902
  fixes pi::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10903
  shows "(pi\<bullet>(IMPRIGHT (A IMP B) X Y Z U)) = IMPRIGHT (A IMP B) (pi\<bullet>X) (pi\<bullet>Y) (pi\<bullet>Z) (pi\<bullet>U)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10904
apply(auto simp add: perm_set_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10905
apply(rule_tac x="pi\<bullet>xb" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10906
apply(rule_tac x="pi\<bullet>a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10907
apply(rule_tac x="pi\<bullet>b" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10908
apply(rule_tac x="pi\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10909
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10910
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10911
apply(drule_tac pi="pi" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10912
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10913
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10914
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10915
apply(rule_tac x="(xb):(M{a:=((rev pi)\<bullet>z).((rev pi)\<bullet>P)})" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10916
apply(perm_simp add: csubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10917
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10918
apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10919
apply(simp add: fresh_right)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10920
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10921
apply(rule_tac x="<a>:(M{xb:=<((rev pi)\<bullet>c)>.((rev pi)\<bullet>Q)})" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10922
apply(perm_simp add: nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10923
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10924
apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10925
apply(simp add: fresh_right)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10926
apply(rule_tac x="(rev pi)\<bullet>(<b>:ImpR xa.<a>.M b)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10927
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10928
apply(rule_tac x="(rev pi)\<bullet>xa" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10929
apply(rule_tac x="(rev pi)\<bullet>a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10930
apply(rule_tac x="(rev pi)\<bullet>b" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10931
apply(rule_tac x="(rev pi)\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10932
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10933
apply(drule_tac pi="rev pi" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10934
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10935
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10936
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10937
apply(drule_tac x="pi\<bullet>z" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10938
apply(drule_tac x="pi\<bullet>P" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10939
apply(simp add: fresh_right)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10940
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10941
apply(rule_tac x="(z):P" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10942
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10943
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10944
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10945
apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10946
apply(perm_simp add: csubst_eqvt fresh_right)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10947
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10948
apply(drule_tac x="pi\<bullet>c" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10949
apply(drule_tac x="pi\<bullet>Q" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10950
apply(simp add: fresh_right)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10951
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10952
apply(rule_tac x="<c>:Q" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10953
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10954
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10955
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10956
apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10957
apply(perm_simp add: nsubst_eqvt fresh_right)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10958
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10959
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10960
consts
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10961
  IMPLEFT::"ty \<Rightarrow> ctrm set \<Rightarrow> ntrm set \<Rightarrow> ntrm set"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10962
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10963
nominal_primrec
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10964
 "IMPLEFT (B IMP C) X Y = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10965
        { (y):ImpL <a>.M (x).N y | x a y M N. fin (ImpL <a>.M (x).N y) y \<and> <a>:M \<in> X \<and> (x):N \<in> Y }"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10966
apply(rule TrueI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10967
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10968
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10969
lemma IMPLEFT_eqvt_name:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10970
  fixes pi::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10971
  shows "(pi\<bullet>(IMPLEFT (A IMP B) X Y)) = IMPLEFT (A IMP B) (pi\<bullet>X) (pi\<bullet>Y)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10972
apply(auto simp add: perm_set_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10973
apply(rule_tac x="pi\<bullet>xb" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10974
apply(rule_tac x="pi\<bullet>a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10975
apply(rule_tac x="pi\<bullet>y" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10976
apply(rule_tac x="pi\<bullet>M" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10977
apply(rule_tac x="pi\<bullet>N" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10978
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10979
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10980
apply(drule_tac pi="pi" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10981
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10982
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10983
apply(rule_tac x="<a>:M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10984
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10985
apply(rule_tac x="(xb):N" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10986
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10987
apply(rule_tac x="(rev pi)\<bullet>((y):ImpL <a>.M (xa).N y)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10988
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10989
apply(rule_tac x="(rev pi)\<bullet>xa" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10990
apply(rule_tac x="(rev pi)\<bullet>a" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10991
apply(rule_tac x="(rev pi)\<bullet>y" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10992
apply(rule_tac x="(rev pi)\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10993
apply(rule_tac x="(rev pi)\<bullet>N" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10994
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10995
apply(drule_tac pi="rev pi" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10996
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10997
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10998
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 10999
apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11000
apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11001
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11002
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11003
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11004
lemma IMPLEFT_eqvt_coname:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11005
  fixes pi::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11006
  shows "(pi\<bullet>(IMPLEFT (A IMP B) X Y)) = IMPLEFT (A IMP B) (pi\<bullet>X) (pi\<bullet>Y)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11007
apply(auto simp add: perm_set_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11008
apply(rule_tac x="pi\<bullet>xb" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11009
apply(rule_tac x="pi\<bullet>a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11010
apply(rule_tac x="pi\<bullet>y" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11011
apply(rule_tac x="pi\<bullet>M" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11012
apply(rule_tac x="pi\<bullet>N" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11013
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11014
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11015
apply(drule_tac pi="pi" in fin.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11016
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11017
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11018
apply(rule_tac x="<a>:M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11019
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11020
apply(rule_tac x="(xb):N" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11021
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11022
apply(rule_tac x="(rev pi)\<bullet>((y):ImpL <a>.M (xa).N y)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11023
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11024
apply(rule_tac x="(rev pi)\<bullet>xa" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11025
apply(rule_tac x="(rev pi)\<bullet>a" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11026
apply(rule_tac x="(rev pi)\<bullet>y" in exI) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11027
apply(rule_tac x="(rev pi)\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11028
apply(rule_tac x="(rev pi)\<bullet>N" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11029
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11030
apply(drule_tac pi="rev pi" in fin.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11031
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11032
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11033
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11034
apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11035
apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11036
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11037
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11038
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11039
lemma sum_cases:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11040
 shows "(\<exists>y. x=Inl y) \<or> (\<exists>y. x=Inr y)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11041
apply(rule_tac s="x" in sumE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11042
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11043
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11044
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11045
function
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11046
  NEGc::"ty \<Rightarrow> ntrm set \<Rightarrow> ctrm set"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11047
and
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11048
  NEGn::"ty \<Rightarrow> ctrm set \<Rightarrow> ntrm set"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11049
where
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11050
  "NEGc (PR A)    X = AXIOMSc (PR A) \<union> BINDINGc (PR A) X"  
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11051
| "NEGc (NOT C)   X = AXIOMSc (NOT C) \<union> BINDINGc (NOT C) X 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11052
                         \<union> NOTRIGHT (NOT C) (lfp (NEGn C \<circ> NEGc C))"  
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11053
| "NEGc (C AND D) X = AXIOMSc (C AND D) \<union> BINDINGc (C AND D) X 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11054
                     \<union> ANDRIGHT (C AND D) (NEGc C (lfp (NEGn C \<circ> NEGc C))) (NEGc D (lfp (NEGn D \<circ> NEGc D)))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11055
| "NEGc (C OR D)  X = AXIOMSc (C OR D) \<union> BINDINGc (C OR D) X  
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11056
                         \<union> ORRIGHT1 (C OR D) (NEGc C (lfp (NEGn C \<circ> NEGc C))) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11057
                         \<union> ORRIGHT2 (C OR D) (NEGc D (lfp (NEGn D \<circ> NEGc D)))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11058
| "NEGc (C IMP D) X = AXIOMSc (C IMP D) \<union> BINDINGc (C IMP D) X 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11059
    \<union> IMPRIGHT (C IMP D) (lfp (NEGn C \<circ> NEGc C)) (NEGc D (lfp (NEGn D \<circ> NEGc D))) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11060
                          (lfp (NEGn D \<circ> NEGc D)) (NEGc C (lfp (NEGn C \<circ> NEGc C)))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11061
| "NEGn (PR A)    X = AXIOMSn (PR A) \<union> BINDINGn (PR A) X"   
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11062
| "NEGn (NOT C)   X = AXIOMSn (NOT C) \<union> BINDINGn (NOT C) X 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11063
                         \<union> NOTLEFT (NOT C) (NEGc C (lfp (NEGn C \<circ> NEGc C)))"  
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11064
| "NEGn (C AND D) X = AXIOMSn (C AND D) \<union> BINDINGn (C AND D) X 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11065
                         \<union> ANDLEFT1 (C AND D) (lfp (NEGn C \<circ> NEGc C)) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11066
                         \<union> ANDLEFT2 (C AND D) (lfp (NEGn D \<circ> NEGc D))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11067
| "NEGn (C OR D)  X = AXIOMSn (C OR D) \<union> BINDINGn (C OR D) X 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11068
                         \<union> ORLEFT (C OR D) (lfp (NEGn C \<circ> NEGc C)) (lfp (NEGn D \<circ> NEGc D))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11069
| "NEGn (C IMP D) X = AXIOMSn (C IMP D) \<union> BINDINGn (C IMP D) X 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11070
                         \<union> IMPLEFT (C IMP D) (NEGc C (lfp (NEGn C \<circ> NEGc C))) (lfp (NEGn D \<circ> NEGc D))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11071
using ty_cases sum_cases 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11072
apply(auto simp add: ty.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11073
apply(drule_tac x="x" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11074
apply(auto simp add: ty.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11075
apply(rotate_tac 10)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11076
apply(drule_tac x="a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11077
apply(auto simp add: ty.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11078
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11079
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11080
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11081
apply(rotate_tac 10)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11082
apply(drule_tac x="a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11083
apply(auto simp add: ty.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11084
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11085
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11086
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11087
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11088
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11089
termination
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11090
apply(relation "measure (sum_case (size\<circ>fst) (size\<circ>fst))")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11091
apply(simp_all)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11092
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11093
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11094
text {* Candidates *}
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11095
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11096
lemma test1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11097
  shows "x\<in>(X\<union>Y) = (x\<in>X \<or> x\<in>Y)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11098
by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11099
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11100
lemma test2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11101
  shows "x\<in>(X\<inter>Y) = (x\<in>X \<and> x\<in>Y)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11102
by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11103
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11104
lemma pt_Collect_eqvt:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11105
  fixes pi::"'x prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11106
  assumes pt: "pt TYPE('a) TYPE('x)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11107
  and     at: "at TYPE('x)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11108
  shows "pi\<bullet>{x::'a. P x} = {x. P ((rev pi)\<bullet>x)}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11109
apply(auto simp add: perm_set_def  pt_rev_pi[OF pt, OF at])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11110
apply(rule_tac x="(rev pi)\<bullet>x" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11111
apply(simp add: pt_pi_rev[OF pt, OF at])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11112
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11113
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11114
lemma big_inter_eqvt:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11115
  fixes pi1::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11116
  and   X::"('a::pt_name) set set"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11117
  and   pi2::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11118
  and   Y::"('b::pt_coname) set set"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11119
  shows "(pi1\<bullet>(\<Inter> X)) = \<Inter> (pi1\<bullet>X)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11120
  and   "(pi2\<bullet>(\<Inter> Y)) = \<Inter> (pi2\<bullet>Y)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11121
apply(auto simp add: perm_set_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11122
apply(rule_tac x="(rev pi1)\<bullet>x" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11123
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11124
apply(rule ballI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11125
apply(drule_tac x="pi1\<bullet>xa" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11126
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11127
apply(drule_tac x="xa" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11128
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11129
apply(rule_tac x="(rev pi1)\<bullet>xb" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11130
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11131
apply(simp add: pt_set_bij1[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11132
apply(simp add: pt_set_bij[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11133
apply(simp add: pt_set_bij1[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11134
apply(rule_tac x="(rev pi2)\<bullet>x" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11135
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11136
apply(rule ballI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11137
apply(drule_tac x="pi2\<bullet>xa" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11138
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11139
apply(drule_tac x="xa" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11140
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11141
apply(rule_tac x="(rev pi2)\<bullet>xb" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11142
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11143
apply(simp add: pt_set_bij1[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11144
apply(simp add: pt_set_bij[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11145
apply(simp add: pt_set_bij1[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11146
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11147
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11148
lemma lfp_eqvt:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11149
  fixes pi1::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11150
  and   f::"'a set \<Rightarrow> ('a::pt_name) set"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11151
  and   pi2::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11152
  and   g::"'b set \<Rightarrow> ('b::pt_coname) set"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11153
  shows "pi1\<bullet>(lfp f) = lfp (pi1\<bullet>f)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11154
  and   "pi2\<bullet>(lfp g) = lfp (pi2\<bullet>g)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11155
apply(simp add: lfp_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11156
apply(simp add: Inf_set_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11157
apply(simp add: big_inter_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11158
apply(simp add: pt_Collect_eqvt[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11159
apply(subgoal_tac "{u. (pi1\<bullet>f) u \<subseteq> u} = {u. ((rev pi1)\<bullet>((pi1\<bullet>f) u)) \<subseteq> ((rev pi1)\<bullet>u)}")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11160
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11161
apply(simp add: pt_subseteq_eqvt[OF pt_name_inst,OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11162
apply(simp add: lfp_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11163
apply(simp add: Inf_set_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11164
apply(simp add: big_inter_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11165
apply(simp add: pt_Collect_eqvt[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11166
apply(subgoal_tac "{u. (pi2\<bullet>g) u \<subseteq> u} = {u. ((rev pi2)\<bullet>((pi2\<bullet>g) u)) \<subseteq> ((rev pi2)\<bullet>u)}")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11167
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11168
apply(simp add: pt_subseteq_eqvt[OF pt_coname_inst,OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11169
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11170
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11171
abbreviation
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11172
  CANDn::"ty \<Rightarrow> ntrm set"  ("\<parallel>'(_')\<parallel>" [60] 60) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11173
where
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11174
  "\<parallel>(B)\<parallel> \<equiv> lfp (NEGn B \<circ> NEGc B)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11175
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11176
abbreviation
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11177
  CANDc::"ty \<Rightarrow> ctrm set"  ("\<parallel><_>\<parallel>" [60] 60)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11178
where
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11179
  "\<parallel><B>\<parallel> \<equiv> NEGc B (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11180
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11181
lemma NEGn_decreasing:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11182
  shows "X\<subseteq>Y \<Longrightarrow> NEGn B Y \<subseteq> NEGn B X"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11183
by (nominal_induct B rule: ty.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11184
   (auto dest: BINDINGn_decreasing)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11185
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11186
lemma NEGc_decreasing:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11187
  shows "X\<subseteq>Y \<Longrightarrow> NEGc B Y \<subseteq> NEGc B X"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11188
by (nominal_induct B rule: ty.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11189
   (auto dest: BINDINGc_decreasing)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11190
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11191
lemma mono_NEGn_NEGc:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11192
  shows "mono (NEGn B \<circ> NEGc B)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11193
  and   "mono (NEGc B \<circ> NEGn B)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11194
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11195
  have "\<forall>X Y. X\<subseteq>Y \<longrightarrow> NEGn B (NEGc B X) \<subseteq> NEGn B (NEGc B Y)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11196
  proof (intro strip)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11197
    fix X::"ntrm set" and Y::"ntrm set"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11198
    assume "X\<subseteq>Y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11199
    then have "NEGc B Y \<subseteq> NEGc B X" by (simp add: NEGc_decreasing)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11200
    then show "NEGn B (NEGc B X) \<subseteq> NEGn B (NEGc B Y)" by (simp add: NEGn_decreasing)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11201
  qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11202
  then show "mono (NEGn B \<circ> NEGc B)" by (simp add: mono_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11203
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11204
  have "\<forall>X Y. X\<subseteq>Y \<longrightarrow> NEGc B (NEGn B X) \<subseteq> NEGc B (NEGn B Y)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11205
  proof (intro strip)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11206
    fix X::"ctrm set" and Y::"ctrm set"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11207
    assume "X\<subseteq>Y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11208
    then have "NEGn B Y \<subseteq> NEGn B X" by (simp add: NEGn_decreasing)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11209
    then show "NEGc B (NEGn B X) \<subseteq> NEGc B (NEGn B Y)" by (simp add: NEGc_decreasing)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11210
  qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11211
  then show "mono (NEGc B \<circ> NEGn B)" by (simp add: mono_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11212
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11213
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11214
lemma NEG_simp:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11215
  shows "\<parallel><B>\<parallel> = NEGc B (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11216
  and   "\<parallel>(B)\<parallel> = NEGn B (\<parallel><B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11217
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11218
  show "\<parallel><B>\<parallel> = NEGc B (\<parallel>(B)\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11219
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11220
  have "\<parallel>(B)\<parallel> \<equiv> lfp (NEGn B \<circ> NEGc B)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11221
  then have "\<parallel>(B)\<parallel> = (NEGn B \<circ> NEGc B) (\<parallel>(B)\<parallel>)" using mono_NEGn_NEGc def_lfp_unfold by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11222
  then show "\<parallel>(B)\<parallel> = NEGn B (\<parallel><B>\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11223
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11224
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11225
lemma NEG_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11226
  shows "M \<in> \<parallel><B>\<parallel> \<Longrightarrow> M \<in> NEGc B (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11227
  and   "N \<in> \<parallel>(B)\<parallel> \<Longrightarrow> N \<in> NEGn B (\<parallel><B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11228
using NEG_simp by (blast)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11229
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11230
lemma NEG_intro:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11231
  shows "M \<in> NEGc B (\<parallel>(B)\<parallel>) \<Longrightarrow> M \<in> \<parallel><B>\<parallel>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11232
  and   "N \<in> NEGn B (\<parallel><B>\<parallel>) \<Longrightarrow> N \<in> \<parallel>(B)\<parallel>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11233
using NEG_simp by (blast)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11234
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11235
lemma NEGc_simps:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11236
  shows "NEGc (PR A) (\<parallel>(PR A)\<parallel>) = AXIOMSc (PR A) \<union> BINDINGc (PR A) (\<parallel>(PR A)\<parallel>)"  
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11237
  and   "NEGc (NOT C) (\<parallel>(NOT C)\<parallel>) = AXIOMSc (NOT C) \<union> BINDINGc (NOT C) (\<parallel>(NOT C)\<parallel>) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11238
                                        \<union> (NOTRIGHT (NOT C) (\<parallel>(C)\<parallel>))"  
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11239
  and   "NEGc (C AND D) (\<parallel>(C AND D)\<parallel>) = AXIOMSc (C AND D) \<union> BINDINGc (C AND D) (\<parallel>(C AND D)\<parallel>) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11240
                                        \<union> (ANDRIGHT (C AND D) (\<parallel><C>\<parallel>) (\<parallel><D>\<parallel>))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11241
  and   "NEGc (C OR D) (\<parallel>(C OR D)\<parallel>) = AXIOMSc (C OR D) \<union> BINDINGc (C OR D)  (\<parallel>(C OR D)\<parallel>)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11242
                                        \<union> (ORRIGHT1 (C OR D) (\<parallel><C>\<parallel>)) \<union> (ORRIGHT2 (C OR D) (\<parallel><D>\<parallel>))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11243
  and   "NEGc (C IMP D) (\<parallel>(C IMP D)\<parallel>) = AXIOMSc (C IMP D) \<union> BINDINGc (C IMP D) (\<parallel>(C IMP D)\<parallel>) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11244
          \<union> (IMPRIGHT (C IMP D) (\<parallel>(C)\<parallel>) (\<parallel><D>\<parallel>) (\<parallel>(D)\<parallel>) (\<parallel><C>\<parallel>))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11245
by (simp_all only: NEGc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11246
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11247
lemma AXIOMS_in_CANDs:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11248
  shows "AXIOMSn B \<subseteq> (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11249
  and   "AXIOMSc B \<subseteq> (\<parallel><B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11250
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11251
  have "AXIOMSn B \<subseteq> NEGn B (\<parallel><B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11252
    by (nominal_induct B rule: ty.induct) (auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11253
  then show "AXIOMSn B \<subseteq> \<parallel>(B)\<parallel>" using NEG_simp by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11254
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11255
  have "AXIOMSc B \<subseteq> NEGc B (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11256
    by (nominal_induct B rule: ty.induct) (auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11257
  then show "AXIOMSc B \<subseteq> \<parallel><B>\<parallel>" using NEG_simp by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11258
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11259
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11260
lemma Ax_in_CANDs:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11261
  shows "(y):Ax x a \<in> \<parallel>(B)\<parallel>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11262
  and   "<b>:Ax x a \<in> \<parallel><B>\<parallel>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11263
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11264
  have "(y):Ax x a \<in> AXIOMSn B" by (auto simp add: AXIOMSn_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11265
  also have "AXIOMSn B \<subseteq> \<parallel>(B)\<parallel>" by (rule AXIOMS_in_CANDs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11266
  finally show "(y):Ax x a \<in> \<parallel>(B)\<parallel>" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11267
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11268
  have "<b>:Ax x a \<in> AXIOMSc B" by (auto simp add: AXIOMSc_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11269
  also have "AXIOMSc B \<subseteq> \<parallel><B>\<parallel>" by (rule AXIOMS_in_CANDs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11270
  finally show "<b>:Ax x a \<in> \<parallel><B>\<parallel>" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11271
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11272
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11273
lemma AXIOMS_eqvt_aux_name:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11274
  fixes pi::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11275
  shows "M \<in> AXIOMSn B \<Longrightarrow> (pi\<bullet>M) \<in> AXIOMSn B" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11276
  and   "N \<in> AXIOMSc B \<Longrightarrow> (pi\<bullet>N) \<in> AXIOMSc B"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11277
apply(auto simp add: AXIOMSn_def AXIOMSc_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11278
apply(rule_tac x="pi\<bullet>x" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11279
apply(rule_tac x="pi\<bullet>y" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11280
apply(rule_tac x="pi\<bullet>b" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11281
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11282
apply(rule_tac x="pi\<bullet>a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11283
apply(rule_tac x="pi\<bullet>y" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11284
apply(rule_tac x="pi\<bullet>b" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11285
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11286
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11287
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11288
lemma AXIOMS_eqvt_aux_coname:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11289
  fixes pi::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11290
  shows "M \<in> AXIOMSn B \<Longrightarrow> (pi\<bullet>M) \<in> AXIOMSn B" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11291
  and   "N \<in> AXIOMSc B \<Longrightarrow> (pi\<bullet>N) \<in> AXIOMSc B"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11292
apply(auto simp add: AXIOMSn_def AXIOMSc_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11293
apply(rule_tac x="pi\<bullet>x" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11294
apply(rule_tac x="pi\<bullet>y" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11295
apply(rule_tac x="pi\<bullet>b" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11296
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11297
apply(rule_tac x="pi\<bullet>a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11298
apply(rule_tac x="pi\<bullet>y" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11299
apply(rule_tac x="pi\<bullet>b" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11300
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11301
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11302
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11303
lemma AXIOMS_eqvt_name:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11304
  fixes pi::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11305
  shows "(pi\<bullet>AXIOMSn B) = AXIOMSn B" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11306
  and   "(pi\<bullet>AXIOMSc B) = AXIOMSc B"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11307
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11308
apply(simp add: pt_set_bij1a[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11309
apply(drule_tac pi="pi" in AXIOMS_eqvt_aux_name(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11310
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11311
apply(drule_tac pi="rev pi" in AXIOMS_eqvt_aux_name(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11312
apply(simp add: pt_set_bij1[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11313
apply(simp add: pt_set_bij1a[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11314
apply(drule_tac pi="pi" in AXIOMS_eqvt_aux_name(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11315
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11316
apply(drule_tac pi="rev pi" in AXIOMS_eqvt_aux_name(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11317
apply(simp add: pt_set_bij1[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11318
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11319
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11320
lemma AXIOMS_eqvt_coname:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11321
  fixes pi::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11322
  shows "(pi\<bullet>AXIOMSn B) = AXIOMSn B" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11323
  and   "(pi\<bullet>AXIOMSc B) = AXIOMSc B"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11324
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11325
apply(simp add: pt_set_bij1a[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11326
apply(drule_tac pi="pi" in AXIOMS_eqvt_aux_coname(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11327
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11328
apply(drule_tac pi="rev pi" in AXIOMS_eqvt_aux_coname(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11329
apply(simp add: pt_set_bij1[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11330
apply(simp add: pt_set_bij1a[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11331
apply(drule_tac pi="pi" in AXIOMS_eqvt_aux_coname(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11332
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11333
apply(drule_tac pi="rev pi" in AXIOMS_eqvt_aux_coname(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11334
apply(simp add: pt_set_bij1[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11335
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11336
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11337
lemma BINDING_eqvt_name:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11338
  fixes pi::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11339
  shows "(pi\<bullet>(BINDINGn B X)) = BINDINGn B (pi\<bullet>X)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11340
  and   "(pi\<bullet>(BINDINGc B Y)) = BINDINGc B (pi\<bullet>Y)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11341
apply(auto simp add: BINDINGn_def BINDINGc_def perm_set_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11342
apply(rule_tac x="pi\<bullet>xb" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11343
apply(rule_tac x="pi\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11344
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11345
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11346
apply(drule_tac x="(rev pi)\<bullet>a" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11347
apply(drule_tac x="(rev pi)\<bullet>P" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11348
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11349
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11350
apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11351
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11352
apply(drule_tac ?pi1.0="pi" in SNa_eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11353
apply(perm_simp add: nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11354
apply(rule_tac x="(rev pi\<bullet>xa):(rev pi\<bullet>M)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11355
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11356
apply(rule_tac x="rev pi\<bullet>xa" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11357
apply(rule_tac x="rev pi\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11358
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11359
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11360
apply(drule_tac x="pi\<bullet>a" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11361
apply(drule_tac x="pi\<bullet>P" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11362
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11363
apply(force)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11364
apply(drule_tac ?pi1.0="rev pi" in SNa_eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11365
apply(perm_simp add: nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11366
apply(rule_tac x="pi\<bullet>a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11367
apply(rule_tac x="pi\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11368
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11369
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11370
apply(drule_tac x="(rev pi)\<bullet>x" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11371
apply(drule_tac x="(rev pi)\<bullet>P" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11372
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11373
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11374
apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11375
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11376
apply(drule_tac ?pi1.0="pi" in SNa_eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11377
apply(perm_simp add: csubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11378
apply(rule_tac x="<(rev pi\<bullet>a)>:(rev pi\<bullet>M)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11379
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11380
apply(rule_tac x="rev pi\<bullet>a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11381
apply(rule_tac x="rev pi\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11382
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11383
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11384
apply(drule_tac x="pi\<bullet>x" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11385
apply(drule_tac x="pi\<bullet>P" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11386
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11387
apply(force)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11388
apply(drule_tac ?pi1.0="rev pi" in SNa_eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11389
apply(perm_simp add: csubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11390
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11391
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11392
lemma BINDING_eqvt_coname:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11393
  fixes pi::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11394
  shows "(pi\<bullet>(BINDINGn B X)) = BINDINGn B (pi\<bullet>X)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11395
  and   "(pi\<bullet>(BINDINGc B Y)) = BINDINGc B (pi\<bullet>Y)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11396
apply(auto simp add: BINDINGn_def BINDINGc_def perm_set_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11397
apply(rule_tac x="pi\<bullet>xb" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11398
apply(rule_tac x="pi\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11399
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11400
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11401
apply(drule_tac x="(rev pi)\<bullet>a" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11402
apply(drule_tac x="(rev pi)\<bullet>P" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11403
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11404
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11405
apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11406
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11407
apply(drule_tac ?pi2.0="pi" in SNa_eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11408
apply(perm_simp add: nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11409
apply(rule_tac x="(rev pi\<bullet>xa):(rev pi\<bullet>M)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11410
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11411
apply(rule_tac x="rev pi\<bullet>xa" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11412
apply(rule_tac x="rev pi\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11413
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11414
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11415
apply(drule_tac x="pi\<bullet>a" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11416
apply(drule_tac x="pi\<bullet>P" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11417
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11418
apply(force)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11419
apply(drule_tac ?pi2.0="rev pi" in SNa_eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11420
apply(perm_simp add: nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11421
apply(rule_tac x="pi\<bullet>a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11422
apply(rule_tac x="pi\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11423
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11424
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11425
apply(drule_tac x="(rev pi)\<bullet>x" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11426
apply(drule_tac x="(rev pi)\<bullet>P" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11427
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11428
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11429
apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11430
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11431
apply(drule_tac ?pi2.0="pi" in SNa_eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11432
apply(perm_simp add: csubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11433
apply(rule_tac x="<(rev pi\<bullet>a)>:(rev pi\<bullet>M)" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11434
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11435
apply(rule_tac x="rev pi\<bullet>a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11436
apply(rule_tac x="rev pi\<bullet>M" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11437
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11438
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11439
apply(drule_tac x="pi\<bullet>x" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11440
apply(drule_tac x="pi\<bullet>P" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11441
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11442
apply(force)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11443
apply(drule_tac ?pi2.0="rev pi" in SNa_eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11444
apply(perm_simp add: csubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11445
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11446
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11447
lemma CAND_eqvt_name:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11448
  fixes pi::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11449
  shows   "(pi\<bullet>(\<parallel>(B)\<parallel>)) = (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11450
  and     "(pi\<bullet>(\<parallel><B>\<parallel>)) = (\<parallel><B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11451
proof (nominal_induct B rule: ty.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11452
  case (PR X)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11453
  { case 1 show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11454
      apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11455
      apply(simp add: lfp_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11456
      apply(simp add: perm_fun_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11457
      apply(simp add: union_eqvt AXIOMS_eqvt_name BINDING_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11458
      apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11459
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11460
  next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11461
    case 2 show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11462
      apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11463
      apply(simp only: NEGc_simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11464
      apply(simp add: union_eqvt AXIOMS_eqvt_name BINDING_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11465
      apply(simp add: lfp_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11466
      apply(simp add: comp_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11467
      apply(simp add: perm_fun_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11468
      apply(simp add: union_eqvt AXIOMS_eqvt_name BINDING_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11469
      apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11470
      done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11471
  }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11472
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11473
  case (NOT B)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11474
  have ih1: "pi\<bullet>(\<parallel>(B)\<parallel>) = (\<parallel>(B)\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11475
  have ih2: "pi\<bullet>(\<parallel><B>\<parallel>) = (\<parallel><B>\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11476
  have g: "pi\<bullet>(\<parallel>(NOT B)\<parallel>) = (\<parallel>(NOT B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11477
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11478
    apply(simp only: lfp_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11479
    apply(simp only: comp_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11480
    apply(simp only: perm_fun_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11481
    apply(simp only: NEGc.simps NEGn.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11482
    apply(simp only: union_eqvt AXIOMS_eqvt_name BINDING_eqvt_name NOTRIGHT_eqvt_name NOTLEFT_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11483
    apply(perm_simp add: ih1 ih2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11484
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11485
  { case 1 show ?case by (rule g)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11486
  next 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11487
    case 2 show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11488
      by (simp only: NEGc_simps union_eqvt AXIOMS_eqvt_name BINDING_eqvt_name NOTRIGHT_eqvt_name ih1 ih2 g)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11489
  }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11490
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11491
  case (AND A B)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11492
  have ih1: "pi\<bullet>(\<parallel>(A)\<parallel>) = (\<parallel>(A)\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11493
  have ih2: "pi\<bullet>(\<parallel><A>\<parallel>) = (\<parallel><A>\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11494
  have ih3: "pi\<bullet>(\<parallel>(B)\<parallel>) = (\<parallel>(B)\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11495
  have ih4: "pi\<bullet>(\<parallel><B>\<parallel>) = (\<parallel><B>\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11496
  have g: "pi\<bullet>(\<parallel>(A AND B)\<parallel>) = (\<parallel>(A AND B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11497
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11498
    apply(simp only: lfp_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11499
    apply(simp only: comp_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11500
    apply(simp only: perm_fun_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11501
    apply(simp only: NEGc.simps NEGn.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11502
    apply(simp only: union_eqvt AXIOMS_eqvt_name BINDING_eqvt_name ANDRIGHT_eqvt_name 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11503
                     ANDLEFT2_eqvt_name ANDLEFT1_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11504
    apply(perm_simp add: ih1 ih2 ih3 ih4)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11505
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11506
  { case 1 show ?case by (rule g)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11507
  next 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11508
    case 2 show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11509
      by (simp only: NEGc_simps union_eqvt AXIOMS_eqvt_name BINDING_eqvt_name 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11510
                     ANDRIGHT_eqvt_name ANDLEFT1_eqvt_name ANDLEFT2_eqvt_name ih1 ih2 ih3 ih4 g)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11511
  }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11512
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11513
  case (OR A B)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11514
  have ih1: "pi\<bullet>(\<parallel>(A)\<parallel>) = (\<parallel>(A)\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11515
  have ih2: "pi\<bullet>(\<parallel><A>\<parallel>) = (\<parallel><A>\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11516
  have ih3: "pi\<bullet>(\<parallel>(B)\<parallel>) = (\<parallel>(B)\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11517
  have ih4: "pi\<bullet>(\<parallel><B>\<parallel>) = (\<parallel><B>\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11518
  have g: "pi\<bullet>(\<parallel>(A OR B)\<parallel>) = (\<parallel>(A OR B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11519
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11520
    apply(simp only: lfp_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11521
    apply(simp only: comp_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11522
    apply(simp only: perm_fun_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11523
    apply(simp only: NEGc.simps NEGn.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11524
    apply(simp only: union_eqvt AXIOMS_eqvt_name BINDING_eqvt_name ORRIGHT1_eqvt_name 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11525
                     ORRIGHT2_eqvt_name ORLEFT_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11526
    apply(perm_simp add: ih1 ih2 ih3 ih4)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11527
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11528
  { case 1 show ?case by (rule g)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11529
  next 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11530
    case 2 show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11531
      by (simp only: NEGc_simps union_eqvt AXIOMS_eqvt_name BINDING_eqvt_name 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11532
                     ORRIGHT1_eqvt_name ORRIGHT2_eqvt_name ORLEFT_eqvt_name ih1 ih2 ih3 ih4 g)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11533
  }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11534
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11535
  case (IMP A B)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11536
  have ih1: "pi\<bullet>(\<parallel>(A)\<parallel>) = (\<parallel>(A)\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11537
  have ih2: "pi\<bullet>(\<parallel><A>\<parallel>) = (\<parallel><A>\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11538
  have ih3: "pi\<bullet>(\<parallel>(B)\<parallel>) = (\<parallel>(B)\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11539
  have ih4: "pi\<bullet>(\<parallel><B>\<parallel>) = (\<parallel><B>\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11540
  have g: "pi\<bullet>(\<parallel>(A IMP B)\<parallel>) = (\<parallel>(A IMP B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11541
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11542
    apply(simp only: lfp_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11543
    apply(simp only: comp_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11544
    apply(simp only: perm_fun_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11545
    apply(simp only: NEGc.simps NEGn.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11546
    apply(simp only: union_eqvt AXIOMS_eqvt_name BINDING_eqvt_name IMPRIGHT_eqvt_name IMPLEFT_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11547
    apply(perm_simp add: ih1 ih2 ih3 ih4)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11548
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11549
  { case 1 show ?case by (rule g)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11550
  next 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11551
    case 2 show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11552
      by (simp only: NEGc_simps union_eqvt AXIOMS_eqvt_name BINDING_eqvt_name 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11553
                     IMPRIGHT_eqvt_name IMPLEFT_eqvt_name ih1 ih2 ih3 ih4 g)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11554
  }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11555
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11556
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11557
lemma CAND_eqvt_coname:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11558
  fixes pi::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11559
  shows   "(pi\<bullet>(\<parallel>(B)\<parallel>)) = (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11560
  and     "(pi\<bullet>(\<parallel><B>\<parallel>)) = (\<parallel><B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11561
proof (nominal_induct B rule: ty.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11562
  case (PR X)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11563
  { case 1 show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11564
      apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11565
      apply(simp add: lfp_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11566
      apply(simp add: perm_fun_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11567
      apply(simp add: union_eqvt AXIOMS_eqvt_coname BINDING_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11568
      apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11569
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11570
  next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11571
    case 2 show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11572
      apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11573
      apply(simp only: NEGc_simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11574
      apply(simp add: union_eqvt AXIOMS_eqvt_coname BINDING_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11575
      apply(simp add: lfp_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11576
      apply(simp add: comp_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11577
      apply(simp add: perm_fun_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11578
      apply(simp add: union_eqvt AXIOMS_eqvt_coname BINDING_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11579
      apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11580
      done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11581
  }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11582
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11583
  case (NOT B)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11584
  have ih1: "pi\<bullet>(\<parallel>(B)\<parallel>) = (\<parallel>(B)\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11585
  have ih2: "pi\<bullet>(\<parallel><B>\<parallel>) = (\<parallel><B>\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11586
  have g: "pi\<bullet>(\<parallel>(NOT B)\<parallel>) = (\<parallel>(NOT B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11587
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11588
    apply(simp only: lfp_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11589
    apply(simp only: comp_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11590
    apply(simp only: perm_fun_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11591
    apply(simp only: NEGc.simps NEGn.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11592
    apply(simp only: union_eqvt AXIOMS_eqvt_coname BINDING_eqvt_coname 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11593
            NOTRIGHT_eqvt_coname NOTLEFT_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11594
    apply(perm_simp add: ih1 ih2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11595
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11596
  { case 1 show ?case by (rule g)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11597
  next 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11598
    case 2 show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11599
      by (simp only: NEGc_simps union_eqvt AXIOMS_eqvt_coname BINDING_eqvt_coname 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11600
              NOTRIGHT_eqvt_coname ih1 ih2 g)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11601
  }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11602
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11603
  case (AND A B)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11604
  have ih1: "pi\<bullet>(\<parallel>(A)\<parallel>) = (\<parallel>(A)\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11605
  have ih2: "pi\<bullet>(\<parallel><A>\<parallel>) = (\<parallel><A>\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11606
  have ih3: "pi\<bullet>(\<parallel>(B)\<parallel>) = (\<parallel>(B)\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11607
  have ih4: "pi\<bullet>(\<parallel><B>\<parallel>) = (\<parallel><B>\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11608
  have g: "pi\<bullet>(\<parallel>(A AND B)\<parallel>) = (\<parallel>(A AND B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11609
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11610
    apply(simp only: lfp_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11611
    apply(simp only: comp_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11612
    apply(simp only: perm_fun_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11613
    apply(simp only: NEGc.simps NEGn.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11614
    apply(simp only: union_eqvt AXIOMS_eqvt_coname BINDING_eqvt_coname ANDRIGHT_eqvt_coname 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11615
                     ANDLEFT2_eqvt_coname ANDLEFT1_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11616
    apply(perm_simp add: ih1 ih2 ih3 ih4)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11617
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11618
  { case 1 show ?case by (rule g)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11619
  next 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11620
    case 2 show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11621
      by (simp only: NEGc_simps union_eqvt AXIOMS_eqvt_coname BINDING_eqvt_coname 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11622
                     ANDRIGHT_eqvt_coname ANDLEFT1_eqvt_coname ANDLEFT2_eqvt_coname ih1 ih2 ih3 ih4 g)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11623
  }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11624
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11625
  case (OR A B)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11626
  have ih1: "pi\<bullet>(\<parallel>(A)\<parallel>) = (\<parallel>(A)\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11627
  have ih2: "pi\<bullet>(\<parallel><A>\<parallel>) = (\<parallel><A>\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11628
  have ih3: "pi\<bullet>(\<parallel>(B)\<parallel>) = (\<parallel>(B)\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11629
  have ih4: "pi\<bullet>(\<parallel><B>\<parallel>) = (\<parallel><B>\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11630
  have g: "pi\<bullet>(\<parallel>(A OR B)\<parallel>) = (\<parallel>(A OR B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11631
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11632
    apply(simp only: lfp_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11633
    apply(simp only: comp_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11634
    apply(simp only: perm_fun_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11635
    apply(simp only: NEGc.simps NEGn.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11636
    apply(simp only: union_eqvt AXIOMS_eqvt_coname BINDING_eqvt_coname ORRIGHT1_eqvt_coname 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11637
                     ORRIGHT2_eqvt_coname ORLEFT_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11638
    apply(perm_simp add: ih1 ih2 ih3 ih4)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11639
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11640
  { case 1 show ?case by (rule g)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11641
  next 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11642
    case 2 show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11643
      by (simp only: NEGc_simps union_eqvt AXIOMS_eqvt_coname BINDING_eqvt_coname 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11644
                     ORRIGHT1_eqvt_coname ORRIGHT2_eqvt_coname ORLEFT_eqvt_coname ih1 ih2 ih3 ih4 g)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11645
  }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11646
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11647
  case (IMP A B)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11648
  have ih1: "pi\<bullet>(\<parallel>(A)\<parallel>) = (\<parallel>(A)\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11649
  have ih2: "pi\<bullet>(\<parallel><A>\<parallel>) = (\<parallel><A>\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11650
  have ih3: "pi\<bullet>(\<parallel>(B)\<parallel>) = (\<parallel>(B)\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11651
  have ih4: "pi\<bullet>(\<parallel><B>\<parallel>) = (\<parallel><B>\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11652
  have g: "pi\<bullet>(\<parallel>(A IMP B)\<parallel>) = (\<parallel>(A IMP B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11653
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11654
    apply(simp only: lfp_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11655
    apply(simp only: comp_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11656
    apply(simp only: perm_fun_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11657
    apply(simp only: NEGc.simps NEGn.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11658
    apply(simp only: union_eqvt AXIOMS_eqvt_coname BINDING_eqvt_coname IMPRIGHT_eqvt_coname 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11659
         IMPLEFT_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11660
    apply(perm_simp add: ih1 ih2 ih3 ih4)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11661
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11662
  { case 1 show ?case by (rule g)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11663
  next 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11664
    case 2 show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11665
      by (simp only: NEGc_simps union_eqvt AXIOMS_eqvt_coname BINDING_eqvt_coname 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11666
                     IMPRIGHT_eqvt_coname IMPLEFT_eqvt_coname ih1 ih2 ih3 ih4 g)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11667
  }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11668
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11669
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11670
text {* Elimination rules for the set-operators *}
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11671
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11672
lemma BINDINGc_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11673
  assumes a: "<a>:M \<in> BINDINGc B (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11674
  shows "\<forall>x P. ((x):P)\<in>(\<parallel>(B)\<parallel>) \<longrightarrow> SNa (M{a:=(x).P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11675
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11676
apply(auto simp add: BINDINGc_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11677
apply(auto simp add: ctrm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11678
apply(drule_tac x="[(a,aa)]\<bullet>x" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11679
apply(drule_tac x="[(a,aa)]\<bullet>P" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11680
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11681
apply(drule_tac pi="[(a,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11682
apply(simp add: CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11683
apply(drule_tac ?pi2.0="[(a,aa)]" in SNa_eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11684
apply(perm_simp add: csubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11685
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11686
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11687
lemma BINDINGn_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11688
  assumes a: "(x):M \<in> BINDINGn B (\<parallel><B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11689
  shows "\<forall>c P. (<c>:P)\<in>(\<parallel><B>\<parallel>) \<longrightarrow> SNa (M{x:=<c>.P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11690
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11691
apply(auto simp add: BINDINGn_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11692
apply(auto simp add: ntrm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11693
apply(drule_tac x="[(x,xa)]\<bullet>c" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11694
apply(drule_tac x="[(x,xa)]\<bullet>P" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11695
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11696
apply(drule_tac pi="[(x,xa)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11697
apply(simp add: CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11698
apply(drule_tac ?pi1.0="[(x,xa)]" in SNa_eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11699
apply(perm_simp add: nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11700
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11701
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11702
lemma NOTRIGHT_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11703
  assumes a: "<a>:M \<in> NOTRIGHT (NOT B) (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11704
  obtains x' M' where "M = NotR (x').M' a" and "fic (NotR (x').M' a) a" and "(x'):M' \<in> (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11705
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11706
apply(auto simp add: ctrm.inject alpha abs_fresh calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11707
apply(drule_tac x="x" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11708
apply(drule_tac x="[(a,aa)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11709
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11710
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11711
apply(drule_tac pi="[(a,aa)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11712
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11713
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11714
apply(drule_tac pi="[(a,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11715
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11716
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11717
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11718
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11719
lemma NOTLEFT_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11720
  assumes a: "(x):M \<in> NOTLEFT (NOT B) (\<parallel><B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11721
  obtains a' M' where "M = NotL <a'>.M' x" and "fin (NotL <a'>.M' x) x" and "<a'>:M' \<in> (\<parallel><B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11722
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11723
apply(auto simp add: ntrm.inject alpha abs_fresh calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11724
apply(drule_tac x="a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11725
apply(drule_tac x="[(x,xa)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11726
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11727
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11728
apply(drule_tac pi="[(x,xa)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11729
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11730
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11731
apply(drule_tac pi="[(x,xa)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11732
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11733
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11734
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11735
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11736
lemma ANDRIGHT_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11737
  assumes a: "<a>:M \<in> ANDRIGHT (B AND C) (\<parallel><B>\<parallel>) (\<parallel><C>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11738
  obtains d' M' e' N' where "M = AndR <d'>.M' <e'>.N' a" and "fic (AndR <d'>.M' <e'>.N' a) a" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11739
                      and "<d'>:M' \<in> (\<parallel><B>\<parallel>)" and "<e'>:N' \<in> (\<parallel><C>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11740
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11741
apply(auto simp add: ctrm.inject alpha abs_fresh calc_atm fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11742
apply(drule_tac x="c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11743
apply(drule_tac x="[(a,c)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11744
apply(drule_tac x="c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11745
apply(drule_tac x="[(a,c)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11746
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11747
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11748
apply(drule_tac pi="[(a,c)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11749
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11750
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11751
apply(drule_tac pi="[(a,c)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11752
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11753
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11754
apply(drule_tac pi="[(a,c)]" and x="<a>:N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11755
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11756
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11757
apply(case_tac "a=b")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11758
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11759
apply(drule_tac x="c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11760
apply(drule_tac x="[(b,c)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11761
apply(drule_tac x="c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11762
apply(drule_tac x="[(b,c)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11763
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11764
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11765
apply(drule_tac pi="[(b,c)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11766
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11767
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11768
apply(drule_tac pi="[(b,c)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11769
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11770
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11771
apply(drule_tac pi="[(b,c)]" and x="<b>:N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11772
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11773
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11774
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11775
apply(case_tac "c=b")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11776
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11777
apply(drule_tac x="b" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11778
apply(drule_tac x="[(a,b)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11779
apply(drule_tac x="a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11780
apply(drule_tac x="[(a,b)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11781
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11782
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11783
apply(drule_tac pi="[(a,b)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11784
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11785
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11786
apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11787
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11788
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11789
apply(drule_tac pi="[(a,b)]" and x="<b>:N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11790
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11791
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11792
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11793
apply(drule_tac x="c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11794
apply(drule_tac x="[(a,c)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11795
apply(drule_tac x="b" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11796
apply(drule_tac x="[(a,c)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11797
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11798
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11799
apply(drule_tac pi="[(a,c)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11800
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11801
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11802
apply(drule_tac pi="[(a,c)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11803
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11804
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11805
apply(drule_tac pi="[(a,c)]" and x="<b>:N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11806
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11807
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11808
apply(case_tac "a=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11809
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11810
apply(drule_tac x="c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11811
apply(drule_tac x="[(aa,c)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11812
apply(drule_tac x="c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11813
apply(drule_tac x="[(aa,c)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11814
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11815
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11816
apply(drule_tac pi="[(aa,c)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11817
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11818
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11819
apply(drule_tac pi="[(aa,c)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11820
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11821
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11822
apply(drule_tac pi="[(aa,c)]" and x="<aa>:N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11823
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11824
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11825
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11826
apply(case_tac "c=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11827
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11828
apply(drule_tac x="a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11829
apply(drule_tac x="[(a,aa)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11830
apply(drule_tac x="aa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11831
apply(drule_tac x="[(a,aa)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11832
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11833
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11834
apply(drule_tac pi="[(a,aa)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11835
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11836
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11837
apply(drule_tac pi="[(a,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11838
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11839
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11840
apply(drule_tac pi="[(a,aa)]" and x="<a>:N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11841
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11842
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11843
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11844
apply(drule_tac x="aa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11845
apply(drule_tac x="[(a,c)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11846
apply(drule_tac x="c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11847
apply(drule_tac x="[(a,c)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11848
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11849
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11850
apply(drule_tac pi="[(a,c)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11851
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11852
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11853
apply(drule_tac pi="[(a,c)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11854
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11855
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11856
apply(drule_tac pi="[(a,c)]" and x="<a>:N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11857
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11858
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11859
apply(case_tac "a=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11860
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11861
apply(case_tac "aa=b")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11862
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11863
apply(drule_tac x="c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11864
apply(drule_tac x="[(b,c)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11865
apply(drule_tac x="c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11866
apply(drule_tac x="[(b,c)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11867
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11868
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11869
apply(drule_tac pi="[(b,c)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11870
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11871
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11872
apply(drule_tac pi="[(b,c)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11873
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11874
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11875
apply(drule_tac pi="[(b,c)]" and x="<b>:N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11876
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11877
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11878
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11879
apply(case_tac "c=b")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11880
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11881
apply(drule_tac x="b" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11882
apply(drule_tac x="[(aa,b)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11883
apply(drule_tac x="aa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11884
apply(drule_tac x="[(aa,b)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11885
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11886
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11887
apply(drule_tac pi="[(aa,b)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11888
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11889
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11890
apply(drule_tac pi="[(aa,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11891
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11892
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11893
apply(drule_tac pi="[(aa,b)]" and x="<b>:N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11894
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11895
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11896
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11897
apply(drule_tac x="c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11898
apply(drule_tac x="[(aa,c)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11899
apply(drule_tac x="b" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11900
apply(drule_tac x="[(aa,c)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11901
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11902
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11903
apply(drule_tac pi="[(aa,c)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11904
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11905
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11906
apply(drule_tac pi="[(aa,c)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11907
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11908
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11909
apply(drule_tac pi="[(aa,c)]" and x="<b>:N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11910
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11911
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11912
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11913
apply(case_tac "c=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11914
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11915
apply(case_tac "a=b")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11916
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11917
apply(drule_tac x="b" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11918
apply(drule_tac x="[(b,aa)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11919
apply(drule_tac x="aa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11920
apply(drule_tac x="[(b,aa)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11921
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11922
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11923
apply(drule_tac pi="[(b,aa)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11924
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11925
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11926
apply(drule_tac pi="[(b,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11927
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11928
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11929
apply(drule_tac pi="[(b,aa)]" and x="<b>:N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11930
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11931
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11932
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11933
apply(case_tac "aa=b")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11934
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11935
apply(drule_tac x="a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11936
apply(drule_tac x="[(a,b)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11937
apply(drule_tac x="a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11938
apply(drule_tac x="[(a,b)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11939
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11940
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11941
apply(drule_tac pi="[(a,b)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11942
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11943
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11944
apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11945
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11946
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11947
apply(drule_tac pi="[(a,b)]" and x="<b>:N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11948
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11949
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11950
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11951
apply(drule_tac x="a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11952
apply(drule_tac x="[(a,aa)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11953
apply(drule_tac x="b" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11954
apply(drule_tac x="[(a,aa)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11955
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11956
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11957
apply(drule_tac pi="[(a,aa)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11958
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11959
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11960
apply(drule_tac pi="[(a,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11961
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11962
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11963
apply(drule_tac pi="[(a,aa)]" and x="<b>:N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11964
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11965
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11966
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11967
apply(case_tac "a=b")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11968
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11969
apply(drule_tac x="aa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11970
apply(drule_tac x="[(b,c)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11971
apply(drule_tac x="c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11972
apply(drule_tac x="[(b,c)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11973
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11974
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11975
apply(drule_tac pi="[(b,c)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11976
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11977
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11978
apply(drule_tac pi="[(b,c)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11979
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11980
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11981
apply(drule_tac pi="[(b,c)]" and x="<b>:N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11982
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11983
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11984
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11985
apply(case_tac "c=b")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11986
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11987
apply(drule_tac x="aa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11988
apply(drule_tac x="[(a,b)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11989
apply(drule_tac x="a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11990
apply(drule_tac x="[(a,b)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11991
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11992
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11993
apply(drule_tac pi="[(a,b)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11994
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11995
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11996
apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11997
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11998
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 11999
apply(drule_tac pi="[(a,b)]" and x="<b>:N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12000
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12001
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12002
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12003
apply(drule_tac x="aa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12004
apply(drule_tac x="[(a,c)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12005
apply(drule_tac x="b" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12006
apply(drule_tac x="[(a,c)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12007
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12008
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12009
apply(drule_tac pi="[(a,c)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12010
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12011
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12012
apply(drule_tac pi="[(a,c)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12013
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12014
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12015
apply(drule_tac pi="[(a,c)]" and x="<b>:N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12016
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12017
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12018
done 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12019
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12020
lemma ANDLEFT1_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12021
  assumes a: "(x):M \<in> ANDLEFT1 (B AND C) (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12022
  obtains x' M' where "M = AndL1 (x').M' x" and "fin (AndL1 (x').M' x) x" and "(x'):M' \<in> (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12023
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12024
apply(auto simp add: ntrm.inject alpha abs_fresh calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12025
apply(drule_tac x="y" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12026
apply(drule_tac x="[(x,y)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12027
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12028
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12029
apply(drule_tac pi="[(x,y)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12030
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12031
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12032
apply(drule_tac pi="[(x,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12033
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12034
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12035
apply(case_tac "x=xa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12036
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12037
apply(drule_tac x="y" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12038
apply(drule_tac x="[(xa,y)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12039
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12040
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12041
apply(drule_tac pi="[(xa,y)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12042
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12043
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12044
apply(drule_tac pi="[(xa,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12045
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12046
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12047
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12048
apply(case_tac "y=xa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12049
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12050
apply(drule_tac x="x" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12051
apply(drule_tac x="[(x,xa)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12052
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12053
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12054
apply(drule_tac pi="[(x,xa)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12055
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12056
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12057
apply(drule_tac pi="[(x,xa)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12058
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12059
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12060
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12061
apply(drule_tac x="xa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12062
apply(drule_tac x="[(x,y)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12063
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12064
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12065
apply(drule_tac pi="[(x,y)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12066
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12067
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12068
apply(drule_tac pi="[(x,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12069
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12070
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12071
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12072
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12073
lemma ANDLEFT2_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12074
  assumes a: "(x):M \<in> ANDLEFT2 (B AND C) (\<parallel>(C)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12075
  obtains x' M' where "M = AndL2 (x').M' x" and "fin (AndL2 (x').M' x) x" and "(x'):M' \<in> (\<parallel>(C)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12076
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12077
apply(auto simp add: ntrm.inject alpha abs_fresh calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12078
apply(drule_tac x="y" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12079
apply(drule_tac x="[(x,y)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12080
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12081
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12082
apply(drule_tac pi="[(x,y)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12083
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12084
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12085
apply(drule_tac pi="[(x,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12086
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12087
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12088
apply(case_tac "x=xa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12089
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12090
apply(drule_tac x="y" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12091
apply(drule_tac x="[(xa,y)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12092
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12093
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12094
apply(drule_tac pi="[(xa,y)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12095
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12096
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12097
apply(drule_tac pi="[(xa,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12098
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12099
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12100
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12101
apply(case_tac "y=xa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12102
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12103
apply(drule_tac x="x" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12104
apply(drule_tac x="[(x,xa)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12105
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12106
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12107
apply(drule_tac pi="[(x,xa)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12108
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12109
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12110
apply(drule_tac pi="[(x,xa)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12111
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12112
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12113
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12114
apply(drule_tac x="xa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12115
apply(drule_tac x="[(x,y)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12116
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12117
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12118
apply(drule_tac pi="[(x,y)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12119
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12120
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12121
apply(drule_tac pi="[(x,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12122
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12123
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12124
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12125
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12126
lemma ORRIGHT1_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12127
  assumes a: "<a>:M \<in> ORRIGHT1 (B OR C) (\<parallel><B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12128
  obtains a' M' where "M = OrR1 <a'>.M' a" and "fic (OrR1 <a'>.M' a) a" and "<a'>:M' \<in> (\<parallel><B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12129
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12130
apply(auto simp add: ctrm.inject alpha abs_fresh calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12131
apply(drule_tac x="b" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12132
apply(drule_tac x="[(a,b)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12133
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12134
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12135
apply(drule_tac pi="[(a,b)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12136
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12137
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12138
apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12139
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12140
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12141
apply(case_tac "a=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12142
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12143
apply(drule_tac x="b" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12144
apply(drule_tac x="[(aa,b)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12145
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12146
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12147
apply(drule_tac pi="[(aa,b)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12148
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12149
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12150
apply(drule_tac pi="[(aa,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12151
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12152
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12153
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12154
apply(case_tac "b=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12155
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12156
apply(drule_tac x="a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12157
apply(drule_tac x="[(a,aa)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12158
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12159
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12160
apply(drule_tac pi="[(a,aa)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12161
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12162
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12163
apply(drule_tac pi="[(a,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12164
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12165
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12166
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12167
apply(drule_tac x="aa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12168
apply(drule_tac x="[(a,b)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12169
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12170
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12171
apply(drule_tac pi="[(a,b)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12172
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12173
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12174
apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12175
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12176
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12177
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12178
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12179
lemma ORRIGHT2_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12180
  assumes a: "<a>:M \<in> ORRIGHT2 (B OR C) (\<parallel><C>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12181
  obtains a' M' where "M = OrR2 <a'>.M' a" and "fic (OrR2 <a'>.M' a) a" and "<a'>:M' \<in> (\<parallel><C>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12182
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12183
apply(auto simp add: ctrm.inject alpha abs_fresh calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12184
apply(drule_tac x="b" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12185
apply(drule_tac x="[(a,b)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12186
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12187
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12188
apply(drule_tac pi="[(a,b)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12189
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12190
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12191
apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12192
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12193
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12194
apply(case_tac "a=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12195
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12196
apply(drule_tac x="b" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12197
apply(drule_tac x="[(aa,b)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12198
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12199
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12200
apply(drule_tac pi="[(aa,b)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12201
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12202
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12203
apply(drule_tac pi="[(aa,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12204
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12205
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12206
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12207
apply(case_tac "b=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12208
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12209
apply(drule_tac x="a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12210
apply(drule_tac x="[(a,aa)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12211
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12212
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12213
apply(drule_tac pi="[(a,aa)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12214
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12215
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12216
apply(drule_tac pi="[(a,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12217
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12218
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12219
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12220
apply(drule_tac x="aa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12221
apply(drule_tac x="[(a,b)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12222
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12223
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12224
apply(drule_tac pi="[(a,b)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12225
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12226
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12227
apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12228
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12229
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12230
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12231
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12232
lemma ORLEFT_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12233
  assumes a: "(x):M \<in> ORLEFT (B OR C) (\<parallel>(B)\<parallel>) (\<parallel>(C)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12234
  obtains y' M' z' N' where "M = OrL (y').M' (z').N' x" and "fin (OrL (y').M' (z').N' x) x" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12235
                      and "(y'):M' \<in> (\<parallel>(B)\<parallel>)" and "(z'):N' \<in> (\<parallel>(C)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12236
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12237
apply(auto simp add: ntrm.inject alpha abs_fresh calc_atm fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12238
apply(drule_tac x="z" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12239
apply(drule_tac x="[(x,z)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12240
apply(drule_tac x="z" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12241
apply(drule_tac x="[(x,z)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12242
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12243
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12244
apply(drule_tac pi="[(x,z)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12245
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12246
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12247
apply(drule_tac pi="[(x,z)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12248
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12249
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12250
apply(drule_tac pi="[(x,z)]" and x="(x):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12251
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12252
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12253
apply(case_tac "x=y")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12254
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12255
apply(drule_tac x="z" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12256
apply(drule_tac x="[(y,z)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12257
apply(drule_tac x="z" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12258
apply(drule_tac x="[(y,z)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12259
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12260
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12261
apply(drule_tac pi="[(y,z)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12262
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12263
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12264
apply(drule_tac pi="[(y,z)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12265
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12266
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12267
apply(drule_tac pi="[(y,z)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12268
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12269
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12270
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12271
apply(case_tac "z=y")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12272
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12273
apply(drule_tac x="y" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12274
apply(drule_tac x="[(x,y)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12275
apply(drule_tac x="x" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12276
apply(drule_tac x="[(x,y)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12277
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12278
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12279
apply(drule_tac pi="[(x,y)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12280
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12281
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12282
apply(drule_tac pi="[(x,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12283
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12284
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12285
apply(drule_tac pi="[(x,y)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12286
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12287
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12288
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12289
apply(drule_tac x="z" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12290
apply(drule_tac x="[(x,z)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12291
apply(drule_tac x="y" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12292
apply(drule_tac x="[(x,z)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12293
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12294
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12295
apply(drule_tac pi="[(x,z)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12296
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12297
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12298
apply(drule_tac pi="[(x,z)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12299
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12300
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12301
apply(drule_tac pi="[(x,z)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12302
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12303
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12304
apply(case_tac "x=xa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12305
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12306
apply(drule_tac x="z" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12307
apply(drule_tac x="[(xa,z)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12308
apply(drule_tac x="z" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12309
apply(drule_tac x="[(xa,z)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12310
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12311
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12312
apply(drule_tac pi="[(xa,z)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12313
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12314
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12315
apply(drule_tac pi="[(xa,z)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12316
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12317
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12318
apply(drule_tac pi="[(xa,z)]" and x="(xa):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12319
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12320
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12321
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12322
apply(case_tac "z=xa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12323
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12324
apply(drule_tac x="x" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12325
apply(drule_tac x="[(x,xa)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12326
apply(drule_tac x="xa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12327
apply(drule_tac x="[(x,xa)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12328
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12329
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12330
apply(drule_tac pi="[(x,xa)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12331
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12332
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12333
apply(drule_tac pi="[(x,xa)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12334
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12335
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12336
apply(drule_tac pi="[(x,xa)]" and x="(x):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12337
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12338
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12339
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12340
apply(drule_tac x="xa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12341
apply(drule_tac x="[(x,z)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12342
apply(drule_tac x="z" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12343
apply(drule_tac x="[(x,z)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12344
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12345
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12346
apply(drule_tac pi="[(x,z)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12347
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12348
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12349
apply(drule_tac pi="[(x,z)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12350
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12351
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12352
apply(drule_tac pi="[(x,z)]" and x="(x):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12353
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12354
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12355
apply(case_tac "x=xa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12356
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12357
apply(case_tac "xa=y")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12358
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12359
apply(drule_tac x="z" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12360
apply(drule_tac x="[(y,z)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12361
apply(drule_tac x="z" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12362
apply(drule_tac x="[(y,z)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12363
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12364
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12365
apply(drule_tac pi="[(y,z)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12366
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12367
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12368
apply(drule_tac pi="[(y,z)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12369
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12370
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12371
apply(drule_tac pi="[(y,z)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12372
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12373
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12374
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12375
apply(case_tac "z=y")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12376
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12377
apply(drule_tac x="y" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12378
apply(drule_tac x="[(xa,y)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12379
apply(drule_tac x="xa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12380
apply(drule_tac x="[(xa,y)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12381
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12382
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12383
apply(drule_tac pi="[(xa,y)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12384
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12385
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12386
apply(drule_tac pi="[(xa,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12387
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12388
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12389
apply(drule_tac pi="[(xa,y)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12390
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12391
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12392
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12393
apply(drule_tac x="z" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12394
apply(drule_tac x="[(xa,z)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12395
apply(drule_tac x="y" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12396
apply(drule_tac x="[(xa,z)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12397
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12398
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12399
apply(drule_tac pi="[(xa,z)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12400
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12401
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12402
apply(drule_tac pi="[(xa,z)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12403
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12404
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12405
apply(drule_tac pi="[(xa,z)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12406
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12407
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12408
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12409
apply(case_tac "z=xa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12410
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12411
apply(case_tac "x=y")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12412
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12413
apply(drule_tac x="y" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12414
apply(drule_tac x="[(y,xa)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12415
apply(drule_tac x="xa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12416
apply(drule_tac x="[(y,xa)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12417
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12418
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12419
apply(drule_tac pi="[(y,xa)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12420
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12421
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12422
apply(drule_tac pi="[(y,xa)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12423
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12424
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12425
apply(drule_tac pi="[(y,xa)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12426
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12427
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12428
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12429
apply(case_tac "xa=y")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12430
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12431
apply(drule_tac x="x" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12432
apply(drule_tac x="[(x,y)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12433
apply(drule_tac x="x" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12434
apply(drule_tac x="[(x,y)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12435
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12436
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12437
apply(drule_tac pi="[(x,y)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12438
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12439
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12440
apply(drule_tac pi="[(x,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12441
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12442
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12443
apply(drule_tac pi="[(x,y)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12444
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12445
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12446
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12447
apply(drule_tac x="x" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12448
apply(drule_tac x="[(x,xa)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12449
apply(drule_tac x="y" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12450
apply(drule_tac x="[(x,xa)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12451
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12452
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12453
apply(drule_tac pi="[(x,xa)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12454
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12455
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12456
apply(drule_tac pi="[(x,xa)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12457
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12458
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12459
apply(drule_tac pi="[(x,xa)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12460
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12461
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12462
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12463
apply(case_tac "x=y")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12464
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12465
apply(drule_tac x="xa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12466
apply(drule_tac x="[(y,z)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12467
apply(drule_tac x="z" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12468
apply(drule_tac x="[(y,z)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12469
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12470
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12471
apply(drule_tac pi="[(y,z)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12472
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12473
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12474
apply(drule_tac pi="[(y,z)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12475
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12476
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12477
apply(drule_tac pi="[(y,z)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12478
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12479
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12480
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12481
apply(case_tac "z=y")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12482
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12483
apply(drule_tac x="xa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12484
apply(drule_tac x="[(x,y)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12485
apply(drule_tac x="x" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12486
apply(drule_tac x="[(x,y)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12487
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12488
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12489
apply(drule_tac pi="[(x,y)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12490
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12491
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12492
apply(drule_tac pi="[(x,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12493
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12494
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12495
apply(drule_tac pi="[(x,y)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12496
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12497
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12498
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12499
apply(drule_tac x="xa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12500
apply(drule_tac x="[(x,z)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12501
apply(drule_tac x="y" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12502
apply(drule_tac x="[(x,z)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12503
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12504
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12505
apply(drule_tac pi="[(x,z)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12506
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12507
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12508
apply(drule_tac pi="[(x,z)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12509
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12510
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12511
apply(drule_tac pi="[(x,z)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12512
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12513
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12514
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12515
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12516
lemma IMPRIGHT_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12517
  assumes a: "<a>:M \<in> IMPRIGHT (B IMP C) (\<parallel>(B)\<parallel>) (\<parallel><C>\<parallel>) (\<parallel>(C)\<parallel>) (\<parallel><B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12518
  obtains x' a' M' where "M = ImpR (x').<a'>.M' a" and "fic (ImpR (x').<a'>.M' a) a" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12519
                   and "\<forall>z P. x'\<sharp>(z,P) \<and> (z):P \<in> \<parallel>(C)\<parallel> \<longrightarrow> (x'):(M'{a':=(z).P}) \<in> \<parallel>(B)\<parallel>" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12520
                   and "\<forall>c Q. a'\<sharp>(c,Q) \<and> <c>:Q \<in> \<parallel><B>\<parallel> \<longrightarrow> <a'>:(M'{x':=<c>.Q}) \<in> \<parallel><C>\<parallel>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12521
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12522
apply(auto simp add: ctrm.inject alpha abs_fresh calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12523
apply(drule_tac x="x" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12524
apply(drule_tac x="b" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12525
apply(drule_tac x="[(a,b)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12526
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12527
apply(drule_tac pi="[(a,b)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12528
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12529
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12530
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12531
apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12532
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12533
apply(drule_tac x="z" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12534
apply(drule_tac x="[(a,b)]\<bullet>P" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12535
apply(simp add: fresh_prod fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12536
apply(drule_tac pi="[(a,b)]" and x="(x):M{a:=(z).([(a,b)]\<bullet>P)}" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12537
                                     in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12538
apply(perm_simp add: calc_atm csubst_eqvt CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12539
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12540
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12541
apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12542
apply(simp add:  CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12543
apply(rotate_tac 2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12544
apply(drule_tac x="[(a,b)]\<bullet>c" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12545
apply(drule_tac x="[(a,b)]\<bullet>Q" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12546
apply(simp add: fresh_prod fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12547
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12548
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12549
apply(drule_tac pi="[(a,b)]" and x="<a>:M{x:=<([(a,b)]\<bullet>c)>.([(a,b)]\<bullet>Q)}" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12550
                                        in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12551
apply(perm_simp add: nsubst_eqvt CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12552
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12553
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12554
apply(case_tac "a=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12555
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12556
apply(drule_tac x="x" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12557
apply(drule_tac x="b" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12558
apply(drule_tac x="[(aa,b)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12559
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12560
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12561
apply(drule_tac pi="[(aa,b)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12562
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12563
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12564
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12565
apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12566
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12567
apply(drule_tac x="z" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12568
apply(drule_tac x="[(a,b)]\<bullet>P" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12569
apply(simp add: fresh_prod fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12570
apply(drule_tac pi="[(a,b)]" and x="(x):M{a:=(z).([(a,b)]\<bullet>P)}" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12571
                                     in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12572
apply(perm_simp add: calc_atm csubst_eqvt  CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12573
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12574
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12575
apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12576
apply(simp add: CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12577
apply(drule_tac x="[(a,b)]\<bullet>c" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12578
apply(drule_tac x="[(a,b)]\<bullet>Q" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12579
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12580
apply(simp add: fresh_prod fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12581
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12582
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12583
apply(drule_tac pi="[(a,b)]" and x="<a>:M{x:=<([(a,b)]\<bullet>c)>.([(a,b)]\<bullet>Q)}" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12584
                                      in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12585
apply(perm_simp add: nsubst_eqvt CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12586
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12587
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12588
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12589
apply(case_tac "b=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12590
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12591
apply(drule_tac x="x" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12592
apply(drule_tac x="a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12593
apply(drule_tac x="[(a,aa)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12594
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12595
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12596
apply(drule_tac pi="[(a,aa)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12597
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12598
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12599
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12600
apply(drule_tac pi="[(a,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12601
apply(simp add: calc_atm CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12602
apply(drule_tac x="z" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12603
apply(drule_tac x="[(a,aa)]\<bullet>P" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12604
apply(simp add: fresh_prod fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12605
apply(drule_tac pi="[(a,aa)]" and x="(x):M{aa:=(z).([(a,aa)]\<bullet>P)}" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12606
                                    in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12607
apply(perm_simp add: calc_atm csubst_eqvt  CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12608
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12609
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12610
apply(drule_tac pi="[(a,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12611
apply(simp add:  CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12612
apply(drule_tac x="[(a,aa)]\<bullet>c" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12613
apply(drule_tac x="[(a,aa)]\<bullet>Q" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12614
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12615
apply(simp add: fresh_prod fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12616
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12617
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12618
apply(drule_tac pi="[(a,aa)]" and x="<aa>:M{x:=<([(a,aa)]\<bullet>c)>.([(a,aa)]\<bullet>Q)}" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12619
                                    in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12620
apply(perm_simp add: nsubst_eqvt  CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12621
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12622
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12623
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12624
apply(drule_tac x="x" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12625
apply(drule_tac x="aa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12626
apply(drule_tac x="[(a,b)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12627
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12628
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12629
apply(drule_tac pi="[(a,b)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12630
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12631
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12632
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12633
apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12634
apply(simp add: calc_atm  CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12635
apply(drule_tac x="z" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12636
apply(drule_tac x="[(a,b)]\<bullet>P" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12637
apply(simp add: fresh_prod fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12638
apply(drule_tac pi="[(a,b)]" and x="(x):M{aa:=(z).([(a,b)]\<bullet>P)}" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12639
                                          in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12640
apply(perm_simp add: calc_atm csubst_eqvt  CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12641
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12642
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12643
apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12644
apply(simp add:  CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12645
apply(drule_tac x="[(a,b)]\<bullet>c" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12646
apply(drule_tac x="[(a,b)]\<bullet>Q" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12647
apply(simp add: fresh_prod fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12648
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12649
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12650
apply(drule_tac pi="[(a,b)]" and x="<aa>:M{x:=<([(a,b)]\<bullet>c)>.([(a,b)]\<bullet>Q)}" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12651
                                        in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12652
apply(perm_simp add: nsubst_eqvt  CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12653
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12654
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12655
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12656
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12657
lemma IMPLEFT_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12658
  assumes a: "(x):M \<in> IMPLEFT (B IMP C) (\<parallel><B>\<parallel>) (\<parallel>(C)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12659
  obtains x' a' M' N' where "M = ImpL <a'>.M' (x').N' x" and "fin (ImpL <a'>.M' (x').N' x) x" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12660
                   and "<a'>:M' \<in> \<parallel><B>\<parallel>" and "(x'):N' \<in> \<parallel>(C)\<parallel>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12661
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12662
apply(auto simp add: ntrm.inject alpha abs_fresh calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12663
apply(drule_tac x="a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12664
apply(drule_tac x="[(x,y)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12665
apply(drule_tac x="y" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12666
apply(drule_tac x="[(x,y)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12667
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12668
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12669
apply(drule_tac pi="[(x,y)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12670
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12671
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12672
apply(drule_tac pi="[(x,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12673
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12674
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12675
apply(drule_tac pi="[(x,y)]" and x="(x):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12676
apply(perm_simp add: calc_atm  CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12677
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12678
apply(case_tac "x=xa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12679
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12680
apply(drule_tac x="a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12681
apply(drule_tac x="[(xa,y)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12682
apply(drule_tac x="y" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12683
apply(drule_tac x="[(xa,y)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12684
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12685
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12686
apply(drule_tac pi="[(xa,y)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12687
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12688
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12689
apply(drule_tac pi="[(xa,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12690
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12691
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12692
apply(drule_tac pi="[(xa,y)]" and x="(xa):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12693
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12694
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12695
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12696
apply(case_tac "y=xa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12697
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12698
apply(drule_tac x="a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12699
apply(drule_tac x="[(x,xa)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12700
apply(drule_tac x="x" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12701
apply(drule_tac x="[(x,xa)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12702
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12703
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12704
apply(drule_tac pi="[(x,xa)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12705
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12706
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12707
apply(drule_tac pi="[(x,xa)]" and x="<a>:M" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12708
apply(simp add: calc_atm  CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12709
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12710
apply(drule_tac pi="[(x,xa)]" and x="(xa):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12711
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12712
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12713
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12714
apply(drule_tac x="a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12715
apply(drule_tac x="[(x,y)]\<bullet>M" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12716
apply(drule_tac x="xa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12717
apply(drule_tac x="[(x,y)]\<bullet>N" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12718
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12719
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12720
apply(drule_tac pi="[(x,y)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12721
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12722
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12723
apply(drule_tac pi="[(x,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12724
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12725
apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12726
apply(drule_tac pi="[(x,y)]" and x="(xa):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12727
apply(simp add: calc_atm CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12728
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12729
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12730
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12731
lemma CANDs_alpha:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12732
  shows "<a>:M \<in> (\<parallel><B>\<parallel>) \<Longrightarrow> [a].M = [b].N \<Longrightarrow> <b>:N \<in> (\<parallel><B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12733
  and   "(x):M \<in> (\<parallel>(B)\<parallel>) \<Longrightarrow> [x].M = [y].N \<Longrightarrow> (y):N \<in> (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12734
apply(auto simp add: alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12735
apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12736
apply(perm_simp add: CAND_eqvt_coname calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12737
apply(drule_tac pi="[(x,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12738
apply(perm_simp add: CAND_eqvt_name calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12739
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12740
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12741
lemma CAND_NotR_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12742
  assumes a: "<a>:NotR (x).M a \<in> (\<parallel><B>\<parallel>)" "<a>:NotR (x).M a \<notin> BINDINGc B (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12743
  shows "\<exists>B'. B = NOT B' \<and> (x):M \<in> (\<parallel>(B')\<parallel>)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12744
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12745
apply(nominal_induct B rule: ty.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12746
apply(simp_all add: ty.inject AXIOMSc_def ctrm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12747
apply(auto intro: CANDs_alpha simp add: trm.inject calc_atm abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12748
apply(drule_tac pi="[(a,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12749
apply(auto simp add: CAND_eqvt_coname calc_atm intro: CANDs_alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12750
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12751
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12752
lemma CAND_NotL_elim_aux:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12753
  assumes a: "(x):NotL <a>.M x \<in> NEGn B (\<parallel><B>\<parallel>)" "(x):NotL <a>.M x \<notin> BINDINGn B (\<parallel><B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12754
  shows "\<exists>B'. B = NOT B' \<and> <a>:M \<in> (\<parallel><B'>\<parallel>)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12755
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12756
apply(nominal_induct B rule: ty.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12757
apply(simp_all add: ty.inject AXIOMSn_def ntrm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12758
apply(auto intro: CANDs_alpha simp add: trm.inject calc_atm abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12759
apply(drule_tac pi="[(x,xa)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12760
apply(auto simp add: CAND_eqvt_name calc_atm intro: CANDs_alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12761
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12762
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12763
lemmas CAND_NotL_elim = CAND_NotL_elim_aux[OF NEG_elim(2)]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12764
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12765
lemma CAND_AndR_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12766
  assumes a: "<a>:AndR <b>.M <c>.N a \<in> (\<parallel><B>\<parallel>)" "<a>:AndR <b>.M <c>.N a \<notin> BINDINGc B (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12767
  shows "\<exists>B1 B2. B = B1 AND B2 \<and> <b>:M \<in> (\<parallel><B1>\<parallel>) \<and> <c>:N \<in> (\<parallel><B2>\<parallel>)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12768
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12769
apply(nominal_induct B rule: ty.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12770
apply(simp_all add: ty.inject AXIOMSc_def ctrm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12771
apply(auto intro: CANDs_alpha simp add: trm.inject calc_atm abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12772
apply(drule_tac pi="[(a,ca)]" and x="<a>:Ma" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12773
apply(simp add: CAND_eqvt_coname calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12774
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12775
apply(drule_tac pi="[(a,ca)]" and x="<a>:Na" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12776
apply(simp add: CAND_eqvt_coname calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12777
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12778
apply(drule_tac pi="[(a,ca)]" and x="<a>:Ma" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12779
apply(simp add: CAND_eqvt_coname calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12780
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12781
apply(case_tac "a=ba")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12782
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12783
apply(drule_tac pi="[(ba,ca)]" and x="<ba>:Na" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12784
apply(simp add: CAND_eqvt_coname calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12785
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12786
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12787
apply(case_tac "ca=ba")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12788
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12789
apply(drule_tac pi="[(a,ba)]" and x="<ba>:Na" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12790
apply(simp add: CAND_eqvt_coname calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12791
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12792
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12793
apply(drule_tac pi="[(a,ca)]" and x="<ba>:Na" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12794
apply(simp add: CAND_eqvt_coname calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12795
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12796
apply(case_tac "a=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12797
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12798
apply(drule_tac pi="[(aa,ca)]" and x="<aa>:Ma" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12799
apply(simp add: CAND_eqvt_coname calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12800
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12801
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12802
apply(case_tac "ca=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12803
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12804
apply(drule_tac pi="[(a,aa)]" and x="<aa>:Ma" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12805
apply(simp add: CAND_eqvt_coname calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12806
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12807
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12808
apply(drule_tac pi="[(a,ca)]" and x="<aa>:Ma" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12809
apply(simp add: CAND_eqvt_coname calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12810
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12811
apply(drule_tac pi="[(a,ca)]" and x="<a>:Na" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12812
apply(simp add: CAND_eqvt_coname calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12813
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12814
apply(case_tac "a=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12815
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12816
apply(drule_tac pi="[(aa,ca)]" and x="<aa>:Ma" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12817
apply(simp add: CAND_eqvt_coname calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12818
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12819
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12820
apply(case_tac "ca=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12821
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12822
apply(drule_tac pi="[(a,aa)]" and x="<aa>:Ma" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12823
apply(simp add: CAND_eqvt_coname calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12824
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12825
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12826
apply(drule_tac pi="[(a,ca)]" and x="<aa>:Ma" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12827
apply(simp add: CAND_eqvt_coname calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12828
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12829
apply(case_tac "a=ba")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12830
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12831
apply(drule_tac pi="[(ba,ca)]" and x="<ba>:Na" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12832
apply(simp add: CAND_eqvt_coname calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12833
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12834
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12835
apply(case_tac "ca=ba")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12836
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12837
apply(drule_tac pi="[(a,ba)]" and x="<ba>:Na" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12838
apply(simp add: CAND_eqvt_coname calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12839
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12840
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12841
apply(drule_tac pi="[(a,ca)]" and x="<ba>:Na" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12842
apply(simp add: CAND_eqvt_coname calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12843
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12844
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12845
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12846
lemma CAND_OrR1_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12847
  assumes a: "<a>:OrR1 <b>.M a \<in> (\<parallel><B>\<parallel>)" "<a>:OrR1 <b>.M a \<notin> BINDINGc B (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12848
  shows "\<exists>B1 B2. B = B1 OR B2 \<and> <b>:M \<in> (\<parallel><B1>\<parallel>)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12849
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12850
apply(nominal_induct B rule: ty.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12851
apply(simp_all add: ty.inject AXIOMSc_def ctrm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12852
apply(auto intro: CANDs_alpha simp add: trm.inject calc_atm abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12853
apply(drule_tac pi="[(a,ba)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12854
apply(auto simp add: CAND_eqvt_coname calc_atm intro: CANDs_alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12855
apply(case_tac "a=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12856
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12857
apply(drule_tac pi="[(aa,ba)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12858
apply(auto simp add: CAND_eqvt_coname calc_atm intro: CANDs_alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12859
apply(case_tac "ba=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12860
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12861
apply(drule_tac pi="[(a,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12862
apply(auto simp add: CAND_eqvt_coname calc_atm intro: CANDs_alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12863
apply(drule_tac pi="[(a,ba)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12864
apply(auto simp add: CAND_eqvt_coname calc_atm intro: CANDs_alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12865
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12866
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12867
lemma CAND_OrR2_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12868
  assumes a: "<a>:OrR2 <b>.M a \<in> (\<parallel><B>\<parallel>)" "<a>:OrR2 <b>.M a \<notin> BINDINGc B (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12869
  shows "\<exists>B1 B2. B = B1 OR B2 \<and> <b>:M \<in> (\<parallel><B2>\<parallel>)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12870
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12871
apply(nominal_induct B rule: ty.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12872
apply(simp_all add: ty.inject AXIOMSc_def ctrm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12873
apply(auto intro: CANDs_alpha simp add: trm.inject calc_atm abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12874
apply(drule_tac pi="[(a,ba)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12875
apply(auto simp add: CAND_eqvt_coname calc_atm intro: CANDs_alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12876
apply(case_tac "a=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12877
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12878
apply(drule_tac pi="[(aa,ba)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12879
apply(auto simp add: CAND_eqvt_coname calc_atm intro: CANDs_alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12880
apply(case_tac "ba=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12881
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12882
apply(drule_tac pi="[(a,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12883
apply(auto simp add: CAND_eqvt_coname calc_atm intro: CANDs_alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12884
apply(drule_tac pi="[(a,ba)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12885
apply(auto simp add: CAND_eqvt_coname calc_atm intro: CANDs_alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12886
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12887
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12888
lemma CAND_OrL_elim_aux:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12889
  assumes a: "(x):(OrL (y).M (z).N x) \<in> NEGn B (\<parallel><B>\<parallel>)" "(x):(OrL (y).M (z).N x) \<notin> BINDINGn B (\<parallel><B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12890
  shows "\<exists>B1 B2. B = B1 OR B2 \<and> (y):M \<in> (\<parallel>(B1)\<parallel>) \<and> (z):N \<in> (\<parallel>(B2)\<parallel>)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12891
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12892
apply(nominal_induct B rule: ty.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12893
apply(simp_all add: ty.inject AXIOMSn_def ntrm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12894
apply(auto intro: CANDs_alpha simp add: trm.inject calc_atm abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12895
apply(drule_tac pi="[(x,za)]" and x="(x):Ma" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12896
apply(simp add: CAND_eqvt_name calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12897
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12898
apply(drule_tac pi="[(x,za)]" and x="(x):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12899
apply(simp add: CAND_eqvt_name calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12900
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12901
apply(drule_tac pi="[(x,za)]" and x="(x):Ma" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12902
apply(simp add: CAND_eqvt_name calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12903
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12904
apply(case_tac "x=ya")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12905
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12906
apply(drule_tac pi="[(ya,za)]" and x="(ya):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12907
apply(simp add: CAND_eqvt_name calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12908
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12909
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12910
apply(case_tac "za=ya")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12911
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12912
apply(drule_tac pi="[(x,ya)]" and x="(ya):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12913
apply(simp add: CAND_eqvt_name calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12914
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12915
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12916
apply(drule_tac pi="[(x,za)]" and x="(ya):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12917
apply(simp add: CAND_eqvt_name calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12918
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12919
apply(case_tac "x=xa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12920
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12921
apply(drule_tac pi="[(xa,za)]" and x="(xa):Ma" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12922
apply(simp add: CAND_eqvt_name calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12923
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12924
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12925
apply(case_tac "za=xa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12926
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12927
apply(drule_tac pi="[(x,xa)]" and x="(xa):Ma" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12928
apply(simp add: CAND_eqvt_name calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12929
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12930
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12931
apply(drule_tac pi="[(x,za)]" and x="(xa):Ma" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12932
apply(simp add: CAND_eqvt_name calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12933
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12934
apply(drule_tac pi="[(x,za)]" and x="(x):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12935
apply(simp add: CAND_eqvt_name calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12936
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12937
apply(case_tac "x=xa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12938
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12939
apply(drule_tac pi="[(xa,za)]" and x="(xa):Ma" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12940
apply(simp add: CAND_eqvt_name calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12941
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12942
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12943
apply(case_tac "za=xa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12944
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12945
apply(drule_tac pi="[(x,xa)]" and x="(xa):Ma" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12946
apply(simp add: CAND_eqvt_name calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12947
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12948
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12949
apply(drule_tac pi="[(x,za)]" and x="(xa):Ma" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12950
apply(simp add: CAND_eqvt_name calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12951
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12952
apply(case_tac "x=ya")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12953
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12954
apply(drule_tac pi="[(ya,za)]" and x="(ya):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12955
apply(simp add: CAND_eqvt_name calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12956
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12957
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12958
apply(case_tac "za=ya")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12959
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12960
apply(drule_tac pi="[(x,ya)]" and x="(ya):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12961
apply(simp add: CAND_eqvt_name calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12962
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12963
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12964
apply(drule_tac pi="[(x,za)]" and x="(ya):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12965
apply(simp add: CAND_eqvt_name calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12966
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12967
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12968
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12969
lemmas CAND_OrL_elim = CAND_OrL_elim_aux[OF NEG_elim(2)]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12970
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12971
lemma CAND_AndL1_elim_aux:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12972
  assumes a: "(x):(AndL1 (y).M x) \<in> NEGn B (\<parallel><B>\<parallel>)" "(x):(AndL1 (y).M x) \<notin> BINDINGn B (\<parallel><B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12973
  shows "\<exists>B1 B2. B = B1 AND B2 \<and> (y):M \<in> (\<parallel>(B1)\<parallel>)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12974
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12975
apply(nominal_induct B rule: ty.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12976
apply(simp_all add: ty.inject AXIOMSn_def ntrm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12977
apply(auto intro: CANDs_alpha simp add: trm.inject calc_atm abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12978
apply(drule_tac pi="[(x,ya)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12979
apply(auto simp add: CAND_eqvt_name calc_atm intro: CANDs_alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12980
apply(case_tac "x=xa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12981
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12982
apply(drule_tac pi="[(xa,ya)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12983
apply(auto simp add: CAND_eqvt_name calc_atm intro: CANDs_alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12984
apply(case_tac "ya=xa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12985
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12986
apply(drule_tac pi="[(x,xa)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12987
apply(auto simp add: CAND_eqvt_name calc_atm intro: CANDs_alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12988
apply(drule_tac pi="[(x,ya)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12989
apply(auto simp add: CAND_eqvt_name calc_atm intro: CANDs_alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12990
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12991
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12992
lemmas CAND_AndL1_elim = CAND_AndL1_elim_aux[OF NEG_elim(2)]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12993
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12994
lemma CAND_AndL2_elim_aux:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12995
  assumes a: "(x):(AndL2 (y).M x) \<in> NEGn B (\<parallel><B>\<parallel>)" "(x):(AndL2 (y).M x) \<notin> BINDINGn B (\<parallel><B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12996
  shows "\<exists>B1 B2. B = B1 AND B2 \<and> (y):M \<in> (\<parallel>(B2)\<parallel>)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12997
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12998
apply(nominal_induct B rule: ty.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 12999
apply(simp_all add: ty.inject AXIOMSn_def ntrm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13000
apply(auto intro: CANDs_alpha simp add: trm.inject calc_atm abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13001
apply(drule_tac pi="[(x,ya)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13002
apply(auto simp add: CAND_eqvt_name calc_atm intro: CANDs_alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13003
apply(case_tac "x=xa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13004
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13005
apply(drule_tac pi="[(xa,ya)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13006
apply(auto simp add: CAND_eqvt_name calc_atm intro: CANDs_alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13007
apply(case_tac "ya=xa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13008
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13009
apply(drule_tac pi="[(x,xa)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13010
apply(auto simp add: CAND_eqvt_name calc_atm intro: CANDs_alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13011
apply(drule_tac pi="[(x,ya)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13012
apply(auto simp add: CAND_eqvt_name calc_atm intro: CANDs_alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13013
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13014
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13015
lemmas CAND_AndL2_elim = CAND_AndL2_elim_aux[OF NEG_elim(2)]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13016
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13017
lemma CAND_ImpL_elim_aux:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13018
  assumes a: "(x):(ImpL <a>.M (z).N x) \<in> NEGn B (\<parallel><B>\<parallel>)" "(x):(ImpL <a>.M (z).N x) \<notin> BINDINGn B (\<parallel><B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13019
  shows "\<exists>B1 B2. B = B1 IMP B2 \<and> <a>:M \<in> (\<parallel><B1>\<parallel>) \<and> (z):N \<in> (\<parallel>(B2)\<parallel>)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13020
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13021
apply(nominal_induct B rule: ty.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13022
apply(simp_all add: ty.inject AXIOMSn_def ntrm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13023
apply(auto intro: CANDs_alpha simp add: trm.inject calc_atm abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13024
apply(drule_tac pi="[(x,y)]" and x="<aa>:Ma" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13025
apply(simp add: CAND_eqvt_name calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13026
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13027
apply(drule_tac pi="[(x,y)]" and x="(x):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13028
apply(simp add: CAND_eqvt_name calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13029
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13030
apply(drule_tac pi="[(x,y)]" and x="<aa>:Ma" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13031
apply(simp add: CAND_eqvt_name calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13032
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13033
apply(case_tac "x=xa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13034
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13035
apply(drule_tac pi="[(xa,y)]" and x="(xa):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13036
apply(simp add: CAND_eqvt_name calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13037
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13038
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13039
apply(case_tac "y=xa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13040
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13041
apply(drule_tac pi="[(x,xa)]" and x="(xa):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13042
apply(simp add: CAND_eqvt_name calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13043
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13044
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13045
apply(drule_tac pi="[(x,y)]" and x="(xa):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13046
apply(simp add: CAND_eqvt_name calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13047
apply(auto intro: CANDs_alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13048
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13049
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13050
lemmas CAND_ImpL_elim = CAND_ImpL_elim_aux[OF NEG_elim(2)]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13051
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13052
lemma CAND_ImpR_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13053
  assumes a: "<a>:ImpR (x).<b>.M a \<in> (\<parallel><B>\<parallel>)" "<a>:ImpR (x).<b>.M a \<notin> BINDINGc B (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13054
  shows "\<exists>B1 B2. B = B1 IMP B2 \<and> 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13055
                 (\<forall>z P. x\<sharp>(z,P) \<and> (z):P \<in> \<parallel>(B2)\<parallel> \<longrightarrow> (x):(M{b:=(z).P}) \<in> \<parallel>(B1)\<parallel>) \<and>
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13056
                 (\<forall>c Q. b\<sharp>(c,Q) \<and> <c>:Q \<in> \<parallel><B1>\<parallel> \<longrightarrow> <b>:(M{x:=<c>.Q}) \<in> \<parallel><B2>\<parallel>)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13057
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13058
apply(nominal_induct B rule: ty.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13059
apply(simp_all add: ty.inject AXIOMSc_def ctrm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13060
apply(auto intro: CANDs_alpha simp add: trm.inject calc_atm abs_fresh fresh_atm fresh_prod fresh_bij)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13061
apply(generate_fresh "name") 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13062
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13063
apply(drule_tac a="ca" and z="c" in alpha_name_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13064
apply(simp) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13065
apply(simp) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13066
apply(simp) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13067
apply(drule_tac x="[(xa,c)]\<bullet>[(aa,ca)]\<bullet>[(b,ca)]\<bullet>[(x,c)]\<bullet>z" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13068
apply(drule_tac x="[(xa,c)]\<bullet>[(aa,ca)]\<bullet>[(b,ca)]\<bullet>[(x,c)]\<bullet>P" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13069
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13070
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13071
apply(auto simp add: calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13072
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13073
apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13074
apply(drule_tac pi="[(x,c)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13075
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13076
apply(drule_tac pi="[(b,ca)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13077
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13078
apply(drule_tac pi="[(aa,ca)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13079
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13080
apply(drule_tac pi="[(xa,c)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13081
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13082
apply(drule_tac pi="[(xa,c)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13083
apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13084
apply(drule_tac pi="[(aa,ca)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13085
apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13086
apply(drule_tac pi="[(b,ca)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13087
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13088
apply(drule_tac pi="[(x,c)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13089
apply(perm_simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13090
apply(auto simp add: calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13091
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13092
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13093
apply(drule_tac a="cb" and z="ca" in alpha_name_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13094
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13095
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13096
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13097
apply(drule_tac x="[(xa,ca)]\<bullet>[(aa,cb)]\<bullet>[(b,cb)]\<bullet>[(x,ca)]\<bullet>c" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13098
apply(drule_tac x="[(xa,ca)]\<bullet>[(aa,cb)]\<bullet>[(b,cb)]\<bullet>[(x,ca)]\<bullet>Q" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13099
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13100
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13101
apply(auto simp add: calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13102
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13103
apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13104
apply(drule_tac pi="[(x,ca)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13105
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13106
apply(drule_tac pi="[(b,cb)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13107
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13108
apply(drule_tac pi="[(aa,cb)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13109
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13110
apply(drule_tac pi="[(xa,ca)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13111
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13112
apply(drule_tac pi="[(xa,ca)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13113
apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13114
apply(drule_tac pi="[(aa,cb)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13115
apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13116
apply(drule_tac pi="[(b,cb)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13117
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13118
apply(drule_tac pi="[(x,ca)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13119
apply(perm_simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13120
apply(auto simp add: calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13121
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13122
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13123
apply(drule_tac a="ca" and z="c" in alpha_name_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13124
apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13125
apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13126
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13127
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13128
apply(drule_tac x="[(a,ba)]\<bullet>[(xa,c)]\<bullet>[(ba,ca)]\<bullet>[(b,ca)]\<bullet>[(x,c)]\<bullet>z" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13129
apply(drule_tac x="[(a,ba)]\<bullet>[(xa,c)]\<bullet>[(ba,ca)]\<bullet>[(b,ca)]\<bullet>[(x,c)]\<bullet>P" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13130
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13131
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13132
apply(auto simp add: calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13133
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13134
apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13135
apply(drule_tac pi="[(x,c)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13136
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13137
apply(drule_tac pi="[(b,ca)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13138
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13139
apply(drule_tac pi="[(ba,ca)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13140
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13141
apply(drule_tac pi="[(xa,c)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13142
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13143
apply(drule_tac pi="[(a,ba)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13144
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13145
apply(drule_tac pi="[(a,ba)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13146
apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13147
apply(drule_tac pi="[(xa,c)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13148
apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13149
apply(drule_tac pi="[(ba,ca)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13150
apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13151
apply(drule_tac pi="[(b,ca)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13152
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13153
apply(drule_tac pi="[(x,c)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13154
apply(perm_simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13155
apply(auto simp add: calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13156
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13157
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13158
apply(drule_tac a="cb" and z="ca" in alpha_name_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13159
apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13160
apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13161
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13162
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13163
apply(drule_tac x="[(a,ba)]\<bullet>[(xa,ca)]\<bullet>[(ba,cb)]\<bullet>[(b,cb)]\<bullet>[(x,ca)]\<bullet>c" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13164
apply(drule_tac x="[(a,ba)]\<bullet>[(xa,ca)]\<bullet>[(ba,cb)]\<bullet>[(b,cb)]\<bullet>[(x,ca)]\<bullet>Q" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13165
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13166
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13167
apply(auto simp add: calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13168
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13169
apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13170
apply(drule_tac pi="[(x,ca)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13171
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13172
apply(drule_tac pi="[(b,cb)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13173
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13174
apply(drule_tac pi="[(ba,cb)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13175
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13176
apply(drule_tac pi="[(xa,ca)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13177
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13178
apply(drule_tac pi="[(a,ba)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13179
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13180
apply(drule_tac pi="[(a,ba)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13181
apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13182
apply(drule_tac pi="[(xa,ca)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13183
apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13184
apply(drule_tac pi="[(ba,cb)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13185
apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13186
apply(drule_tac pi="[(b,cb)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13187
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13188
apply(drule_tac pi="[(x,ca)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13189
apply(perm_simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13190
apply(auto simp add: calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13191
apply(case_tac "a=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13192
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13193
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13194
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13195
apply(drule_tac a="ca" and z="c" in alpha_name_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13196
apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13197
apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13198
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13199
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13200
apply(drule_tac x="[(aa,ba)]\<bullet>[(xa,c)]\<bullet>[(ba,ca)]\<bullet>[(b,ca)]\<bullet>[(x,c)]\<bullet>z" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13201
apply(drule_tac x="[(aa,ba)]\<bullet>[(xa,c)]\<bullet>[(ba,ca)]\<bullet>[(b,ca)]\<bullet>[(x,c)]\<bullet>P" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13202
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13203
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13204
apply(auto simp add: calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13205
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13206
apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13207
apply(drule_tac pi="[(x,c)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13208
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13209
apply(drule_tac pi="[(b,ca)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13210
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13211
apply(drule_tac pi="[(ba,ca)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13212
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13213
apply(drule_tac pi="[(xa,c)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13214
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13215
apply(drule_tac pi="[(aa,ba)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13216
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13217
apply(drule_tac pi="[(aa,ba)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13218
apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13219
apply(drule_tac pi="[(xa,c)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13220
apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13221
apply(drule_tac pi="[(ba,ca)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13222
apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13223
apply(drule_tac pi="[(b,ca)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13224
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13225
apply(drule_tac pi="[(x,c)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13226
apply(perm_simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13227
apply(auto simp add: calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13228
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13229
apply(case_tac "ba=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13230
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13231
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13232
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13233
apply(drule_tac a="ca" and z="c" in alpha_name_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13234
apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13235
apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13236
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13237
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13238
apply(drule_tac x="[(a,aa)]\<bullet>[(xa,c)]\<bullet>[(a,ca)]\<bullet>[(b,ca)]\<bullet>[(x,c)]\<bullet>z" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13239
apply(drule_tac x="[(a,aa)]\<bullet>[(xa,c)]\<bullet>[(a,ca)]\<bullet>[(b,ca)]\<bullet>[(x,c)]\<bullet>P" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13240
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13241
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13242
apply(auto simp add: calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13243
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13244
apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13245
apply(drule_tac pi="[(x,c)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13246
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13247
apply(drule_tac pi="[(b,ca)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13248
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13249
apply(drule_tac pi="[(a,ca)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13250
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13251
apply(drule_tac pi="[(xa,c)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13252
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13253
apply(drule_tac pi="[(a,aa)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13254
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13255
apply(drule_tac pi="[(a,aa)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13256
apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13257
apply(drule_tac pi="[(xa,c)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13258
apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13259
apply(drule_tac pi="[(a,ca)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13260
apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13261
apply(drule_tac pi="[(b,ca)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13262
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13263
apply(drule_tac pi="[(x,c)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13264
apply(perm_simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13265
apply(auto simp add: calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13266
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13267
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13268
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13269
apply(drule_tac a="ca" and z="c" in alpha_name_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13270
apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13271
apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13272
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13273
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13274
apply(drule_tac x="[(a,ba)]\<bullet>[(xa,c)]\<bullet>[(aa,ca)]\<bullet>[(b,ca)]\<bullet>[(x,c)]\<bullet>z" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13275
apply(drule_tac x="[(a,ba)]\<bullet>[(xa,c)]\<bullet>[(aa,ca)]\<bullet>[(b,ca)]\<bullet>[(x,c)]\<bullet>P" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13276
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13277
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13278
apply(auto simp add: calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13279
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13280
apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13281
apply(drule_tac pi="[(x,c)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13282
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13283
apply(drule_tac pi="[(b,ca)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13284
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13285
apply(drule_tac pi="[(aa,ca)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13286
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13287
apply(drule_tac pi="[(xa,c)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13288
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13289
apply(drule_tac pi="[(a,ba)]" and X="\<parallel>(ty2)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13290
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13291
apply(drule_tac pi="[(a,ba)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13292
apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13293
apply(drule_tac pi="[(xa,c)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13294
apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13295
apply(drule_tac pi="[(aa,ca)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13296
apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13297
apply(drule_tac pi="[(b,ca)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13298
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13299
apply(drule_tac pi="[(x,c)]" and X="\<parallel>(ty1)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13300
apply(perm_simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13301
apply(auto simp add: calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13302
apply(case_tac "a=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13303
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13304
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13305
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13306
apply(drule_tac a="cb" and z="ca" in alpha_name_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13307
apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13308
apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13309
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13310
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13311
apply(drule_tac x="[(aa,ba)]\<bullet>[(xa,ca)]\<bullet>[(ba,cb)]\<bullet>[(b,cb)]\<bullet>[(x,ca)]\<bullet>c" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13312
apply(drule_tac x="[(aa,ba)]\<bullet>[(xa,ca)]\<bullet>[(ba,cb)]\<bullet>[(b,cb)]\<bullet>[(x,ca)]\<bullet>Q" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13313
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13314
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13315
apply(auto simp add: calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13316
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13317
apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13318
apply(drule_tac pi="[(x,ca)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13319
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13320
apply(drule_tac pi="[(b,cb)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13321
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13322
apply(drule_tac pi="[(ba,cb)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13323
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13324
apply(drule_tac pi="[(xa,ca)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13325
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13326
apply(drule_tac pi="[(aa,ba)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13327
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13328
apply(drule_tac pi="[(aa,ba)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13329
apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13330
apply(drule_tac pi="[(xa,ca)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13331
apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13332
apply(drule_tac pi="[(ba,cb)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13333
apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13334
apply(drule_tac pi="[(b,cb)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13335
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13336
apply(drule_tac pi="[(x,ca)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13337
apply(perm_simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13338
apply(auto simp add: calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13339
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13340
apply(case_tac "ba=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13341
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13342
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13343
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13344
apply(drule_tac a="cb" and z="ca" in alpha_name_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13345
apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13346
apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13347
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13348
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13349
apply(drule_tac x="[(a,aa)]\<bullet>[(xa,ca)]\<bullet>[(a,cb)]\<bullet>[(b,cb)]\<bullet>[(x,ca)]\<bullet>c" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13350
apply(drule_tac x="[(a,aa)]\<bullet>[(xa,ca)]\<bullet>[(a,cb)]\<bullet>[(b,cb)]\<bullet>[(x,ca)]\<bullet>Q" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13351
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13352
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13353
apply(auto simp add: calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13354
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13355
apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13356
apply(drule_tac pi="[(x,ca)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13357
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13358
apply(drule_tac pi="[(b,cb)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13359
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13360
apply(drule_tac pi="[(a,cb)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13361
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13362
apply(drule_tac pi="[(xa,ca)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13363
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13364
apply(drule_tac pi="[(a,aa)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13365
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13366
apply(drule_tac pi="[(a,aa)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13367
apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13368
apply(drule_tac pi="[(xa,ca)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13369
apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13370
apply(drule_tac pi="[(a,cb)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13371
apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13372
apply(drule_tac pi="[(b,cb)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13373
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13374
apply(drule_tac pi="[(x,ca)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13375
apply(perm_simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13376
apply(auto simp add: calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13377
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13378
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13379
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13380
apply(drule_tac a="cb" and z="ca" in alpha_name_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13381
apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13382
apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13383
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13384
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13385
apply(drule_tac x="[(a,ba)]\<bullet>[(xa,ca)]\<bullet>[(aa,cb)]\<bullet>[(b,cb)]\<bullet>[(x,ca)]\<bullet>c" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13386
apply(drule_tac x="[(a,ba)]\<bullet>[(xa,ca)]\<bullet>[(aa,cb)]\<bullet>[(b,cb)]\<bullet>[(x,ca)]\<bullet>Q" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13387
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13388
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13389
apply(auto simp add: calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13390
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13391
apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13392
apply(drule_tac pi="[(x,ca)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13393
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13394
apply(drule_tac pi="[(b,cb)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13395
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13396
apply(drule_tac pi="[(aa,cb)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13397
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13398
apply(drule_tac pi="[(xa,ca)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13399
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13400
apply(drule_tac pi="[(a,ba)]" and X="\<parallel><ty1>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13401
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13402
apply(drule_tac pi="[(a,ba)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13403
apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13404
apply(drule_tac pi="[(xa,ca)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13405
apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13406
apply(drule_tac pi="[(aa,cb)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13407
apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13408
apply(drule_tac pi="[(b,cb)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13409
apply(simp add: CAND_eqvt_name CAND_eqvt_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13410
apply(drule_tac pi="[(x,ca)]" and X="\<parallel><ty2>\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13411
apply(perm_simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13412
apply(auto simp add: calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13413
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13414
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13415
text {* Main lemma 1 *}
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13416
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13417
lemma AXIOMS_imply_SNa:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13418
  shows "<a>:M \<in> AXIOMSc B \<Longrightarrow> SNa M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13419
  and   "(x):M \<in> AXIOMSn B \<Longrightarrow> SNa M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13420
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13421
apply(auto simp add: AXIOMSn_def AXIOMSc_def ntrm.inject ctrm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13422
apply(rule Ax_in_SNa)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13423
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13424
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13425
lemma BINDING_imply_SNa:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13426
  shows "<a>:M \<in> BINDINGc B (\<parallel>(B)\<parallel>) \<Longrightarrow> SNa M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13427
  and   "(x):M \<in> BINDINGn B (\<parallel><B>\<parallel>) \<Longrightarrow> SNa M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13428
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13429
apply(auto simp add: BINDINGn_def BINDINGc_def ntrm.inject ctrm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13430
apply(drule_tac x="x" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13431
apply(drule_tac x="Ax x a" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13432
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13433
apply(rule Ax_in_CANDs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13434
apply(drule a_star_preserves_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13435
apply(rule subst_with_ax2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13436
apply(simp add: crename_id)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13437
apply(drule_tac x="x" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13438
apply(drule_tac x="Ax x aa" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13439
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13440
apply(rule Ax_in_CANDs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13441
apply(drule a_star_preserves_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13442
apply(rule subst_with_ax2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13443
apply(simp add: crename_id SNa_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13444
apply(drule_tac x="a" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13445
apply(drule_tac x="Ax x a" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13446
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13447
apply(rule Ax_in_CANDs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13448
apply(drule a_star_preserves_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13449
apply(rule subst_with_ax1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13450
apply(simp add: nrename_id)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13451
apply(drule_tac x="a" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13452
apply(drule_tac x="Ax xa a" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13453
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13454
apply(rule Ax_in_CANDs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13455
apply(drule a_star_preserves_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13456
apply(rule subst_with_ax1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13457
apply(simp add: nrename_id SNa_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13458
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13459
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13460
lemma CANDs_imply_SNa:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13461
  shows "<a>:M \<in> \<parallel><B>\<parallel> \<Longrightarrow> SNa M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13462
  and   "(x):M \<in> \<parallel>(B)\<parallel> \<Longrightarrow> SNa M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13463
proof(induct B arbitrary: a x M rule: ty.weak_induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13464
  case (PR X)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13465
  { case 1 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13466
    have "<a>:M \<in> \<parallel><PR X>\<parallel>" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13467
    then have "<a>:M \<in> NEGc (PR X) (\<parallel>(PR X)\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13468
    then have "<a>:M \<in> AXIOMSc (PR X) \<union> BINDINGc (PR X) (\<parallel>(PR X)\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13469
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13470
    { assume "<a>:M \<in> AXIOMSc (PR X)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13471
      then have "SNa M" by (simp add: AXIOMS_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13472
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13473
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13474
    { assume "<a>:M \<in> BINDINGc (PR X) (\<parallel>(PR X)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13475
      then have "SNa M" by (simp add: BINDING_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13476
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13477
    ultimately show "SNa M" by blast 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13478
  next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13479
    case 2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13480
    have "(x):M \<in> (\<parallel>(PR X)\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13481
    then have "(x):M \<in> NEGn (PR X) (\<parallel><PR X>\<parallel>)" using NEG_simp by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13482
    then have "(x):M \<in> AXIOMSn (PR X) \<union> BINDINGn (PR X) (\<parallel><PR X>\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13483
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13484
    { assume "(x):M \<in> AXIOMSn (PR X)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13485
      then have "SNa M" by (simp add: AXIOMS_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13486
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13487
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13488
    { assume "(x):M \<in> BINDINGn (PR X) (\<parallel><PR X>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13489
      then have "SNa M" by (simp only: BINDING_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13490
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13491
    ultimately show "SNa M" by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13492
  }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13493
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13494
  case (NOT B)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13495
  have ih1: "\<And>a M. <a>:M \<in> \<parallel><B>\<parallel> \<Longrightarrow> SNa M" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13496
  have ih2: "\<And>x M. (x):M \<in> \<parallel>(B)\<parallel> \<Longrightarrow> SNa M" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13497
  { case 1
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13498
    have "<a>:M \<in> (\<parallel><NOT B>\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13499
    then have "<a>:M \<in> NEGc (NOT B) (\<parallel>(NOT B)\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13500
    then have "<a>:M \<in> AXIOMSc (NOT B) \<union> BINDINGc (NOT B) (\<parallel>(NOT B)\<parallel>) \<union> NOTRIGHT (NOT B) (\<parallel>(B)\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13501
     moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13502
    { assume "<a>:M \<in> AXIOMSc (NOT B)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13503
      then have "SNa M" by (simp add: AXIOMS_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13504
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13505
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13506
    { assume "<a>:M \<in> BINDINGc (NOT B) (\<parallel>(NOT B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13507
      then have "SNa M" by (simp only: BINDING_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13508
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13509
     moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13510
    { assume "<a>:M \<in> NOTRIGHT (NOT B) (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13511
      then obtain x' M' where eq: "M = NotR (x').M' a" and "(x'):M' \<in> (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13512
	using NOTRIGHT_elim by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13513
      then have "SNa M'" using ih2 by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13514
      then have "SNa M" using eq by (simp add: NotR_in_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13515
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13516
    ultimately show "SNa M" by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13517
  next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13518
    case 2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13519
    have "(x):M \<in> (\<parallel>(NOT B)\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13520
    then have "(x):M \<in> NEGn (NOT B) (\<parallel><NOT B>\<parallel>)" using NEG_simp by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13521
    then have "(x):M \<in> AXIOMSn (NOT B) \<union> BINDINGn (NOT B) (\<parallel><NOT B>\<parallel>) \<union> NOTLEFT (NOT B) (\<parallel><B>\<parallel>)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13522
      by (simp only: NEGn.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13523
     moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13524
    { assume "(x):M \<in> AXIOMSn (NOT B)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13525
      then have "SNa M" by (simp add: AXIOMS_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13526
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13527
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13528
    { assume "(x):M \<in> BINDINGn (NOT B) (\<parallel><NOT B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13529
      then have "SNa M" by (simp only: BINDING_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13530
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13531
     moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13532
    { assume "(x):M \<in> NOTLEFT (NOT B) (\<parallel><B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13533
      then obtain a' M' where eq: "M = NotL <a'>.M' x" and "<a'>:M' \<in> (\<parallel><B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13534
	using NOTLEFT_elim by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13535
      then have "SNa M'" using ih1 by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13536
      then have "SNa M" using eq by (simp add: NotL_in_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13537
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13538
    ultimately show "SNa M" by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13539
  }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13540
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13541
  case (AND A B)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13542
  have ih1: "\<And>a M. <a>:M \<in> \<parallel><A>\<parallel> \<Longrightarrow> SNa M" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13543
  have ih2: "\<And>x M. (x):M \<in> \<parallel>(A)\<parallel> \<Longrightarrow> SNa M" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13544
  have ih3: "\<And>a M. <a>:M \<in> \<parallel><B>\<parallel> \<Longrightarrow> SNa M" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13545
  have ih4: "\<And>x M. (x):M \<in> \<parallel>(B)\<parallel> \<Longrightarrow> SNa M" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13546
  { case 1
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13547
    have "<a>:M \<in> (\<parallel><A AND B>\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13548
    then have "<a>:M \<in> NEGc (A AND B) (\<parallel>(A AND B)\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13549
    then have "<a>:M \<in> AXIOMSc (A AND B) \<union> BINDINGc (A AND B) (\<parallel>(A AND B)\<parallel>) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13550
                                  \<union> ANDRIGHT (A AND B) (\<parallel><A>\<parallel>) (\<parallel><B>\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13551
     moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13552
    { assume "<a>:M \<in> AXIOMSc (A AND B)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13553
      then have "SNa M" by (simp add: AXIOMS_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13554
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13555
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13556
    { assume "<a>:M \<in> BINDINGc (A AND B) (\<parallel>(A AND B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13557
      then have "SNa M" by (simp only: BINDING_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13558
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13559
     moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13560
    { assume "<a>:M \<in> ANDRIGHT (A AND B) (\<parallel><A>\<parallel>) (\<parallel><B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13561
      then obtain a' M' b' N' where eq: "M = AndR <a'>.M' <b'>.N' a" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13562
                                and "<a'>:M' \<in> (\<parallel><A>\<parallel>)" and "<b'>:N' \<in> (\<parallel><B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13563
	by (erule_tac ANDRIGHT_elim, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13564
      then have "SNa M'" and "SNa N'" using ih1 ih3 by blast+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13565
      then have "SNa M" using eq by (simp add: AndR_in_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13566
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13567
    ultimately show "SNa M" by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13568
  next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13569
    case 2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13570
    have "(x):M \<in> (\<parallel>(A AND B)\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13571
    then have "(x):M \<in> NEGn (A AND B) (\<parallel><A AND B>\<parallel>)" using NEG_simp by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13572
    then have "(x):M \<in> AXIOMSn (A AND B) \<union> BINDINGn (A AND B) (\<parallel><A AND B>\<parallel>) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13573
                       \<union> ANDLEFT1 (A AND B) (\<parallel>(A)\<parallel>) \<union> ANDLEFT2 (A AND B) (\<parallel>(B)\<parallel>)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13574
      by (simp only: NEGn.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13575
     moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13576
    { assume "(x):M \<in> AXIOMSn (A AND B)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13577
      then have "SNa M" by (simp add: AXIOMS_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13578
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13579
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13580
    { assume "(x):M \<in> BINDINGn (A AND B) (\<parallel><A AND B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13581
      then have "SNa M" by (simp only: BINDING_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13582
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13583
     moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13584
    { assume "(x):M \<in> ANDLEFT1 (A AND B) (\<parallel>(A)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13585
      then obtain x' M' where eq: "M = AndL1 (x').M' x" and "(x'):M' \<in> (\<parallel>(A)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13586
	using ANDLEFT1_elim by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13587
      then have "SNa M'" using ih2 by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13588
      then have "SNa M" using eq by (simp add: AndL1_in_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13589
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13590
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13591
    { assume "(x):M \<in> ANDLEFT2 (A AND B) (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13592
      then obtain x' M' where eq: "M = AndL2 (x').M' x" and "(x'):M' \<in> (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13593
	using ANDLEFT2_elim by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13594
      then have "SNa M'" using ih4 by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13595
      then have "SNa M" using eq by (simp add: AndL2_in_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13596
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13597
    ultimately show "SNa M" by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13598
  }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13599
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13600
  case (OR A B)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13601
  have ih1: "\<And>a M. <a>:M \<in> \<parallel><A>\<parallel> \<Longrightarrow> SNa M" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13602
  have ih2: "\<And>x M. (x):M \<in> \<parallel>(A)\<parallel> \<Longrightarrow> SNa M" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13603
  have ih3: "\<And>a M. <a>:M \<in> \<parallel><B>\<parallel> \<Longrightarrow> SNa M" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13604
  have ih4: "\<And>x M. (x):M \<in> \<parallel>(B)\<parallel> \<Longrightarrow> SNa M" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13605
  { case 1
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13606
    have "<a>:M \<in> (\<parallel><A OR B>\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13607
    then have "<a>:M \<in> NEGc (A OR B) (\<parallel>(A OR B)\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13608
    then have "<a>:M \<in> AXIOMSc (A OR B) \<union> BINDINGc (A OR B) (\<parallel>(A OR B)\<parallel>) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13609
                                  \<union> ORRIGHT1 (A OR B) (\<parallel><A>\<parallel>) \<union> ORRIGHT2 (A OR B) (\<parallel><B>\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13610
     moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13611
    { assume "<a>:M \<in> AXIOMSc (A OR B)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13612
      then have "SNa M" by (simp add: AXIOMS_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13613
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13614
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13615
    { assume "<a>:M \<in> BINDINGc (A OR B) (\<parallel>(A OR B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13616
      then have "SNa M" by (simp only: BINDING_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13617
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13618
     moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13619
    { assume "<a>:M \<in> ORRIGHT1 (A OR B) (\<parallel><A>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13620
      then obtain a' M' where eq: "M = OrR1 <a'>.M' a" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13621
                                and "<a'>:M' \<in> (\<parallel><A>\<parallel>)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13622
	by (erule_tac ORRIGHT1_elim, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13623
      then have "SNa M'" using ih1 by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13624
      then have "SNa M" using eq by (simp add: OrR1_in_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13625
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13626
     moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13627
    { assume "<a>:M \<in> ORRIGHT2 (A OR B) (\<parallel><B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13628
      then obtain a' M' where eq: "M = OrR2 <a'>.M' a" and "<a'>:M' \<in> (\<parallel><B>\<parallel>)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13629
	using ORRIGHT2_elim by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13630
      then have "SNa M'" using ih3 by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13631
      then have "SNa M" using eq by (simp add: OrR2_in_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13632
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13633
    ultimately show "SNa M" by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13634
  next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13635
    case 2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13636
    have "(x):M \<in> (\<parallel>(A OR B)\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13637
    then have "(x):M \<in> NEGn (A OR B) (\<parallel><A OR B>\<parallel>)" using NEG_simp by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13638
    then have "(x):M \<in> AXIOMSn (A OR B) \<union> BINDINGn (A OR B) (\<parallel><A OR B>\<parallel>) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13639
                       \<union> ORLEFT (A OR B) (\<parallel>(A)\<parallel>) (\<parallel>(B)\<parallel>)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13640
      by (simp only: NEGn.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13641
     moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13642
    { assume "(x):M \<in> AXIOMSn (A OR B)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13643
      then have "SNa M" by (simp add: AXIOMS_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13644
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13645
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13646
    { assume "(x):M \<in> BINDINGn (A OR B) (\<parallel><A OR B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13647
      then have "SNa M" by (simp only: BINDING_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13648
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13649
     moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13650
    { assume "(x):M \<in> ORLEFT (A OR B) (\<parallel>(A)\<parallel>) (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13651
      then obtain x' M' y' N' where eq: "M = OrL (x').M' (y').N' x" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13652
                                and "(x'):M' \<in> (\<parallel>(A)\<parallel>)" and  "(y'):N' \<in> (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13653
	by (erule_tac ORLEFT_elim, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13654
      then have "SNa M'" and "SNa N'" using ih2 ih4 by blast+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13655
      then have "SNa M" using eq by (simp add: OrL_in_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13656
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13657
    ultimately show "SNa M" by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13658
  }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13659
next 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13660
  case (IMP A B)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13661
  have ih1: "\<And>a M. <a>:M \<in> \<parallel><A>\<parallel> \<Longrightarrow> SNa M" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13662
  have ih2: "\<And>x M. (x):M \<in> \<parallel>(A)\<parallel> \<Longrightarrow> SNa M" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13663
  have ih3: "\<And>a M. <a>:M \<in> \<parallel><B>\<parallel> \<Longrightarrow> SNa M" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13664
  have ih4: "\<And>x M. (x):M \<in> \<parallel>(B)\<parallel> \<Longrightarrow> SNa M" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13665
  { case 1
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13666
    have "<a>:M \<in> (\<parallel><A IMP B>\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13667
    then have "<a>:M \<in> NEGc (A IMP B) (\<parallel>(A IMP B)\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13668
    then have "<a>:M \<in> AXIOMSc (A IMP B) \<union> BINDINGc (A IMP B) (\<parallel>(A IMP B)\<parallel>) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13669
                                  \<union> IMPRIGHT (A IMP B) (\<parallel>(A)\<parallel>) (\<parallel><B>\<parallel>) (\<parallel>(B)\<parallel>) (\<parallel><A>\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13670
     moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13671
    { assume "<a>:M \<in> AXIOMSc (A IMP B)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13672
      then have "SNa M" by (simp add: AXIOMS_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13673
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13674
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13675
    { assume "<a>:M \<in> BINDINGc (A IMP B) (\<parallel>(A IMP B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13676
      then have "SNa M" by (simp only: BINDING_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13677
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13678
     moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13679
    { assume "<a>:M \<in> IMPRIGHT (A IMP B) (\<parallel>(A)\<parallel>) (\<parallel><B>\<parallel>) (\<parallel>(B)\<parallel>) (\<parallel><A>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13680
      then obtain x' a' M' where eq: "M = ImpR (x').<a'>.M' a" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13681
                           and imp: "\<forall>z P. x'\<sharp>(z,P) \<and> (z):P \<in> \<parallel>(B)\<parallel> \<longrightarrow> (x'):(M'{a':=(z).P}) \<in> \<parallel>(A)\<parallel>"    
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13682
	by (erule_tac IMPRIGHT_elim, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13683
      obtain z::"name" where fs: "z\<sharp>x'" by (rule_tac exists_fresh, rule fin_supp, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13684
      have "(z):Ax z a'\<in> \<parallel>(B)\<parallel>" by (simp add: Ax_in_CANDs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13685
      with imp fs have "(x'):(M'{a':=(z).Ax z a'}) \<in> \<parallel>(A)\<parallel>" by (simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13686
      then have "SNa (M'{a':=(z).Ax z a'})" using ih2 by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13687
      moreover 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13688
      have "M'{a':=(z).Ax z a'} \<longrightarrow>\<^isub>a* M'[a'\<turnstile>c>a']" by (simp add: subst_with_ax2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13689
      ultimately have "SNa (M'[a'\<turnstile>c>a'])" by (simp add: a_star_preserves_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13690
      then have "SNa M'" by (simp add: crename_id)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13691
      then have "SNa M" using eq by (simp add: ImpR_in_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13692
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13693
    ultimately show "SNa M" by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13694
  next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13695
    case 2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13696
    have "(x):M \<in> (\<parallel>(A IMP B)\<parallel>)" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13697
    then have "(x):M \<in> NEGn (A IMP B) (\<parallel><A IMP B>\<parallel>)" using NEG_simp by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13698
    then have "(x):M \<in> AXIOMSn (A IMP B) \<union> BINDINGn (A IMP B) (\<parallel><A IMP B>\<parallel>) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13699
                       \<union> IMPLEFT (A IMP B) (\<parallel><A>\<parallel>) (\<parallel>(B)\<parallel>)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13700
      by (simp only: NEGn.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13701
     moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13702
    { assume "(x):M \<in> AXIOMSn (A IMP B)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13703
      then have "SNa M" by (simp add: AXIOMS_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13704
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13705
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13706
    { assume "(x):M \<in> BINDINGn (A IMP B) (\<parallel><A IMP B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13707
      then have "SNa M" by (simp only: BINDING_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13708
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13709
     moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13710
    { assume "(x):M \<in> IMPLEFT (A IMP B) (\<parallel><A>\<parallel>) (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13711
      then obtain a' M' y' N' where eq: "M = ImpL <a'>.M' (y').N' x" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13712
                                and "<a'>:M' \<in> (\<parallel><A>\<parallel>)" and  "(y'):N' \<in> (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13713
	by (erule_tac IMPLEFT_elim, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13714
      then have "SNa M'" and "SNa N'" using ih1 ih4 by blast+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13715
      then have "SNa M" using eq by (simp add: ImpL_in_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13716
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13717
    ultimately show "SNa M" by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13718
  }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13719
qed 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13720
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13721
text {* Main lemma 2 *}
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13722
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13723
lemma AXIOMS_preserved:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13724
  shows "<a>:M \<in> AXIOMSc B \<Longrightarrow> M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> <a>:M' \<in> AXIOMSc B"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13725
  and   "(x):M \<in> AXIOMSn B \<Longrightarrow> M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> (x):M' \<in> AXIOMSn B"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13726
  apply(simp_all add: AXIOMSc_def AXIOMSn_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13727
  apply(auto simp add: ntrm.inject ctrm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13728
  apply(drule ax_do_not_a_star_reduce)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13729
  apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13730
  apply(drule ax_do_not_a_star_reduce)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13731
  apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13732
  apply(drule ax_do_not_a_star_reduce)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13733
  apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13734
  apply(drule ax_do_not_a_star_reduce)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13735
  apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13736
  done	
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13737
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13738
lemma BINDING_preserved:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13739
  shows "<a>:M \<in> BINDINGc B (\<parallel>(B)\<parallel>) \<Longrightarrow> M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> <a>:M' \<in> BINDINGc B (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13740
  and   "(x):M \<in> BINDINGn B (\<parallel><B>\<parallel>) \<Longrightarrow> M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> (x):M' \<in> BINDINGn B (\<parallel><B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13741
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13742
  assume red: "M \<longrightarrow>\<^isub>a* M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13743
  assume asm: "<a>:M \<in> BINDINGc B (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13744
  {
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13745
    fix x::"name" and  P::"trm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13746
    from asm have "((x):P) \<in> (\<parallel>(B)\<parallel>) \<Longrightarrow> SNa (M{a:=(x).P})" by (simp add: BINDINGc_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13747
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13748
    have "M{a:=(x).P} \<longrightarrow>\<^isub>a* M'{a:=(x).P}" using red by (simp add: a_star_subst2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13749
    ultimately 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13750
    have "((x):P) \<in> (\<parallel>(B)\<parallel>) \<Longrightarrow> SNa (M'{a:=(x).P})" by (simp add: a_star_preserves_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13751
  }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13752
  then show "<a>:M' \<in> BINDINGc B (\<parallel>(B)\<parallel>)" by (auto simp add: BINDINGc_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13753
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13754
  assume red: "M \<longrightarrow>\<^isub>a* M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13755
  assume asm: "(x):M \<in> BINDINGn B (\<parallel><B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13756
  {
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13757
    fix c::"coname" and  P::"trm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13758
    from asm have "(<c>:P) \<in> (\<parallel><B>\<parallel>) \<Longrightarrow> SNa (M{x:=<c>.P})" by (simp add: BINDINGn_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13759
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13760
    have "M{x:=<c>.P} \<longrightarrow>\<^isub>a* M'{x:=<c>.P}" using red by (simp add: a_star_subst1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13761
    ultimately 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13762
    have "(<c>:P) \<in> (\<parallel><B>\<parallel>) \<Longrightarrow> SNa (M'{x:=<c>.P})" by (simp add: a_star_preserves_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13763
  }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13764
  then show "(x):M' \<in> BINDINGn B (\<parallel><B>\<parallel>)" by (auto simp add: BINDINGn_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13765
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13766
    
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13767
lemma CANDs_preserved:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13768
  shows "<a>:M \<in> \<parallel><B>\<parallel> \<Longrightarrow> M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> <a>:M' \<in> \<parallel><B>\<parallel>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13769
  and   "(x):M \<in> \<parallel>(B)\<parallel> \<Longrightarrow> M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> (x):M' \<in> \<parallel>(B)\<parallel>" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13770
proof(nominal_induct B arbitrary: a x M M' rule: ty.induct) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13771
  case (PR X)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13772
  { case 1 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13773
    have asm: "M \<longrightarrow>\<^isub>a* M'" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13774
    have "<a>:M \<in> \<parallel><PR X>\<parallel>" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13775
    then have "<a>:M \<in> NEGc (PR X) (\<parallel>(PR X)\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13776
    then have "<a>:M \<in> AXIOMSc (PR X) \<union> BINDINGc (PR X) (\<parallel>(PR X)\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13777
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13778
    { assume "<a>:M \<in> AXIOMSc (PR X)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13779
      then have "<a>:M' \<in> AXIOMSc (PR X)" using asm by (simp only: AXIOMS_preserved)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13780
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13781
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13782
    { assume "<a>:M \<in> BINDINGc (PR X) (\<parallel>(PR X)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13783
      then have "<a>:M' \<in> BINDINGc (PR X) (\<parallel>(PR X)\<parallel>)" using asm by (simp add: BINDING_preserved)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13784
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13785
    ultimately have "<a>:M' \<in> AXIOMSc (PR X) \<union> BINDINGc (PR X) (\<parallel>(PR X)\<parallel>)" by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13786
    then have "<a>:M' \<in> NEGc (PR X) (\<parallel>(PR X)\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13787
    then show "<a>:M' \<in> (\<parallel><PR X>\<parallel>)" using NEG_simp by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13788
  next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13789
    case 2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13790
    have asm: "M \<longrightarrow>\<^isub>a* M'" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13791
    have "(x):M \<in> \<parallel>(PR X)\<parallel>" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13792
    then have "(x):M \<in> NEGn (PR X) (\<parallel><PR X>\<parallel>)" using NEG_simp by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13793
    then have "(x):M \<in> AXIOMSn (PR X) \<union> BINDINGn (PR X) (\<parallel><PR X>\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13794
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13795
    { assume "(x):M \<in> AXIOMSn (PR X)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13796
      then have "(x):M' \<in> AXIOMSn (PR X)" using asm by (simp only: AXIOMS_preserved) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13797
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13798
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13799
    { assume "(x):M \<in> BINDINGn (PR X) (\<parallel><PR X>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13800
      then have "(x):M' \<in> BINDINGn (PR X) (\<parallel><PR X>\<parallel>)" using asm by (simp only: BINDING_preserved)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13801
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13802
    ultimately have "(x):M' \<in> AXIOMSn (PR X) \<union> BINDINGn (PR X) (\<parallel><PR X>\<parallel>)" by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13803
    then have "(x):M' \<in> NEGn (PR X) (\<parallel><PR X>\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13804
    then show "(x):M' \<in> (\<parallel>(PR X)\<parallel>)" using NEG_simp by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13805
  }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13806
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13807
  case (IMP A B)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13808
  have ih1: "\<And>a M M'. \<lbrakk><a>:M \<in> \<parallel><A>\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> <a>:M' \<in> \<parallel><A>\<parallel>" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13809
  have ih2: "\<And>x M M'. \<lbrakk>(x):M \<in> \<parallel>(A)\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> (x):M' \<in> \<parallel>(A)\<parallel>" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13810
  have ih3: "\<And>a M M'. \<lbrakk><a>:M \<in> \<parallel><B>\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> <a>:M' \<in> \<parallel><B>\<parallel>" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13811
  have ih4: "\<And>x M M'. \<lbrakk>(x):M \<in> \<parallel>(B)\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> (x):M' \<in> \<parallel>(B)\<parallel>" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13812
  { case 1 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13813
    have asm: "M \<longrightarrow>\<^isub>a* M'" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13814
    have "<a>:M \<in> \<parallel><A IMP B>\<parallel>" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13815
    then have "<a>:M \<in> NEGc (A IMP B) (\<parallel>(A IMP B)\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13816
    then have "<a>:M \<in> AXIOMSc (A IMP B) \<union> BINDINGc (A IMP B) (\<parallel>(A IMP B)\<parallel>) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13817
                            \<union> IMPRIGHT (A IMP B) (\<parallel>(A)\<parallel>) (\<parallel><B>\<parallel>) (\<parallel>(B)\<parallel>) (\<parallel><A>\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13818
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13819
    { assume "<a>:M \<in> AXIOMSc (A IMP B)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13820
      then have "<a>:M' \<in> AXIOMSc (A IMP B)" using asm by (simp only: AXIOMS_preserved)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13821
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13822
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13823
    { assume "<a>:M \<in> BINDINGc (A IMP B) (\<parallel>(A IMP B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13824
      then have "<a>:M' \<in> BINDINGc (A IMP B) (\<parallel>(A IMP B)\<parallel>)" using asm by (simp only: BINDING_preserved)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13825
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13826
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13827
    { assume "<a>:M \<in> IMPRIGHT (A IMP B) (\<parallel>(A)\<parallel>) (\<parallel><B>\<parallel>) (\<parallel>(B)\<parallel>) (\<parallel><A>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13828
      then obtain x' a' N' where eq: "M = ImpR (x').<a'>.N' a" and fic: "fic (ImpR (x').<a'>.N' a) a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13829
                           and imp1: "\<forall>z P. x'\<sharp>(z,P) \<and> (z):P \<in> \<parallel>(B)\<parallel> \<longrightarrow> (x'):(N'{a':=(z).P}) \<in> \<parallel>(A)\<parallel>" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13830
                           and imp2: "\<forall>c Q. a'\<sharp>(c,Q) \<and> <c>:Q \<in> \<parallel><A>\<parallel> \<longrightarrow> <a'>:(N'{x':=<c>.Q}) \<in> \<parallel><B>\<parallel>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13831
	using IMPRIGHT_elim by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13832
      from eq asm obtain N'' where eq': "M' = ImpR (x').<a'>.N'' a" and red: "N' \<longrightarrow>\<^isub>a* N''" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13833
	using a_star_redu_ImpR_elim by (blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13834
      from imp1 have "\<forall>z P. x'\<sharp>(z,P) \<and> (z):P \<in> \<parallel>(B)\<parallel> \<longrightarrow> (x'):(N''{a':=(z).P}) \<in> \<parallel>(A)\<parallel>" using red ih2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13835
	apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13836
	apply(drule_tac x="z" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13837
	apply(drule_tac x="P" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13838
	apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13839
	apply(drule_tac a_star_subst2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13840
	apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13841
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13842
      moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13843
      from imp2 have "\<forall>c Q. a'\<sharp>(c,Q) \<and> <c>:Q \<in> \<parallel><A>\<parallel> \<longrightarrow> <a'>:(N''{x':=<c>.Q}) \<in> \<parallel><B>\<parallel>" using red ih3
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13844
	apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13845
	apply(drule_tac x="c" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13846
	apply(drule_tac x="Q" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13847
	apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13848
	apply(drule_tac a_star_subst1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13849
	apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13850
	done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13851
      moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13852
      from fic have "fic M' a" using eq asm by (simp add: fic_a_star_reduce)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13853
      ultimately have "<a>:M' \<in> IMPRIGHT (A IMP B) (\<parallel>(A)\<parallel>) (\<parallel><B>\<parallel>) (\<parallel>(B)\<parallel>) (\<parallel><A>\<parallel>)" using eq' by auto
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13854
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13855
    ultimately have "<a>:M' \<in> AXIOMSc (A IMP B) \<union> BINDINGc (A IMP B) (\<parallel>(A IMP B)\<parallel>)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13856
                                            \<union> IMPRIGHT (A IMP B) (\<parallel>(A)\<parallel>) (\<parallel><B>\<parallel>) (\<parallel>(B)\<parallel>) (\<parallel><A>\<parallel>)" by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13857
    then have "<a>:M' \<in> NEGc (A IMP B) (\<parallel>(A IMP B)\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13858
    then show "<a>:M' \<in> (\<parallel><A IMP B>\<parallel>)" using NEG_simp by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13859
  next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13860
    case 2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13861
    have asm: "M \<longrightarrow>\<^isub>a* M'" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13862
    have "(x):M \<in> \<parallel>(A IMP B)\<parallel>" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13863
    then have "(x):M \<in> NEGn (A IMP B) (\<parallel><A IMP B>\<parallel>)" using NEG_simp by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13864
    then have "(x):M \<in> AXIOMSn (A IMP B) \<union> BINDINGn (A IMP B) (\<parallel><A IMP B>\<parallel>) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13865
                                              \<union> IMPLEFT (A IMP B) (\<parallel><A>\<parallel>) (\<parallel>(B)\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13866
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13867
    { assume "(x):M \<in> AXIOMSn (A IMP B)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13868
      then have "(x):M' \<in> AXIOMSn (A IMP B)" using asm by (simp only: AXIOMS_preserved)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13869
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13870
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13871
    { assume "(x):M \<in> BINDINGn (A IMP B) (\<parallel><A IMP B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13872
      then have "(x):M' \<in> BINDINGn (A IMP B) (\<parallel><A IMP B>\<parallel>)" using asm by (simp only: BINDING_preserved)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13873
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13874
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13875
    { assume "(x):M \<in> IMPLEFT (A IMP B) (\<parallel><A>\<parallel>) (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13876
      then obtain a' T' y' N' where eq: "M = ImpL <a'>.T' (y').N' x" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13877
                             and fin: "fin (ImpL <a'>.T' (y').N' x) x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13878
                             and imp1: "<a'>:T' \<in> \<parallel><A>\<parallel>" and imp2: "(y'):N' \<in> \<parallel>(B)\<parallel>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13879
	by (erule_tac IMPLEFT_elim, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13880
      from eq asm obtain T'' N'' where eq': "M' = ImpL <a'>.T'' (y').N'' x" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13881
                                 and red1: "T' \<longrightarrow>\<^isub>a* T''"  and red2: "N' \<longrightarrow>\<^isub>a* N''"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13882
	using a_star_redu_ImpL_elim by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13883
      from fin have "fin M' x" using eq asm by (simp add: fin_a_star_reduce)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13884
      moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13885
      from imp1 red1 have "<a'>:T'' \<in> \<parallel><A>\<parallel>" using ih1 by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13886
      moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13887
      from imp2 red2 have "(y'):N'' \<in> \<parallel>(B)\<parallel>" using ih4 by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13888
      ultimately have "(x):M' \<in> IMPLEFT (A IMP B) (\<parallel><A>\<parallel>) (\<parallel>(B)\<parallel>)" using eq' by (simp, blast) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13889
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13890
    ultimately have "(x):M' \<in> AXIOMSn (A IMP B) \<union> BINDINGn (A IMP B) (\<parallel><A IMP B>\<parallel>)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13891
                                              \<union> IMPLEFT (A IMP B) (\<parallel><A>\<parallel>) (\<parallel>(B)\<parallel>)" by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13892
    then have "(x):M' \<in> NEGn (A IMP B) (\<parallel><A IMP B>\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13893
    then show "(x):M' \<in> (\<parallel>(A IMP B)\<parallel>)" using NEG_simp by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13894
  }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13895
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13896
  case (AND A B)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13897
  have ih1: "\<And>a M M'. \<lbrakk><a>:M \<in> \<parallel><A>\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> <a>:M' \<in> \<parallel><A>\<parallel>" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13898
  have ih2: "\<And>x M M'. \<lbrakk>(x):M \<in> \<parallel>(A)\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> (x):M' \<in> \<parallel>(A)\<parallel>" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13899
  have ih3: "\<And>a M M'. \<lbrakk><a>:M \<in> \<parallel><B>\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> <a>:M' \<in> \<parallel><B>\<parallel>" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13900
  have ih4: "\<And>x M M'. \<lbrakk>(x):M \<in> \<parallel>(B)\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> (x):M' \<in> \<parallel>(B)\<parallel>" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13901
  { case 1 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13902
    have asm: "M \<longrightarrow>\<^isub>a* M'" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13903
    have "<a>:M \<in> \<parallel><A AND B>\<parallel>" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13904
    then have "<a>:M \<in> NEGc (A AND B) (\<parallel>(A AND B)\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13905
    then have "<a>:M \<in> AXIOMSc (A AND B) \<union> BINDINGc (A AND B) (\<parallel>(A AND B)\<parallel>) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13906
                                              \<union> ANDRIGHT (A AND B) (\<parallel><A>\<parallel>) (\<parallel><B>\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13907
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13908
    { assume "<a>:M \<in> AXIOMSc (A AND B)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13909
      then have "<a>:M' \<in> AXIOMSc (A AND B)" using asm by (simp only: AXIOMS_preserved)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13910
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13911
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13912
    { assume "<a>:M \<in> BINDINGc (A AND B) (\<parallel>(A AND B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13913
      then have "<a>:M' \<in> BINDINGc (A AND B) (\<parallel>(A AND B)\<parallel>)" using asm by (simp only: BINDING_preserved)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13914
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13915
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13916
    { assume "<a>:M \<in> ANDRIGHT (A AND B) (\<parallel><A>\<parallel>) (\<parallel><B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13917
      then obtain a' T' b' N' where eq: "M = AndR <a'>.T' <b'>.N' a" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13918
                              and fic: "fic (AndR <a'>.T' <b'>.N' a) a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13919
                           and imp1: "<a'>:T' \<in> \<parallel><A>\<parallel>" and imp2: "<b'>:N' \<in> \<parallel><B>\<parallel>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13920
	using ANDRIGHT_elim by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13921
      from eq asm obtain T'' N'' where eq': "M' = AndR <a'>.T'' <b'>.N'' a" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13922
                          and red1: "T' \<longrightarrow>\<^isub>a* T''" and red2: "N' \<longrightarrow>\<^isub>a* N''" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13923
	using a_star_redu_AndR_elim by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13924
      from fic have "fic M' a" using eq asm by (simp add: fic_a_star_reduce)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13925
      moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13926
      from imp1 red1 have "<a'>:T'' \<in> \<parallel><A>\<parallel>" using ih1 by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13927
      moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13928
      from imp2 red2 have "<b'>:N'' \<in> \<parallel><B>\<parallel>" using ih3 by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13929
      ultimately have "<a>:M' \<in> ANDRIGHT (A AND B) (\<parallel><A>\<parallel>) (\<parallel><B>\<parallel>)" using eq' by (simp, blast) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13930
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13931
    ultimately have "<a>:M' \<in> AXIOMSc (A AND B) \<union> BINDINGc (A AND B) (\<parallel>(A AND B)\<parallel>)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13932
                                              \<union> ANDRIGHT (A AND B) (\<parallel><A>\<parallel>) (\<parallel><B>\<parallel>)" by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13933
    then have "<a>:M' \<in> NEGc (A AND B) (\<parallel>(A AND B)\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13934
    then show "<a>:M' \<in> (\<parallel><A AND B>\<parallel>)" using NEG_simp by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13935
  next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13936
    case 2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13937
    have asm: "M \<longrightarrow>\<^isub>a* M'" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13938
    have "(x):M \<in> \<parallel>(A AND B)\<parallel>" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13939
    then have "(x):M \<in> NEGn (A AND B) (\<parallel><A AND B>\<parallel>)" using NEG_simp by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13940
    then have "(x):M \<in> AXIOMSn (A AND B) \<union> BINDINGn (A AND B) (\<parallel><A AND B>\<parallel>) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13941
                                     \<union> ANDLEFT1 (A AND B) (\<parallel>(A)\<parallel>) \<union> ANDLEFT2 (A AND B) (\<parallel>(B)\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13942
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13943
    { assume "(x):M \<in> AXIOMSn (A AND B)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13944
      then have "(x):M' \<in> AXIOMSn (A AND B)" using asm by (simp only: AXIOMS_preserved)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13945
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13946
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13947
    { assume "(x):M \<in> BINDINGn (A AND B) (\<parallel><A AND B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13948
      then have "(x):M' \<in> BINDINGn (A AND B) (\<parallel><A AND B>\<parallel>)" using asm by (simp only: BINDING_preserved)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13949
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13950
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13951
    { assume "(x):M \<in> ANDLEFT1 (A AND B) (\<parallel>(A)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13952
      then obtain y' N' where eq: "M = AndL1 (y').N' x" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13953
                             and fin: "fin (AndL1 (y').N' x) x" and imp: "(y'):N' \<in> \<parallel>(A)\<parallel>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13954
	by (erule_tac ANDLEFT1_elim, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13955
      from eq asm obtain N'' where eq': "M' = AndL1 (y').N'' x" and red1: "N' \<longrightarrow>\<^isub>a* N''"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13956
	using a_star_redu_AndL1_elim by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13957
      from fin have "fin M' x" using eq asm by (simp add: fin_a_star_reduce)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13958
      moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13959
      from imp red1 have "(y'):N'' \<in> \<parallel>(A)\<parallel>" using ih2 by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13960
      ultimately have "(x):M' \<in> ANDLEFT1 (A AND B) (\<parallel>(A)\<parallel>)" using eq' by (simp, blast) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13961
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13962
     moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13963
    { assume "(x):M \<in> ANDLEFT2 (A AND B) (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13964
      then obtain y' N' where eq: "M = AndL2 (y').N' x" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13965
                             and fin: "fin (AndL2 (y').N' x) x" and imp: "(y'):N' \<in> \<parallel>(B)\<parallel>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13966
	by (erule_tac ANDLEFT2_elim, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13967
      from eq asm obtain N'' where eq': "M' = AndL2 (y').N'' x" and red1: "N' \<longrightarrow>\<^isub>a* N''"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13968
	using a_star_redu_AndL2_elim by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13969
      from fin have "fin M' x" using eq asm by (simp add: fin_a_star_reduce)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13970
      moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13971
      from imp red1 have "(y'):N'' \<in> \<parallel>(B)\<parallel>" using ih4 by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13972
      ultimately have "(x):M' \<in> ANDLEFT2 (A AND B) (\<parallel>(B)\<parallel>)" using eq' by (simp, blast) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13973
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13974
    ultimately have "(x):M' \<in> AXIOMSn (A AND B) \<union> BINDINGn (A AND B) (\<parallel><A AND B>\<parallel>)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13975
                               \<union> ANDLEFT1 (A AND B) (\<parallel>(A)\<parallel>) \<union> ANDLEFT2 (A AND B) (\<parallel>(B)\<parallel>)" by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13976
    then have "(x):M' \<in> NEGn (A AND B) (\<parallel><A AND B>\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13977
    then show "(x):M' \<in> (\<parallel>(A AND B)\<parallel>)" using NEG_simp by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13978
  }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13979
next    
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13980
 case (OR A B)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13981
  have ih1: "\<And>a M M'. \<lbrakk><a>:M \<in> \<parallel><A>\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> <a>:M' \<in> \<parallel><A>\<parallel>" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13982
  have ih2: "\<And>x M M'. \<lbrakk>(x):M \<in> \<parallel>(A)\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> (x):M' \<in> \<parallel>(A)\<parallel>" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13983
  have ih3: "\<And>a M M'. \<lbrakk><a>:M \<in> \<parallel><B>\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> <a>:M' \<in> \<parallel><B>\<parallel>" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13984
  have ih4: "\<And>x M M'. \<lbrakk>(x):M \<in> \<parallel>(B)\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> (x):M' \<in> \<parallel>(B)\<parallel>" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13985
  { case 1 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13986
    have asm: "M \<longrightarrow>\<^isub>a* M'" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13987
    have "<a>:M \<in> \<parallel><A OR B>\<parallel>" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13988
    then have "<a>:M \<in> NEGc (A OR B) (\<parallel>(A OR B)\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13989
    then have "<a>:M \<in> AXIOMSc (A OR B) \<union> BINDINGc (A OR B) (\<parallel>(A OR B)\<parallel>) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13990
                          \<union> ORRIGHT1 (A OR B) (\<parallel><A>\<parallel>) \<union> ORRIGHT2 (A OR B) (\<parallel><B>\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13991
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13992
    { assume "<a>:M \<in> AXIOMSc (A OR B)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13993
      then have "<a>:M' \<in> AXIOMSc (A OR B)" using asm by (simp only: AXIOMS_preserved)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13994
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13995
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13996
    { assume "<a>:M \<in> BINDINGc (A OR B) (\<parallel>(A OR B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13997
      then have "<a>:M' \<in> BINDINGc (A OR B) (\<parallel>(A OR B)\<parallel>)" using asm by (simp only: BINDING_preserved)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13998
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 13999
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14000
    { assume "<a>:M \<in> ORRIGHT1 (A OR B) (\<parallel><A>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14001
      then obtain a' N' where eq: "M = OrR1 <a'>.N' a" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14002
                              and fic: "fic (OrR1 <a'>.N' a) a" and imp1: "<a'>:N' \<in> \<parallel><A>\<parallel>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14003
	using ORRIGHT1_elim by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14004
      from eq asm obtain N'' where eq': "M' = OrR1 <a'>.N'' a" and red1: "N' \<longrightarrow>\<^isub>a* N''" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14005
	using a_star_redu_OrR1_elim by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14006
      from fic have "fic M' a" using eq asm by (simp add: fic_a_star_reduce)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14007
      moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14008
      from imp1 red1 have "<a'>:N'' \<in> \<parallel><A>\<parallel>" using ih1 by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14009
      ultimately have "<a>:M' \<in> ORRIGHT1 (A OR B) (\<parallel><A>\<parallel>)" using eq' by (simp, blast) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14010
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14011
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14012
    { assume "<a>:M \<in> ORRIGHT2 (A OR B) (\<parallel><B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14013
      then obtain a' N' where eq: "M = OrR2 <a'>.N' a" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14014
                              and fic: "fic (OrR2 <a'>.N' a) a" and imp1: "<a'>:N' \<in> \<parallel><B>\<parallel>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14015
	using ORRIGHT2_elim by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14016
      from eq asm obtain N'' where eq': "M' = OrR2 <a'>.N'' a" and red1: "N' \<longrightarrow>\<^isub>a* N''" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14017
	using a_star_redu_OrR2_elim by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14018
      from fic have "fic M' a" using eq asm by (simp add: fic_a_star_reduce)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14019
      moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14020
      from imp1 red1 have "<a'>:N'' \<in> \<parallel><B>\<parallel>" using ih3 by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14021
      ultimately have "<a>:M' \<in> ORRIGHT2 (A OR B) (\<parallel><B>\<parallel>)" using eq' by (simp, blast) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14022
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14023
    ultimately have "<a>:M' \<in> AXIOMSc (A OR B) \<union> BINDINGc (A OR B) (\<parallel>(A OR B)\<parallel>)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14024
                                \<union> ORRIGHT1 (A OR B) (\<parallel><A>\<parallel>) \<union> ORRIGHT2 (A OR B) (\<parallel><B>\<parallel>)" by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14025
    then have "<a>:M' \<in> NEGc (A OR B) (\<parallel>(A OR B)\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14026
    then show "<a>:M' \<in> (\<parallel><A OR B>\<parallel>)" using NEG_simp by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14027
  next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14028
    case 2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14029
    have asm: "M \<longrightarrow>\<^isub>a* M'" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14030
    have "(x):M \<in> \<parallel>(A OR B)\<parallel>" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14031
    then have "(x):M \<in> NEGn (A OR B) (\<parallel><A OR B>\<parallel>)" using NEG_simp by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14032
    then have "(x):M \<in> AXIOMSn (A OR B) \<union> BINDINGn (A OR B) (\<parallel><A OR B>\<parallel>) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14033
                                     \<union> ORLEFT (A OR B) (\<parallel>(A)\<parallel>) (\<parallel>(B)\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14034
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14035
    { assume "(x):M \<in> AXIOMSn (A OR B)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14036
      then have "(x):M' \<in> AXIOMSn (A OR B)" using asm by (simp only: AXIOMS_preserved)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14037
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14038
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14039
    { assume "(x):M \<in> BINDINGn (A OR B) (\<parallel><A OR B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14040
      then have "(x):M' \<in> BINDINGn (A OR B) (\<parallel><A OR B>\<parallel>)" using asm by (simp only: BINDING_preserved)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14041
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14042
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14043
    { assume "(x):M \<in> ORLEFT (A OR B) (\<parallel>(A)\<parallel>) (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14044
      then obtain y' T' z' N' where eq: "M = OrL (y').T' (z').N' x" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14045
                             and fin: "fin (OrL (y').T' (z').N' x) x" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14046
                             and imp1: "(y'):T' \<in> \<parallel>(A)\<parallel>" and imp2: "(z'):N' \<in> \<parallel>(B)\<parallel>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14047
	by (erule_tac ORLEFT_elim, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14048
      from eq asm obtain T'' N'' where eq': "M' = OrL (y').T'' (z').N'' x" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14049
                and red1: "T' \<longrightarrow>\<^isub>a* T''" and red2: "N' \<longrightarrow>\<^isub>a* N''"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14050
	using a_star_redu_OrL_elim by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14051
      from fin have "fin M' x" using eq asm by (simp add: fin_a_star_reduce)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14052
      moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14053
      from imp1 red1 have "(y'):T'' \<in> \<parallel>(A)\<parallel>" using ih2 by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14054
      moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14055
      from imp2 red2 have "(z'):N'' \<in> \<parallel>(B)\<parallel>" using ih4 by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14056
      ultimately have "(x):M' \<in> ORLEFT (A OR B) (\<parallel>(A)\<parallel>) (\<parallel>(B)\<parallel>)" using eq' by (simp, blast) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14057
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14058
    ultimately have "(x):M' \<in> AXIOMSn (A OR B) \<union> BINDINGn (A OR B) (\<parallel><A OR B>\<parallel>)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14059
                               \<union> ORLEFT (A OR B) (\<parallel>(A)\<parallel>) (\<parallel>(B)\<parallel>)" by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14060
    then have "(x):M' \<in> NEGn (A OR B) (\<parallel><A OR B>\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14061
    then show "(x):M' \<in> (\<parallel>(A OR B)\<parallel>)" using NEG_simp by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14062
  }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14063
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14064
  case (NOT A)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14065
  have ih1: "\<And>a M M'. \<lbrakk><a>:M \<in> \<parallel><A>\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> <a>:M' \<in> \<parallel><A>\<parallel>" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14066
  have ih2: "\<And>x M M'. \<lbrakk>(x):M \<in> \<parallel>(A)\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> (x):M' \<in> \<parallel>(A)\<parallel>" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14067
  { case 1 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14068
    have asm: "M \<longrightarrow>\<^isub>a* M'" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14069
    have "<a>:M \<in> \<parallel><NOT A>\<parallel>" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14070
    then have "<a>:M \<in> NEGc (NOT A) (\<parallel>(NOT A)\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14071
    then have "<a>:M \<in> AXIOMSc (NOT A) \<union> BINDINGc (NOT A) (\<parallel>(NOT A)\<parallel>) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14072
                                              \<union> NOTRIGHT (NOT A) (\<parallel>(A)\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14073
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14074
    { assume "<a>:M \<in> AXIOMSc (NOT A)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14075
      then have "<a>:M' \<in> AXIOMSc (NOT A)" using asm by (simp only: AXIOMS_preserved)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14076
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14077
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14078
    { assume "<a>:M \<in> BINDINGc (NOT A) (\<parallel>(NOT A)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14079
      then have "<a>:M' \<in> BINDINGc (NOT A) (\<parallel>(NOT A)\<parallel>)" using asm by (simp only: BINDING_preserved)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14080
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14081
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14082
    { assume "<a>:M \<in> NOTRIGHT (NOT A) (\<parallel>(A)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14083
      then obtain y' N' where eq: "M = NotR (y').N' a" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14084
                              and fic: "fic (NotR (y').N' a) a" and imp: "(y'):N' \<in> \<parallel>(A)\<parallel>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14085
	using NOTRIGHT_elim by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14086
      from eq asm obtain N'' where eq': "M' = NotR (y').N'' a" and red: "N' \<longrightarrow>\<^isub>a* N''" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14087
	using a_star_redu_NotR_elim by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14088
      from fic have "fic M' a" using eq asm by (simp add: fic_a_star_reduce)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14089
      moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14090
      from imp red have "(y'):N'' \<in> \<parallel>(A)\<parallel>" using ih2 by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14091
      ultimately have "<a>:M' \<in> NOTRIGHT (NOT A) (\<parallel>(A)\<parallel>)" using eq' by (simp, blast) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14092
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14093
    ultimately have "<a>:M' \<in> AXIOMSc (NOT A) \<union> BINDINGc (NOT A) (\<parallel>(NOT A)\<parallel>)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14094
                                              \<union> NOTRIGHT (NOT A) (\<parallel>(A)\<parallel>)" by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14095
    then have "<a>:M' \<in> NEGc (NOT A) (\<parallel>(NOT A)\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14096
    then show "<a>:M' \<in> (\<parallel><NOT A>\<parallel>)" using NEG_simp by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14097
  next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14098
    case 2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14099
    have asm: "M \<longrightarrow>\<^isub>a* M'" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14100
    have "(x):M \<in> \<parallel>(NOT A)\<parallel>" by fact
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14101
    then have "(x):M \<in> NEGn (NOT A) (\<parallel><NOT A>\<parallel>)" using NEG_simp by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14102
    then have "(x):M \<in> AXIOMSn (NOT A) \<union> BINDINGn (NOT A) (\<parallel><NOT A>\<parallel>) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14103
                                     \<union> NOTLEFT (NOT A) (\<parallel><A>\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14104
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14105
    { assume "(x):M \<in> AXIOMSn (NOT A)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14106
      then have "(x):M' \<in> AXIOMSn (NOT A)" using asm by (simp only: AXIOMS_preserved)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14107
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14108
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14109
    { assume "(x):M \<in> BINDINGn (NOT A) (\<parallel><NOT A>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14110
      then have "(x):M' \<in> BINDINGn (NOT A) (\<parallel><NOT A>\<parallel>)" using asm by (simp only: BINDING_preserved)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14111
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14112
    moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14113
    { assume "(x):M \<in> NOTLEFT (NOT A) (\<parallel><A>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14114
      then obtain a' N' where eq: "M = NotL <a'>.N' x" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14115
                             and fin: "fin (NotL <a'>.N' x) x" and imp: "<a'>:N' \<in> \<parallel><A>\<parallel>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14116
	by (erule_tac NOTLEFT_elim, blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14117
      from eq asm obtain N'' where eq': "M' = NotL <a'>.N'' x" and red1: "N' \<longrightarrow>\<^isub>a* N''"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14118
	using a_star_redu_NotL_elim by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14119
      from fin have "fin M' x" using eq asm by (simp add: fin_a_star_reduce)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14120
      moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14121
      from imp red1 have "<a'>:N'' \<in> \<parallel><A>\<parallel>" using ih1 by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14122
      ultimately have "(x):M' \<in> NOTLEFT (NOT A) (\<parallel><A>\<parallel>)" using eq' by (simp, blast) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14123
    }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14124
    ultimately have "(x):M' \<in> AXIOMSn (NOT A) \<union> BINDINGn (NOT A) (\<parallel><NOT A>\<parallel>)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14125
                               \<union> NOTLEFT (NOT A) (\<parallel><A>\<parallel>)" by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14126
    then have "(x):M' \<in> NEGn (NOT A) (\<parallel><NOT A>\<parallel>)" by simp
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14127
    then show "(x):M' \<in> (\<parallel>(NOT A)\<parallel>)" using NEG_simp by blast
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14128
  }
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14129
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14130
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14131
lemma CANDs_preserved_single:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14132
  shows "<a>:M \<in> \<parallel><B>\<parallel> \<Longrightarrow> M \<longrightarrow>\<^isub>a M' \<Longrightarrow> <a>:M' \<in> \<parallel><B>\<parallel>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14133
  and   "(x):M \<in> \<parallel>(B)\<parallel> \<Longrightarrow> M \<longrightarrow>\<^isub>a M' \<Longrightarrow> (x):M' \<in> \<parallel>(B)\<parallel>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14134
by (auto simp add: a_starI CANDs_preserved)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14135
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14136
lemma fic_CANDS:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14137
  assumes a: "\<not>fic M a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14138
  and     b: "<a>:M \<in> \<parallel><B>\<parallel>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14139
  shows "<a>:M \<in> AXIOMSc B \<or> <a>:M \<in> BINDINGc B (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14140
using a b
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14141
apply(nominal_induct B rule: ty.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14142
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14143
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14144
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14145
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14146
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14147
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14148
apply(auto simp add: ctrm.inject)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14149
apply(simp add: alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14150
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14151
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14152
apply(auto simp add: calc_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14153
apply(drule_tac pi="[(a,aa)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14154
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14155
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14156
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14157
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14158
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14159
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14160
apply(auto simp add: ctrm.inject)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14161
apply(simp add: alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14162
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14163
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14164
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14165
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14166
apply(drule_tac pi="[(a,c)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14167
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14168
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14169
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14170
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14171
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14172
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14173
apply(auto simp add: ctrm.inject)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14174
apply(simp add: alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14175
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14176
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14177
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14178
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14179
apply(drule_tac pi="[(a,b)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14180
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14181
apply(simp add: alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14182
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14183
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14184
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14185
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14186
apply(drule_tac pi="[(a,b)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14187
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14188
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14189
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14190
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14191
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14192
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14193
apply(auto simp add: ctrm.inject)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14194
apply(simp add: alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14195
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14196
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14197
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14198
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14199
apply(drule_tac pi="[(a,b)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14200
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14201
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14202
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14203
lemma fin_CANDS_aux:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14204
  assumes a: "\<not>fin M x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14205
  and     b: "(x):M \<in> (NEGn B (\<parallel><B>\<parallel>))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14206
  shows "(x):M \<in> AXIOMSn B \<or> (x):M \<in> BINDINGn B (\<parallel><B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14207
using a b
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14208
apply(nominal_induct B rule: ty.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14209
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14210
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14211
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14212
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14213
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14214
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14215
apply(auto simp add: ntrm.inject)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14216
apply(simp add: alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14217
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14218
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14219
apply(auto simp add: calc_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14220
apply(drule_tac pi="[(x,xa)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14221
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14222
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14223
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14224
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14225
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14226
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14227
apply(auto simp add: ntrm.inject)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14228
apply(simp add: alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14229
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14230
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14231
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14232
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14233
apply(drule_tac pi="[(x,y)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14234
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14235
apply(simp add: alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14236
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14237
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14238
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14239
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14240
apply(drule_tac pi="[(x,y)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14241
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14242
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14243
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14244
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14245
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14246
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14247
apply(auto simp add: ntrm.inject)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14248
apply(simp add: alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14249
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14250
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14251
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14252
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14253
apply(drule_tac pi="[(x,z)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14254
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14255
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14256
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14257
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14258
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14259
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14260
apply(auto simp add: ntrm.inject)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14261
apply(simp add: alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14262
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14263
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14264
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14265
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14266
apply(drule_tac pi="[(x,y)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14267
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14268
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14269
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14270
lemma fin_CANDS:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14271
  assumes a: "\<not>fin M x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14272
  and     b: "(x):M \<in> (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14273
  shows "(x):M \<in> AXIOMSn B \<or> (x):M \<in> BINDINGn B (\<parallel><B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14274
apply(rule fin_CANDS_aux)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14275
apply(rule a)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14276
apply(rule NEG_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14277
apply(rule b)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14278
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14279
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14280
lemma BINDING_implies_CAND:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14281
  shows "<c>:M \<in> BINDINGc B (\<parallel>(B)\<parallel>) \<Longrightarrow> <c>:M \<in> (\<parallel><B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14282
  and   "(x):N \<in> BINDINGn B (\<parallel><B>\<parallel>) \<Longrightarrow> (x):N \<in> (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14283
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14284
apply(nominal_induct B rule: ty.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14285
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14286
apply(rule NEG_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14287
apply(nominal_induct B rule: ty.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14288
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14289
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14290
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14291
text {* 3rd Main Lemma *}
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14292
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14293
lemma Cut_a_redu_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14294
  assumes a: "Cut <a>.M (x).N \<longrightarrow>\<^isub>a R"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14295
  shows "(\<exists>M'. R = Cut <a>.M' (x).N \<and> M \<longrightarrow>\<^isub>a M') \<or>
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14296
         (\<exists>N'. R = Cut <a>.M (x).N' \<and> N \<longrightarrow>\<^isub>a N') \<or>
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14297
         (Cut <a>.M (x).N \<longrightarrow>\<^isub>c R) \<or>
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14298
         (Cut <a>.M (x).N \<longrightarrow>\<^isub>l R)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14299
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14300
apply(erule_tac a_redu.cases)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14301
apply(simp_all)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14302
apply(simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14303
apply(rule disjI1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14304
apply(auto simp add: alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14305
apply(rule_tac x="[(a,aa)]\<bullet>M'" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14306
apply(perm_simp add: fresh_left calc_atm a_redu.eqvt fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14307
apply(rule_tac x="[(a,aa)]\<bullet>M'" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14308
apply(perm_simp add: fresh_left calc_atm a_redu.eqvt fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14309
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14310
apply(rule disjI1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14311
apply(auto simp add: alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14312
apply(rule_tac x="[(x,xa)]\<bullet>N'" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14313
apply(perm_simp add: fresh_left calc_atm a_redu.eqvt fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14314
apply(rule_tac x="[(x,xa)]\<bullet>N'" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14315
apply(perm_simp add: fresh_left calc_atm a_redu.eqvt fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14316
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14317
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14318
lemma Cut_c_redu_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14319
  assumes a: "Cut <a>.M (x).N \<longrightarrow>\<^isub>c R"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14320
  shows "(R = M{a:=(x).N} \<and> \<not>fic M a) \<or>
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14321
         (R = N{x:=<a>.M} \<and> \<not>fin N x)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14322
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14323
apply(erule_tac c_redu.cases)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14324
apply(simp_all)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14325
apply(simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14326
apply(rule disjI1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14327
apply(auto simp add: alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14328
apply(simp add: subst_rename fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14329
apply(simp add: subst_rename fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14330
apply(drule_tac pi="[(a,aa)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14331
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14332
apply(simp add: subst_rename fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14333
apply(drule_tac pi="[(a,aa)]" in fic.eqvt(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14334
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14335
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14336
apply(auto simp add: alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14337
apply(simp add: subst_rename fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14338
apply(drule_tac pi="[(x,xa)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14339
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14340
apply(simp add: subst_rename fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14341
apply(simp add: subst_rename fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14342
apply(drule_tac pi="[(x,xa)]" in fin.eqvt(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14343
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14344
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14345
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14346
lemma not_fic_crename_aux:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14347
  assumes a: "fic M c" "c\<sharp>(a,b)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14348
  shows "fic (M[a\<turnstile>c>b]) c" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14349
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14350
apply(nominal_induct M avoiding: c a b rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14351
apply(auto dest!: fic_elims intro!: fic.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14352
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14353
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14354
lemma not_fic_crename:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14355
  assumes a: "\<not>(fic (M[a\<turnstile>c>b]) c)" "c\<sharp>(a,b)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14356
  shows "\<not>(fic M c)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14357
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14358
apply(auto dest:  not_fic_crename_aux)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14359
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14360
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14361
lemma not_fin_crename_aux:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14362
  assumes a: "fin M y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14363
  shows "fin (M[a\<turnstile>c>b]) y" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14364
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14365
apply(nominal_induct M avoiding: a b rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14366
apply(auto dest!: fin_elims intro!: fin.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14367
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14368
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14369
lemma not_fin_crename:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14370
  assumes a: "\<not>(fin (M[a\<turnstile>c>b]) y)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14371
  shows "\<not>(fin M y)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14372
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14373
apply(auto dest:  not_fin_crename_aux)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14374
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14375
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14376
lemma crename_fresh_interesting1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14377
  fixes c::"coname"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14378
  assumes a: "c\<sharp>(M[a\<turnstile>c>b])" "c\<sharp>(a,b)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14379
  shows "c\<sharp>M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14380
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14381
apply(nominal_induct M avoiding: c a b rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14382
apply(auto split: if_splits simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14383
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14384
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14385
lemma crename_fresh_interesting2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14386
  fixes x::"name"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14387
  assumes a: "x\<sharp>(M[a\<turnstile>c>b])" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14388
  shows "x\<sharp>M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14389
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14390
apply(nominal_induct M avoiding: x a b rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14391
apply(auto split: if_splits simp add: abs_fresh abs_supp fin_supp fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14392
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14393
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14394
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14395
lemma fic_crename:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14396
  assumes a: "fic (M[a\<turnstile>c>b]) c" "c\<sharp>(a,b)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14397
  shows "fic M c" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14398
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14399
apply(nominal_induct M avoiding: c a b rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14400
apply(auto dest!: fic_elims intro!: fic.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14401
           split: if_splits)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14402
apply(auto dest: crename_fresh_interesting1 simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14403
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14404
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14405
lemma fin_crename:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14406
  assumes a: "fin (M[a\<turnstile>c>b]) x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14407
  shows "fin M x" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14408
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14409
apply(nominal_induct M avoiding: x a b rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14410
apply(auto dest!: fin_elims intro!: fin.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14411
           split: if_splits)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14412
apply(auto dest: crename_fresh_interesting2 simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14413
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14414
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14415
lemma crename_Cut:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14416
  assumes a: "R[a\<turnstile>c>b] = Cut <c>.M (x).N" "c\<sharp>(a,b,N,R)" "x\<sharp>(M,R)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14417
  shows "\<exists>M' N'. R = Cut <c>.M' (x).N' \<and> M'[a\<turnstile>c>b] = M \<and> N'[a\<turnstile>c>b] = N \<and> c\<sharp>N' \<and> x\<sharp>M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14418
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14419
apply(nominal_induct R avoiding: a b c x M N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14420
apply(auto split: if_splits)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14421
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14422
apply(auto simp add: alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14423
apply(rule_tac x="[(name,x)]\<bullet>trm2" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14424
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14425
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14426
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14427
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14428
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14429
apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14430
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14431
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14432
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14433
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14434
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14435
apply(auto simp add: fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14436
apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14437
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14438
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14439
apply(rule_tac x="[(name,x)]\<bullet>trm2" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14440
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14441
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14442
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14443
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14444
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14445
apply(auto simp add: fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14446
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14447
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14448
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14449
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14450
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14451
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14452
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14453
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14454
lemma crename_NotR:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14455
  assumes a: "R[a\<turnstile>c>b] = NotR (x).N c" "x\<sharp>R" "c\<sharp>(a,b)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14456
  shows "\<exists>N'. (R = NotR (x).N' c) \<and> N'[a\<turnstile>c>b] = N" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14457
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14458
apply(nominal_induct R avoiding: a b c x N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14459
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14460
apply(rule_tac x="[(name,x)]\<bullet>trm" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14461
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14462
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14463
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14464
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14465
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14466
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14467
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14468
lemma crename_NotR':
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14469
  assumes a: "R[a\<turnstile>c>b] = NotR (x).N c" "x\<sharp>R" "c\<sharp>a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14470
  shows "(\<exists>N'. (R = NotR (x).N' c) \<and> N'[a\<turnstile>c>b] = N) \<or> (\<exists>N'. (R = NotR (x).N' a) \<and> b=c \<and> N'[a\<turnstile>c>b] = N)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14471
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14472
apply(nominal_induct R avoiding: a b c x N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14473
apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14474
apply(rule_tac x="[(name,x)]\<bullet>trm" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14475
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14476
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14477
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14478
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14479
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14480
apply(rule_tac x="[(name,x)]\<bullet>trm" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14481
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14482
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14483
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14484
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14485
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14486
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14487
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14488
lemma crename_NotR_aux:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14489
  assumes a: "R[a\<turnstile>c>b] = NotR (x).N c" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14490
  shows "(a=c \<and> a=b) \<or> (a\<noteq>c)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14491
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14492
apply(nominal_induct R avoiding: a b c x N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14493
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14494
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14495
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14496
lemma crename_NotL:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14497
  assumes a: "R[a\<turnstile>c>b] = NotL <c>.N y" "c\<sharp>(R,a,b)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14498
  shows "\<exists>N'. (R = NotL <c>.N' y) \<and> N'[a\<turnstile>c>b] = N" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14499
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14500
apply(nominal_induct R avoiding: a b c y N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14501
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14502
apply(rule_tac x="[(coname,c)]\<bullet>trm" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14503
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14504
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14505
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14506
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14507
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14508
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14509
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14510
lemma crename_AndL1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14511
  assumes a: "R[a\<turnstile>c>b] = AndL1 (x).N y" "x\<sharp>R"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14512
  shows "\<exists>N'. (R = AndL1 (x).N' y) \<and> N'[a\<turnstile>c>b] = N" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14513
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14514
apply(nominal_induct R avoiding: a b x y N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14515
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14516
apply(rule_tac x="[(name1,x)]\<bullet>trm" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14517
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14518
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14519
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14520
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14521
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14522
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14523
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14524
lemma crename_AndL2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14525
  assumes a: "R[a\<turnstile>c>b] = AndL2 (x).N y" "x\<sharp>R"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14526
  shows "\<exists>N'. (R = AndL2 (x).N' y) \<and> N'[a\<turnstile>c>b] = N" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14527
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14528
apply(nominal_induct R avoiding: a b x y N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14529
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14530
apply(rule_tac x="[(name1,x)]\<bullet>trm" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14531
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14532
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14533
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14534
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14535
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14536
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14537
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14538
lemma crename_AndR_aux:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14539
  assumes a: "R[a\<turnstile>c>b] = AndR <c>.M <d>.N e" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14540
  shows "(a=e \<and> a=b) \<or> (a\<noteq>e)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14541
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14542
apply(nominal_induct R avoiding: a b c d e M N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14543
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14544
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14545
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14546
lemma crename_AndR:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14547
  assumes a: "R[a\<turnstile>c>b] = AndR <c>.M <d>.N e" "c\<sharp>(a,b,d,e,N,R)" "d\<sharp>(a,b,c,e,M,R)" "e\<sharp>(a,b)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14548
  shows "\<exists>M' N'. R = AndR <c>.M' <d>.N' e \<and> M'[a\<turnstile>c>b] = M \<and> N'[a\<turnstile>c>b] = N \<and> c\<sharp>N' \<and> d\<sharp>M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14549
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14550
apply(nominal_induct R avoiding: a b c d e M N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14551
apply(auto split: if_splits simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14552
apply(simp add: fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14553
apply(rule_tac x="[(coname2,d)]\<bullet>trm2" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14554
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14555
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14556
apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14557
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14558
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14559
apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14560
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14561
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14562
apply(rule_tac x="[(coname2,d)]\<bullet>trm2" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14563
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14564
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14565
apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14566
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14567
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14568
apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14569
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14570
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14571
apply(rule_tac x="[(coname2,d)]\<bullet>trm2" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14572
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14573
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14574
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14575
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14576
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14577
apply(drule_tac s="trm2[a\<turnstile>c>b]" in  sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14578
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14579
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14580
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14581
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14582
lemma crename_AndR':
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14583
  assumes a: "R[a\<turnstile>c>b] = AndR <c>.M <d>.N e" "c\<sharp>(a,b,d,e,N,R)" "d\<sharp>(a,b,c,e,M,R)" "e\<sharp>a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14584
  shows "(\<exists>M' N'. R = AndR <c>.M' <d>.N' e \<and> M'[a\<turnstile>c>b] = M \<and> N'[a\<turnstile>c>b] = N \<and> c\<sharp>N' \<and> d\<sharp>M') \<or>
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14585
         (\<exists>M' N'. R = AndR <c>.M' <d>.N' a \<and> b=e \<and> M'[a\<turnstile>c>b] = M \<and> N'[a\<turnstile>c>b] = N \<and> c\<sharp>N' \<and> d\<sharp>M')"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14586
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14587
apply(nominal_induct R avoiding: a b c d e M N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14588
apply(auto split: if_splits simp add: trm.inject alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14589
apply(auto split: if_splits simp add: trm.inject alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14590
apply(auto split: if_splits simp add: trm.inject alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14591
apply(auto split: if_splits simp add: trm.inject alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14592
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm trm.inject alpha)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14593
apply(case_tac "coname3=a")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14594
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14595
apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14596
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14597
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14598
apply(rule_tac x="[(coname2,d)]\<bullet>trm2" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14599
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14600
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm trm.inject alpha split: if_splits)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14601
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14602
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14603
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14604
apply(drule_tac s="trm2[a\<turnstile>c>e]" in  sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14605
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14606
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14607
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14608
apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14609
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14610
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14611
apply(rule_tac x="[(coname2,d)]\<bullet>trm2" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14612
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14613
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm trm.inject alpha split: if_splits)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14614
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14615
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14616
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14617
apply(drule_tac s="trm2[a\<turnstile>c>b]" in  sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14618
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14619
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14620
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14621
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14622
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14623
lemma crename_OrR1_aux:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14624
  assumes a: "R[a\<turnstile>c>b] = OrR1 <c>.M e" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14625
  shows "(a=e \<and> a=b) \<or> (a\<noteq>e)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14626
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14627
apply(nominal_induct R avoiding: a b c e M rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14628
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14629
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14630
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14631
lemma crename_OrR1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14632
  assumes a: "R[a\<turnstile>c>b] = OrR1 <c>.N d" "c\<sharp>(R,a,b)" "d\<sharp>(a,b)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14633
  shows "\<exists>N'. (R = OrR1 <c>.N' d) \<and> N'[a\<turnstile>c>b] = N" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14634
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14635
apply(nominal_induct R avoiding: a b c d N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14636
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14637
apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14638
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14639
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14640
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14641
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14642
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14643
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14644
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14645
lemma crename_OrR1':
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14646
  assumes a: "R[a\<turnstile>c>b] = OrR1 <c>.N d" "c\<sharp>(R,a,b)" "d\<sharp>a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14647
  shows "(\<exists>N'. (R = OrR1 <c>.N' d) \<and> N'[a\<turnstile>c>b] = N) \<or>
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14648
         (\<exists>N'. (R = OrR1 <c>.N' a) \<and> b=d \<and> N'[a\<turnstile>c>b] = N)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14649
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14650
apply(nominal_induct R avoiding: a b c d N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14651
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14652
apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14653
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14654
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14655
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14656
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14657
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14658
apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14659
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14660
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14661
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14662
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14663
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14664
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14665
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14666
lemma crename_OrR2_aux:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14667
  assumes a: "R[a\<turnstile>c>b] = OrR2 <c>.M e" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14668
  shows "(a=e \<and> a=b) \<or> (a\<noteq>e)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14669
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14670
apply(nominal_induct R avoiding: a b c e M rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14671
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14672
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14673
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14674
lemma crename_OrR2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14675
  assumes a: "R[a\<turnstile>c>b] = OrR2 <c>.N d" "c\<sharp>(R,a,b)" "d\<sharp>(a,b)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14676
  shows "\<exists>N'. (R = OrR2 <c>.N' d) \<and> N'[a\<turnstile>c>b] = N" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14677
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14678
apply(nominal_induct R avoiding: a b c d N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14679
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14680
apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14681
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14682
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14683
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14684
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14685
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14686
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14687
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14688
lemma crename_OrR2':
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14689
  assumes a: "R[a\<turnstile>c>b] = OrR2 <c>.N d" "c\<sharp>(R,a,b)" "d\<sharp>a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14690
  shows "(\<exists>N'. (R = OrR2 <c>.N' d) \<and> N'[a\<turnstile>c>b] = N) \<or>
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14691
         (\<exists>N'. (R = OrR2 <c>.N' a) \<and> b=d \<and> N'[a\<turnstile>c>b] = N)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14692
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14693
apply(nominal_induct R avoiding: a b c d N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14694
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14695
apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14696
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14697
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14698
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14699
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14700
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14701
apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14702
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14703
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14704
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14705
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14706
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14707
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14708
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14709
lemma crename_OrL:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14710
  assumes a: "R[a\<turnstile>c>b] = OrL (x).M (y).N z" "x\<sharp>(y,z,N,R)" "y\<sharp>(x,z,M,R)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14711
  shows "\<exists>M' N'. R = OrL (x).M' (y).N' z \<and> M'[a\<turnstile>c>b] = M \<and> N'[a\<turnstile>c>b] = N \<and> x\<sharp>N' \<and> y\<sharp>M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14712
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14713
apply(nominal_induct R avoiding: a b x y z M N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14714
apply(auto split: if_splits simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14715
apply(rule_tac x="[(name2,y)]\<bullet>trm2" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14716
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14717
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14718
apply(rule_tac x="[(name1,x)]\<bullet>trm1" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14719
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14720
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14721
apply(rule_tac x="[(name1,x)]\<bullet>trm1" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14722
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14723
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14724
apply(rule_tac x="[(name2,y)]\<bullet>trm2" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14725
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14726
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14727
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14728
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14729
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14730
apply(drule_tac s="trm2[a\<turnstile>c>b]" in  sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14731
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14732
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14733
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14734
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14735
lemma crename_ImpL:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14736
  assumes a: "R[a\<turnstile>c>b] = ImpL <c>.M (y).N z" "c\<sharp>(a,b,N,R)" "y\<sharp>(z,M,R)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14737
  shows "\<exists>M' N'. R = ImpL <c>.M' (y).N' z \<and> M'[a\<turnstile>c>b] = M \<and> N'[a\<turnstile>c>b] = N \<and> c\<sharp>N' \<and> y\<sharp>M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14738
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14739
apply(nominal_induct R avoiding: a b c y z M N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14740
apply(auto split: if_splits simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14741
apply(rule_tac x="[(name1,y)]\<bullet>trm2" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14742
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14743
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14744
apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14745
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14746
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14747
apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14748
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14749
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14750
apply(rule_tac x="[(name1,y)]\<bullet>trm2" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14751
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14752
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14753
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14754
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14755
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14756
apply(drule_tac s="trm2[a\<turnstile>c>b]" in  sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14757
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14758
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14759
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14760
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14761
lemma crename_ImpR_aux:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14762
  assumes a: "R[a\<turnstile>c>b] = ImpR (x).<c>.M e" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14763
  shows "(a=e \<and> a=b) \<or> (a\<noteq>e)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14764
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14765
apply(nominal_induct R avoiding: x a b c e M rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14766
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14767
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14768
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14769
lemma crename_ImpR:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14770
  assumes a: "R[a\<turnstile>c>b] = ImpR (x).<c>.N d" "c\<sharp>(R,a,b)" "d\<sharp>(a,b)" "x\<sharp>R" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14771
  shows "\<exists>N'. (R = ImpR (x).<c>.N' d) \<and> N'[a\<turnstile>c>b] = N" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14772
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14773
apply(nominal_induct R avoiding: a b x c d N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14774
apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_perm alpha abs_fresh trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14775
apply(rule_tac x="[(name,x)]\<bullet>trm" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14776
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14777
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14778
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14779
apply(rule_tac x="[(name,x)]\<bullet>[(coname1, c)]\<bullet>trm" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14780
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14781
apply(simp add: abs_supp fin_supp abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14782
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14783
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14784
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14785
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14786
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14787
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14788
lemma crename_ImpR':
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14789
  assumes a: "R[a\<turnstile>c>b] = ImpR (x).<c>.N d" "c\<sharp>(R,a,b)" "x\<sharp>R" "d\<sharp>a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14790
  shows "(\<exists>N'. (R = ImpR (x).<c>.N' d) \<and> N'[a\<turnstile>c>b] = N) \<or>
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14791
         (\<exists>N'. (R = ImpR (x).<c>.N' a) \<and> b=d \<and> N'[a\<turnstile>c>b] = N)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14792
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14793
apply(nominal_induct R avoiding: x a b c d N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14794
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject abs_perm calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14795
apply(rule_tac x="[(name,x)]\<bullet>[(coname1,c)]\<bullet>trm" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14796
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14797
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod abs_supp fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14798
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14799
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14800
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14801
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14802
apply(rule_tac x="[(name,x)]\<bullet>[(coname1,c)]\<bullet>trm" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14803
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14804
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod abs_supp fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14805
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14806
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14807
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14808
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14809
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14810
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14811
lemma crename_ax2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14812
  assumes a: "N[a\<turnstile>c>b] = Ax x c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14813
  shows "\<exists>d. N = Ax x d"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14814
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14815
apply(nominal_induct N avoiding: a b rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14816
apply(auto split: if_splits)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14817
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14818
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14819
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14820
lemma crename_interesting1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14821
  assumes a: "distinct [a,b,c]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14822
  shows "M[a\<turnstile>c>c][c\<turnstile>c>b] = M[c\<turnstile>c>b][a\<turnstile>c>b]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14823
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14824
apply(nominal_induct M avoiding: a c b rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14825
apply(auto simp add: rename_fresh simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14826
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14827
apply(rotate_tac 12)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14828
apply(drule_tac x="a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14829
apply(rotate_tac 15)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14830
apply(drule_tac x="c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14831
apply(rotate_tac 15)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14832
apply(drule_tac x="b" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14833
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14834
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14835
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14836
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14837
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14838
lemma crename_interesting2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14839
  assumes a: "a\<noteq>c" "a\<noteq>d" "a\<noteq>b" "c\<noteq>d" "b\<noteq>c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14840
  shows "M[a\<turnstile>c>b][c\<turnstile>c>d] = M[c\<turnstile>c>d][a\<turnstile>c>b]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14841
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14842
apply(nominal_induct M avoiding: a c b d rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14843
apply(auto simp add: rename_fresh simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14844
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14845
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14846
lemma crename_interesting3:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14847
  shows "M[a\<turnstile>c>c][x\<turnstile>n>y] = M[x\<turnstile>n>y][a\<turnstile>c>c]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14848
apply(nominal_induct M avoiding: a c x y rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14849
apply(auto simp add: rename_fresh simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14850
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14851
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14852
lemma crename_credu:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14853
  assumes a: "(M[a\<turnstile>c>b]) \<longrightarrow>\<^isub>c M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14854
  shows "\<exists>M0. M0[a\<turnstile>c>b]=M' \<and> M \<longrightarrow>\<^isub>c M0"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14855
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14856
apply(nominal_induct M\<equiv>"M[a\<turnstile>c>b]" M' avoiding: M a b rule: c_redu.strong_induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14857
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14858
apply(drule crename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14859
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14860
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14861
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14862
apply(rule_tac x="M'{a:=(x).N'}" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14863
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14864
apply(simp add: fresh_atm abs_fresh subst_comm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14865
apply(rule c_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14866
apply(auto dest: not_fic_crename)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14867
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14868
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14869
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14870
apply(drule crename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14871
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14872
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14873
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14874
apply(rule_tac x="N'{x:=<a>.M'}" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14875
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14876
apply(simp add: fresh_atm abs_fresh subst_comm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14877
apply(rule c_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14878
apply(auto dest: not_fin_crename)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14879
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14880
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14881
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14882
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14883
lemma crename_lredu:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14884
  assumes a: "(M[a\<turnstile>c>b]) \<longrightarrow>\<^isub>l M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14885
  shows "\<exists>M0. M0[a\<turnstile>c>b]=M' \<and> M \<longrightarrow>\<^isub>l M0"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14886
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14887
apply(nominal_induct M\<equiv>"M[a\<turnstile>c>b]" M' avoiding: M a b rule: l_redu.strong_induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14888
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14889
apply(drule crename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14890
apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14891
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14892
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14893
apply(case_tac "aa=ba")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14894
apply(simp add: crename_id)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14895
apply(rule l_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14896
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14897
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14898
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14899
apply(frule crename_ax2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14900
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14901
apply(case_tac "d=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14902
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14903
apply(rule_tac x="M'[a\<turnstile>c>aa]" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14904
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14905
apply(rule crename_interesting1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14906
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14907
apply(rule l_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14908
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14909
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14910
apply(auto dest: fic_crename simp add: fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14911
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14912
apply(rule_tac x="M'[a\<turnstile>c>b]" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14913
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14914
apply(rule crename_interesting2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14915
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14916
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14917
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14918
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14919
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14920
apply(rule l_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14921
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14922
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14923
apply(auto dest: fic_crename simp add: fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14924
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14925
apply(drule crename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14926
apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14927
apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14928
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14929
apply(case_tac "aa=b")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14930
apply(simp add: crename_id)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14931
apply(rule l_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14932
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14933
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14934
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14935
apply(frule crename_ax2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14936
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14937
apply(case_tac "d=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14938
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14939
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14940
apply(rule_tac x="N'[x\<turnstile>n>y]" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14941
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14942
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14943
apply(rule crename_interesting3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14944
apply(rule l_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14945
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14946
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14947
apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14948
(* LNot *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14949
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14950
apply(drule crename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14951
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14952
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14953
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14954
apply(drule crename_NotR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14955
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14956
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14957
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14958
apply(drule crename_NotL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14959
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14960
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14961
apply(rule_tac x="Cut <b>.N'b (x).N'a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14962
apply(simp add: fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14963
apply(rule l_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14964
apply(auto simp add: fresh_prod intro: crename_fresh_interesting2)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14965
apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting2)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14966
apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14967
apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14968
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14969
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14970
(* LAnd1 *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14971
apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14972
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14973
apply(drule crename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14974
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14975
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14976
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14977
apply(drule crename_AndR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14978
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14979
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14980
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14981
apply(drule crename_AndL1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14982
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14983
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14984
apply(rule_tac x="Cut <a1>.M'a (x).N'a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14985
apply(simp add: fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14986
apply(rule l_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14987
apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: crename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14988
apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting2)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14989
apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14990
apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14991
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14992
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14993
(* LAnd2 *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14994
apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14995
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14996
apply(drule crename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14997
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14998
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 14999
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15000
apply(drule crename_AndR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15001
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15002
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15003
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15004
apply(drule crename_AndL2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15005
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15006
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15007
apply(rule_tac x="Cut <a2>.N'b (x).N'a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15008
apply(simp add: fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15009
apply(rule l_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15010
apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: crename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15011
apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting2)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15012
apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15013
apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15014
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15015
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15016
(* LOr1 *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15017
apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15018
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15019
apply(drule crename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15020
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15021
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15022
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15023
apply(drule crename_OrL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15024
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15025
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15026
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15027
apply(drule crename_OrR1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15028
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15029
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15030
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15031
apply(rule_tac x="Cut <a>.N' (x1).M'a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15032
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15033
apply(simp add: abs_fresh fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15034
apply(rule l_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15035
apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: crename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15036
apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting2)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15037
apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15038
apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15039
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15040
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15041
(* LOr2 *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15042
apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15043
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15044
apply(drule crename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15045
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15046
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15047
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15048
apply(drule crename_OrL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15049
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15050
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15051
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15052
apply(drule crename_OrR2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15053
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15054
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15055
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15056
apply(rule_tac x="Cut <a>.N' (x2).N'a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15057
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15058
apply(simp add: abs_fresh fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15059
apply(rule l_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15060
apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: crename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15061
apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting2)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15062
apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15063
apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15064
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15065
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15066
(* ImpL *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15067
apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15068
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15069
apply(drule crename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15070
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15071
apply(simp add: fresh_prod abs_fresh fresh_atm abs_supp fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15072
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15073
apply(drule crename_ImpL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15074
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15075
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15076
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15077
apply(drule crename_ImpR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15078
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15079
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15080
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15081
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15082
apply(rule_tac x="Cut <a>.(Cut <c>.M'a (x).N') (y).N'a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15083
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15084
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15085
apply(rule l_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15086
apply(auto simp add: fresh_atm abs_fresh abs_supp fin_supp fresh_prod intro: crename_fresh_interesting2)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15087
apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: crename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15088
apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: crename_fresh_interesting2)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15089
apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: crename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15090
apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: crename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15091
apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: crename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15092
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15093
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15094
lemma crename_aredu:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15095
  assumes a: "(M[a\<turnstile>c>b]) \<longrightarrow>\<^isub>a M'" "a\<noteq>b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15096
  shows "\<exists>M0. M0[a\<turnstile>c>b]=M' \<and> M \<longrightarrow>\<^isub>a M0"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15097
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15098
apply(nominal_induct M\<equiv>"M[a\<turnstile>c>b]" M' avoiding: M a b rule: a_redu.strong_induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15099
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15100
apply(drule  crename_lredu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15101
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15102
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15103
apply(drule  crename_credu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15104
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15105
(* Cut *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15106
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15107
apply(drule crename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15108
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15109
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15110
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15111
apply(drule_tac x="M'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15112
apply(drule_tac x="aa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15113
apply(drule_tac x="b" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15114
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15115
apply(rule_tac x="Cut <a>.M0 (x).N'" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15116
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15117
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15118
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15119
apply(rule crename.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15120
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15121
apply(drule crename_fresh_interesting2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15122
apply(simp add: fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15123
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15124
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15125
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15126
apply(drule crename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15127
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15128
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15129
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15130
apply(drule_tac x="N'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15131
apply(drule_tac x="aa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15132
apply(drule_tac x="b" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15133
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15134
apply(rule_tac x="Cut <a>.M' (x).M0" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15135
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15136
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15137
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15138
apply(rule crename.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15139
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15140
apply(drule crename_fresh_interesting1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15141
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15142
apply(simp add: fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15143
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15144
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15145
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15146
(* NotL *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15147
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15148
apply(drule crename_NotL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15149
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15150
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15151
apply(drule_tac x="N'" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15152
apply(drule_tac x="aa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15153
apply(drule_tac x="b" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15154
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15155
apply(rule_tac x="NotL <a>.M0 x" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15156
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15157
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15158
(* NotR *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15159
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15160
apply(frule crename_NotR_aux)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15161
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15162
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15163
apply(drule crename_NotR')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15164
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15165
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15166
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15167
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15168
apply(drule_tac x="N'" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15169
apply(drule_tac x="aa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15170
apply(drule_tac x="b" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15171
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15172
apply(rule_tac x="NotR (x).M0 a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15173
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15174
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15175
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15176
apply(drule_tac x="N'" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15177
apply(drule_tac x="aa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15178
apply(drule_tac x="a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15179
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15180
apply(rule_tac x="NotR (x).M0 aa" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15181
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15182
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15183
(* AndR *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15184
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15185
apply(frule crename_AndR_aux)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15186
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15187
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15188
apply(drule crename_AndR')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15189
apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15190
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15191
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15192
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15193
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15194
apply(drule_tac x="M'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15195
apply(drule_tac x="aa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15196
apply(drule_tac x="ba" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15197
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15198
apply(rule_tac x="AndR <a>.M0 <b>.N' c" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15199
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15200
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15201
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15202
apply(rule crename.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15203
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15204
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15205
apply(auto intro: fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15206
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15207
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15208
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15209
apply(drule_tac x="M'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15210
apply(drule_tac x="aa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15211
apply(drule_tac x="c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15212
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15213
apply(rule_tac x="AndR <a>.M0 <b>.N' aa" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15214
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15215
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15216
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15217
apply(rule crename.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15218
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15219
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15220
apply(auto intro: fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15221
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15222
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15223
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15224
apply(frule crename_AndR_aux)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15225
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15226
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15227
apply(drule crename_AndR')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15228
apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15229
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15230
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15231
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15232
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15233
apply(drule_tac x="N'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15234
apply(drule_tac x="aa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15235
apply(drule_tac x="ba" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15236
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15237
apply(rule_tac x="AndR <a>.M' <b>.M0 c" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15238
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15239
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15240
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15241
apply(rule crename.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15242
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15243
apply(auto intro: fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15244
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15245
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15246
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15247
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15248
apply(drule_tac x="N'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15249
apply(drule_tac x="aa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15250
apply(drule_tac x="c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15251
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15252
apply(rule_tac x="AndR <a>.M' <b>.M0 aa" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15253
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15254
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15255
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15256
apply(rule crename.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15257
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15258
apply(auto intro: fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15259
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15260
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15261
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15262
(* AndL1 *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15263
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15264
apply(drule crename_AndL1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15265
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15266
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15267
apply(drule_tac x="N'" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15268
apply(drule_tac x="a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15269
apply(drule_tac x="b" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15270
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15271
apply(rule_tac x="AndL1 (x).M0 y" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15272
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15273
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15274
(* AndL2 *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15275
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15276
apply(drule crename_AndL2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15277
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15278
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15279
apply(drule_tac x="N'" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15280
apply(drule_tac x="a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15281
apply(drule_tac x="b" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15282
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15283
apply(rule_tac x="AndL2 (x).M0 y" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15284
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15285
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15286
(* OrL *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15287
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15288
apply(drule crename_OrL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15289
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15290
apply(auto simp add: fresh_atm fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15291
apply(auto simp add: fresh_atm fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15292
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15293
apply(drule_tac x="M'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15294
apply(drule_tac x="a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15295
apply(drule_tac x="b" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15296
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15297
apply(rule_tac x="OrL (x).M0 (y).N' z" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15298
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15299
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15300
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15301
apply(rule crename.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15302
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15303
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15304
apply(auto intro: fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15305
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15306
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15307
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15308
apply(drule crename_OrL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15309
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15310
apply(auto simp add: fresh_atm fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15311
apply(auto simp add: fresh_atm fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15312
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15313
apply(drule_tac x="N'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15314
apply(drule_tac x="a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15315
apply(drule_tac x="b" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15316
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15317
apply(rule_tac x="OrL (x).M' (y).M0 z" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15318
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15319
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15320
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15321
apply(rule crename.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15322
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15323
apply(auto intro: fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15324
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15325
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15326
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15327
(* OrR1 *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15328
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15329
apply(frule crename_OrR1_aux)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15330
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15331
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15332
apply(drule crename_OrR1')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15333
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15334
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15335
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15336
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15337
apply(drule_tac x="N'" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15338
apply(drule_tac x="aa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15339
apply(drule_tac x="ba" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15340
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15341
apply(rule_tac x="OrR1 <a>.M0 b" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15342
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15343
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15344
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15345
apply(drule_tac x="N'" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15346
apply(drule_tac x="aa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15347
apply(drule_tac x="b" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15348
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15349
apply(rule_tac x="OrR1 <a>.M0 aa" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15350
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15351
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15352
(* OrR2 *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15353
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15354
apply(frule crename_OrR2_aux)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15355
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15356
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15357
apply(drule crename_OrR2')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15358
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15359
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15360
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15361
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15362
apply(drule_tac x="N'" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15363
apply(drule_tac x="aa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15364
apply(drule_tac x="ba" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15365
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15366
apply(rule_tac x="OrR2 <a>.M0 b" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15367
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15368
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15369
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15370
apply(drule_tac x="N'" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15371
apply(drule_tac x="aa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15372
apply(drule_tac x="b" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15373
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15374
apply(rule_tac x="OrR2 <a>.M0 aa" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15375
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15376
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15377
(* ImpL *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15378
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15379
apply(drule crename_ImpL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15380
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15381
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15382
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15383
apply(drule_tac x="M'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15384
apply(drule_tac x="aa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15385
apply(drule_tac x="b" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15386
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15387
apply(rule_tac x="ImpL <a>.M0 (x).N' y" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15388
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15389
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15390
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15391
apply(rule crename.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15392
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15393
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15394
apply(auto intro: fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15395
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15396
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15397
apply(drule crename_ImpL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15398
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15399
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15400
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15401
apply(drule_tac x="N'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15402
apply(drule_tac x="aa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15403
apply(drule_tac x="b" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15404
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15405
apply(rule_tac x="ImpL <a>.M' (x).M0 y" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15406
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15407
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15408
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15409
apply(rule crename.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15410
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15411
apply(auto intro: fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15412
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15413
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15414
(* ImpR *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15415
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15416
apply(frule crename_ImpR_aux)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15417
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15418
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15419
apply(drule crename_ImpR')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15420
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15421
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15422
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15423
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15424
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15425
apply(drule_tac x="N'" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15426
apply(drule_tac x="aa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15427
apply(drule_tac x="ba" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15428
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15429
apply(rule_tac x="ImpR (x).<a>.M0 b" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15430
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15431
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15432
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15433
apply(drule_tac x="N'" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15434
apply(drule_tac x="aa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15435
apply(drule_tac x="b" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15436
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15437
apply(rule_tac x="ImpR (x).<a>.M0 aa" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15438
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15439
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15440
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15441
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15442
lemma SNa_preserved_renaming1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15443
  assumes a: "SNa M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15444
  shows "SNa (M[a\<turnstile>c>b])"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15445
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15446
apply(induct rule: SNa_induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15447
apply(case_tac "a=b")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15448
apply(simp add: crename_id)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15449
apply(rule SNaI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15450
apply(drule crename_aredu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15451
apply(blast)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15452
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15453
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15454
lemma nrename_interesting1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15455
  assumes a: "distinct [x,y,z]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15456
  shows "M[x\<turnstile>n>z][z\<turnstile>n>y] = M[z\<turnstile>n>y][x\<turnstile>n>y]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15457
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15458
apply(nominal_induct M avoiding: x y z rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15459
apply(auto simp add: rename_fresh simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15460
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15461
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15462
apply(rotate_tac 12)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15463
apply(drule_tac x="x" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15464
apply(rotate_tac 15)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15465
apply(drule_tac x="y" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15466
apply(rotate_tac 15)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15467
apply(drule_tac x="z" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15468
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15469
apply(rotate_tac 11)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15470
apply(drule_tac x="x" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15471
apply(rotate_tac 14)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15472
apply(drule_tac x="y" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15473
apply(rotate_tac 14)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15474
apply(drule_tac x="z" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15475
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15476
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15477
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15478
lemma nrename_interesting2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15479
  assumes a: "x\<noteq>z" "x\<noteq>u" "x\<noteq>y" "z\<noteq>u" "y\<noteq>z"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15480
  shows "M[x\<turnstile>n>y][z\<turnstile>n>u] = M[z\<turnstile>n>u][x\<turnstile>n>y]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15481
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15482
apply(nominal_induct M avoiding: x y z u rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15483
apply(auto simp add: rename_fresh simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15484
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15485
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15486
lemma not_fic_nrename_aux:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15487
  assumes a: "fic M c" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15488
  shows "fic (M[x\<turnstile>n>y]) c" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15489
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15490
apply(nominal_induct M avoiding: c x y rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15491
apply(auto dest!: fic_elims intro!: fic.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15492
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15493
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15494
lemma not_fic_nrename:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15495
  assumes a: "\<not>(fic (M[x\<turnstile>n>y]) c)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15496
  shows "\<not>(fic M c)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15497
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15498
apply(auto dest:  not_fic_nrename_aux)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15499
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15500
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15501
lemma fin_nrename:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15502
  assumes a: "fin M z" "z\<sharp>(x,y)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15503
  shows "fin (M[x\<turnstile>n>y]) z" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15504
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15505
apply(nominal_induct M avoiding: x y z rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15506
apply(auto dest!: fin_elims intro!: fin.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15507
           split: if_splits)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15508
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15509
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15510
lemma nrename_fresh_interesting1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15511
  fixes z::"name"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15512
  assumes a: "z\<sharp>(M[x\<turnstile>n>y])" "z\<sharp>(x,y)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15513
  shows "z\<sharp>M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15514
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15515
apply(nominal_induct M avoiding: x y z rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15516
apply(auto split: if_splits simp add: abs_fresh abs_supp fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15517
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15518
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15519
lemma nrename_fresh_interesting2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15520
  fixes c::"coname"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15521
  assumes a: "c\<sharp>(M[x\<turnstile>n>y])" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15522
  shows "c\<sharp>M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15523
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15524
apply(nominal_induct M avoiding: x y c rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15525
apply(auto split: if_splits simp add: abs_fresh abs_supp fin_supp fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15526
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15527
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15528
lemma fin_nrename2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15529
  assumes a: "fin (M[x\<turnstile>n>y]) z" "z\<sharp>(x,y)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15530
  shows "fin M z" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15531
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15532
apply(nominal_induct M avoiding: x y z rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15533
apply(auto dest!: fin_elims intro!: fin.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15534
           split: if_splits)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15535
apply(auto dest: nrename_fresh_interesting1 simp add: fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15536
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15537
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15538
lemma nrename_Cut:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15539
  assumes a: "R[x\<turnstile>n>y] = Cut <c>.M (z).N" "c\<sharp>(N,R)" "z\<sharp>(x,y,M,R)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15540
  shows "\<exists>M' N'. R = Cut <c>.M' (z).N' \<and> M'[x\<turnstile>n>y] = M \<and> N'[x\<turnstile>n>y] = N \<and> c\<sharp>N' \<and> z\<sharp>M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15541
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15542
apply(nominal_induct R avoiding: c y x z M N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15543
apply(auto split: if_splits)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15544
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15545
apply(auto simp add: alpha fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15546
apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15547
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15548
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15549
apply(rule_tac x="[(name,z)]\<bullet>trm2" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15550
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15551
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15552
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15553
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15554
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15555
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15556
apply(auto simp add: fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15557
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15558
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15559
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15560
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15561
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15562
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15563
lemma nrename_NotR:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15564
  assumes a: "R[x\<turnstile>n>y] = NotR (z).N c" "z\<sharp>(R,x,y)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15565
  shows "\<exists>N'. (R = NotR (z).N' c) \<and> N'[x\<turnstile>n>y] = N" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15566
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15567
apply(nominal_induct R avoiding: x y c z N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15568
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15569
apply(rule_tac x="[(name,z)]\<bullet>trm" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15570
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15571
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15572
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15573
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15574
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15575
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15576
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15577
lemma nrename_NotL:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15578
  assumes a: "R[x\<turnstile>n>y] = NotL <c>.N z" "c\<sharp>R" "z\<sharp>(x,y)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15579
  shows "\<exists>N'. (R = NotL <c>.N' z) \<and> N'[x\<turnstile>n>y] = N" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15580
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15581
apply(nominal_induct R avoiding: x y c z N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15582
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15583
apply(rule_tac x="[(coname,c)]\<bullet>trm" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15584
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15585
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15586
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15587
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15588
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15589
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15590
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15591
lemma nrename_NotL':
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15592
  assumes a: "R[x\<turnstile>n>y] = NotL <c>.N u" "c\<sharp>R" "x\<noteq>y" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15593
  shows "(\<exists>N'. (R = NotL <c>.N' u) \<and> N'[x\<turnstile>n>y] = N) \<or> (\<exists>N'. (R = NotL <c>.N' x) \<and> y=u \<and> N'[x\<turnstile>n>y] = N)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15594
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15595
apply(nominal_induct R avoiding: y u c x N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15596
apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15597
apply(rule_tac x="[(coname,c)]\<bullet>trm" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15598
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15599
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15600
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15601
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15602
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15603
apply(rule_tac x="[(coname,c)]\<bullet>trm" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15604
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15605
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15606
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15607
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15608
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15609
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15610
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15611
lemma nrename_NotL_aux:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15612
  assumes a: "R[x\<turnstile>n>y] = NotL <c>.N u" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15613
  shows "(x=u \<and> x=y) \<or> (x\<noteq>u)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15614
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15615
apply(nominal_induct R avoiding: y u c x N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15616
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15617
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15618
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15619
lemma nrename_AndL1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15620
  assumes a: "R[x\<turnstile>n>y] = AndL1 (z).N u" "z\<sharp>(R,x,y)" "u\<sharp>(x,y)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15621
  shows "\<exists>N'. (R = AndL1 (z).N' u) \<and> N'[x\<turnstile>n>y] = N" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15622
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15623
apply(nominal_induct R avoiding: z u x y N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15624
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15625
apply(rule_tac x="[(name1,z)]\<bullet>trm" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15626
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15627
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15628
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15629
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15630
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15631
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15632
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15633
lemma nrename_AndL1':
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15634
  assumes a: "R[x\<turnstile>n>y] = AndL1 (v).N u" "v\<sharp>(R,u,x,y)" "x\<noteq>y" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15635
  shows "(\<exists>N'. (R = AndL1 (v).N' u) \<and> N'[x\<turnstile>n>y] = N) \<or> (\<exists>N'. (R = AndL1 (v).N' x) \<and> y=u \<and> N'[x\<turnstile>n>y] = N)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15636
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15637
apply(nominal_induct R avoiding: y u v x N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15638
apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15639
apply(rule_tac x="[(name1,v)]\<bullet>trm" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15640
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15641
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15642
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15643
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15644
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15645
apply(rule_tac x="[(name1,v)]\<bullet>trm" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15646
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15647
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15648
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15649
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15650
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15651
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15652
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15653
lemma nrename_AndL1_aux:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15654
  assumes a: "R[x\<turnstile>n>y] = AndL1 (v).N u" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15655
  shows "(x=u \<and> x=y) \<or> (x\<noteq>u)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15656
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15657
apply(nominal_induct R avoiding: y u v x N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15658
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15659
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15660
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15661
lemma nrename_AndL2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15662
  assumes a: "R[x\<turnstile>n>y] = AndL2 (z).N u" "z\<sharp>(R,x,y)" "u\<sharp>(x,y)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15663
  shows "\<exists>N'. (R = AndL2 (z).N' u) \<and> N'[x\<turnstile>n>y] = N" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15664
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15665
apply(nominal_induct R avoiding: z u x y N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15666
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15667
apply(rule_tac x="[(name1,z)]\<bullet>trm" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15668
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15669
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15670
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15671
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15672
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15673
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15674
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15675
lemma nrename_AndL2':
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15676
  assumes a: "R[x\<turnstile>n>y] = AndL2 (v).N u" "v\<sharp>(R,u,x,y)" "x\<noteq>y" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15677
  shows "(\<exists>N'. (R = AndL2 (v).N' u) \<and> N'[x\<turnstile>n>y] = N) \<or> (\<exists>N'. (R = AndL2 (v).N' x) \<and> y=u \<and> N'[x\<turnstile>n>y] = N)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15678
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15679
apply(nominal_induct R avoiding: y u v x N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15680
apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15681
apply(rule_tac x="[(name1,v)]\<bullet>trm" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15682
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15683
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15684
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15685
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15686
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15687
apply(rule_tac x="[(name1,v)]\<bullet>trm" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15688
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15689
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15690
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15691
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15692
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15693
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15694
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15695
lemma nrename_AndL2_aux:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15696
  assumes a: "R[x\<turnstile>n>y] = AndL2 (v).N u" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15697
  shows "(x=u \<and> x=y) \<or> (x\<noteq>u)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15698
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15699
apply(nominal_induct R avoiding: y u v x N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15700
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15701
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15702
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15703
lemma nrename_AndR:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15704
  assumes a: "R[x\<turnstile>n>y] = AndR <c>.M <d>.N e" "c\<sharp>(d,e,N,R)" "d\<sharp>(c,e,M,R)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15705
  shows "\<exists>M' N'. R = AndR <c>.M' <d>.N' e \<and> M'[x\<turnstile>n>y] = M \<and> N'[x\<turnstile>n>y] = N \<and> c\<sharp>N' \<and> d\<sharp>M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15706
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15707
apply(nominal_induct R avoiding: x y c d e M N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15708
apply(auto split: if_splits simp add: trm.inject alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15709
apply(simp add: fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15710
apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15711
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15712
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15713
apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15714
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15715
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15716
apply(rule_tac x="[(coname2,d)]\<bullet>trm2" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15717
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15718
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15719
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15720
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15721
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15722
apply(drule_tac s="trm2[x\<turnstile>n>y]" in  sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15723
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15724
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15725
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15726
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15727
lemma nrename_OrR1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15728
  assumes a: "R[x\<turnstile>n>y] = OrR1 <c>.N d" "c\<sharp>(R,d)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15729
  shows "\<exists>N'. (R = OrR1 <c>.N' d) \<and> N'[x\<turnstile>n>y] = N" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15730
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15731
apply(nominal_induct R avoiding: x y c d N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15732
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15733
apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15734
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15735
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15736
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15737
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15738
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15739
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15740
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15741
lemma nrename_OrR2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15742
  assumes a: "R[x\<turnstile>n>y] = OrR2 <c>.N d" "c\<sharp>(R,d)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15743
  shows "\<exists>N'. (R = OrR2 <c>.N' d) \<and> N'[x\<turnstile>n>y] = N" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15744
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15745
apply(nominal_induct R avoiding: x y c d N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15746
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15747
apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15748
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15749
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15750
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15751
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15752
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15753
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15754
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15755
lemma nrename_OrL:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15756
  assumes a: "R[u\<turnstile>n>v] = OrL (x).M (y).N z" "x\<sharp>(y,z,u,v,N,R)" "y\<sharp>(x,z,u,v,M,R)" "z\<sharp>(u,v)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15757
  shows "\<exists>M' N'. R = OrL (x).M' (y).N' z \<and> M'[u\<turnstile>n>v] = M \<and> N'[u\<turnstile>n>v] = N \<and> x\<sharp>N' \<and> y\<sharp>M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15758
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15759
apply(nominal_induct R avoiding: u v x y z M N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15760
apply(auto split: if_splits simp add: trm.inject alpha fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15761
apply(rule_tac x="[(name1,x)]\<bullet>trm1" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15762
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15763
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15764
apply(rule_tac x="[(name2,y)]\<bullet>trm2" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15765
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15766
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15767
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15768
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15769
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15770
apply(drule_tac s="trm2[u\<turnstile>n>v]" in  sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15771
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15772
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15773
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15774
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15775
lemma nrename_OrL':
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15776
  assumes a: "R[x\<turnstile>n>y] = OrL (v).M (w).N u" "v\<sharp>(R,N,u,x,y)" "w\<sharp>(R,M,u,x,y)" "x\<noteq>y" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15777
  shows "(\<exists>M' N'. (R = OrL (v).M' (w).N' u) \<and> M'[x\<turnstile>n>y] = M \<and> N'[x\<turnstile>n>y] = N) \<or> 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15778
         (\<exists>M' N'. (R = OrL (v).M' (w).N' x) \<and> y=u \<and> M'[x\<turnstile>n>y] = M \<and> N'[x\<turnstile>n>y] = N)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15779
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15780
apply(nominal_induct R avoiding: y x u v w M N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15781
apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15782
apply(rule_tac x="[(name1,v)]\<bullet>trm1" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15783
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15784
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15785
apply(rule_tac x="[(name2,w)]\<bullet>trm2" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15786
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15787
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15788
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15789
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15790
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15791
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15792
apply(drule_tac s="trm2[x\<turnstile>n>u]" in sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15793
apply(drule_tac pt_bij1[OF pt_name_inst,OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15794
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15795
apply(rule_tac x="[(name1,v)]\<bullet>trm1" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15796
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15797
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15798
apply(rule_tac x="[(name2,w)]\<bullet>trm2" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15799
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15800
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15801
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15802
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15803
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15804
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15805
apply(drule_tac s="trm2[x\<turnstile>n>y]" in sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15806
apply(drule_tac pt_bij1[OF pt_name_inst,OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15807
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15808
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15809
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15810
lemma nrename_OrL_aux:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15811
  assumes a: "R[x\<turnstile>n>y] = OrL (v).M (w).N u" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15812
  shows "(x=u \<and> x=y) \<or> (x\<noteq>u)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15813
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15814
apply(nominal_induct R avoiding: y x w u v M N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15815
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15816
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15817
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15818
lemma nrename_ImpL:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15819
  assumes a: "R[x\<turnstile>n>y] = ImpL <c>.M (u).N z" "c\<sharp>(N,R)" "u\<sharp>(y,x,z,M,R)" "z\<sharp>(x,y)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15820
  shows "\<exists>M' N'. R = ImpL <c>.M' (u).N' z \<and> M'[x\<turnstile>n>y] = M \<and> N'[x\<turnstile>n>y] = N \<and> c\<sharp>N' \<and> u\<sharp>M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15821
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15822
apply(nominal_induct R avoiding: u x c y z M N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15823
apply(auto split: if_splits simp add: trm.inject alpha fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15824
apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15825
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15826
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15827
apply(rule_tac x="[(name1,u)]\<bullet>trm2" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15828
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15829
apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15830
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15831
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15832
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15833
apply(drule_tac s="trm2[x\<turnstile>n>y]" in  sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15834
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15835
apply(simp add: eqvts calc_atm fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15836
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15837
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15838
lemma nrename_ImpL':
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15839
  assumes a: "R[x\<turnstile>n>y] = ImpL <c>.M (w).N u" "c\<sharp>(R,N)" "w\<sharp>(R,M,u,x,y)" "x\<noteq>y" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15840
  shows "(\<exists>M' N'. (R = ImpL <c>.M' (w).N' u) \<and> M'[x\<turnstile>n>y] = M \<and> N'[x\<turnstile>n>y] = N) \<or> 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15841
         (\<exists>M' N'. (R = ImpL <c>.M' (w).N' x) \<and> y=u \<and> M'[x\<turnstile>n>y] = M \<and> N'[x\<turnstile>n>y] = N)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15842
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15843
apply(nominal_induct R avoiding: y x u c w M N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15844
apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15845
apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15846
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15847
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15848
apply(rule_tac x="[(name1,w)]\<bullet>trm2" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15849
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15850
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15851
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15852
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15853
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15854
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15855
apply(drule_tac s="trm2[x\<turnstile>n>u]" in sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15856
apply(drule_tac pt_bij1[OF pt_name_inst,OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15857
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15858
apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15859
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15860
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15861
apply(rule_tac x="[(name1,w)]\<bullet>trm2" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15862
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15863
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15864
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15865
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15866
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15867
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15868
apply(drule_tac s="trm2[x\<turnstile>n>y]" in sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15869
apply(drule_tac pt_bij1[OF pt_name_inst,OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15870
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15871
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15872
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15873
lemma nrename_ImpL_aux:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15874
  assumes a: "R[x\<turnstile>n>y] = ImpL <c>.M (w).N u" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15875
  shows "(x=u \<and> x=y) \<or> (x\<noteq>u)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15876
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15877
apply(nominal_induct R avoiding: y x w u c M N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15878
apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15879
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15880
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15881
lemma nrename_ImpR:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15882
  assumes a: "R[u\<turnstile>n>v] = ImpR (x).<c>.N d" "c\<sharp>(R,d)" "x\<sharp>(R,u,v)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15883
  shows "\<exists>N'. (R = ImpR (x).<c>.N' d) \<and> N'[u\<turnstile>n>v] = N" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15884
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15885
apply(nominal_induct R avoiding: u v x c d N rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15886
apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_perm alpha abs_fresh trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15887
apply(rule_tac x="[(name,x)]\<bullet>trm" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15888
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15889
apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15890
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15891
apply(rule_tac x="[(name,x)]\<bullet>[(coname1, c)]\<bullet>trm" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15892
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15893
apply(simp add: abs_supp fin_supp abs_fresh fresh_left calc_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15894
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15895
apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15896
apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15897
apply(simp add: eqvts calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15898
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15899
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15900
lemma nrename_credu:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15901
  assumes a: "(M[x\<turnstile>n>y]) \<longrightarrow>\<^isub>c M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15902
  shows "\<exists>M0. M0[x\<turnstile>n>y]=M' \<and> M \<longrightarrow>\<^isub>c M0"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15903
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15904
apply(nominal_induct M\<equiv>"M[x\<turnstile>n>y]" M' avoiding: M x y rule: c_redu.strong_induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15905
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15906
apply(drule nrename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15907
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15908
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15909
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15910
apply(rule_tac x="M'{a:=(x).N'}" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15911
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15912
apply(simp add: fresh_atm abs_fresh subst_comm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15913
apply(rule c_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15914
apply(auto dest: not_fic_nrename)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15915
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15916
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15917
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15918
apply(drule nrename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15919
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15920
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15921
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15922
apply(rule_tac x="N'{x:=<a>.M'}" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15923
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15924
apply(simp add: fresh_atm abs_fresh subst_comm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15925
apply(rule c_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15926
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15927
apply(drule_tac x="xa" and y="y" in fin_nrename)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15928
apply(auto simp add: fresh_prod abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15929
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15930
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15931
lemma nrename_ax2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15932
  assumes a: "N[x\<turnstile>n>y] = Ax z c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15933
  shows "\<exists>z. N = Ax z c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15934
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15935
apply(nominal_induct N avoiding: x y rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15936
apply(auto split: if_splits)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15937
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15938
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15939
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15940
lemma fic_nrename:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15941
  assumes a: "fic (M[x\<turnstile>n>y]) c" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15942
  shows "fic M c" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15943
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15944
apply(nominal_induct M avoiding: c x y rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15945
apply(auto dest!: fic_elims intro!: fic.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15946
           split: if_splits)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15947
apply(auto dest: nrename_fresh_interesting2 simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15948
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15949
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15950
lemma nrename_lredu:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15951
  assumes a: "(M[x\<turnstile>n>y]) \<longrightarrow>\<^isub>l M'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15952
  shows "\<exists>M0. M0[x\<turnstile>n>y]=M' \<and> M \<longrightarrow>\<^isub>l M0"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15953
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15954
apply(nominal_induct M\<equiv>"M[x\<turnstile>n>y]" M' avoiding: M x y rule: l_redu.strong_induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15955
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15956
apply(drule nrename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15957
apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15958
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15959
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15960
apply(case_tac "xa=y")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15961
apply(simp add: nrename_id)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15962
apply(rule l_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15963
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15964
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15965
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15966
apply(frule nrename_ax2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15967
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15968
apply(case_tac "z=xa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15969
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15970
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15971
apply(rule_tac x="M'[a\<turnstile>c>b]" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15972
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15973
apply(rule crename_interesting3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15974
apply(rule l_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15975
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15976
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15977
apply(auto dest: fic_nrename simp add: fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15978
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15979
apply(drule nrename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15980
apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15981
apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15982
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15983
apply(case_tac "xa=ya")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15984
apply(simp add: nrename_id)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15985
apply(rule l_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15986
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15987
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15988
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15989
apply(frule nrename_ax2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15990
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15991
apply(case_tac "z=xa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15992
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15993
apply(rule_tac x="N'[x\<turnstile>n>xa]" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15994
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15995
apply(rule nrename_interesting1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15996
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15997
apply(rule l_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15998
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 15999
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16000
apply(auto dest: fin_nrename2 simp add: fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16001
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16002
apply(rule_tac x="N'[x\<turnstile>n>y]" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16003
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16004
apply(rule nrename_interesting2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16005
apply(simp_all)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16006
apply(rule l_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16007
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16008
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16009
apply(auto dest: fin_nrename2 simp add: fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16010
(* LNot *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16011
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16012
apply(drule nrename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16013
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16014
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16015
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16016
apply(drule nrename_NotR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16017
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16018
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16019
apply(drule nrename_NotL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16020
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16021
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16022
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16023
apply(rule_tac x="Cut <b>.N'b (x).N'a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16024
apply(simp add: fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16025
apply(rule l_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16026
apply(auto simp add: fresh_prod fresh_atm intro: nrename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16027
apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16028
apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting2)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16029
apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting2)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16030
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16031
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16032
(* LAnd1 *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16033
apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16034
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16035
apply(drule nrename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16036
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16037
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16038
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16039
apply(drule nrename_AndR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16040
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16041
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16042
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16043
apply(drule nrename_AndL1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16044
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16045
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16046
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16047
apply(rule_tac x="Cut <a1>.M'a (x).N'b" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16048
apply(simp add: fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16049
apply(rule l_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16050
apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: nrename_fresh_interesting2)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16051
apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16052
apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16053
apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16054
apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16055
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16056
(* LAnd2 *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16057
apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16058
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16059
apply(drule nrename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16060
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16061
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16062
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16063
apply(drule nrename_AndR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16064
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16065
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16066
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16067
apply(drule nrename_AndL2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16068
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16069
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16070
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16071
apply(rule_tac x="Cut <a2>.N'a (x).N'b" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16072
apply(simp add: fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16073
apply(rule l_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16074
apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: nrename_fresh_interesting2)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16075
apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16076
apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16077
apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16078
apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16079
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16080
(* LOr1 *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16081
apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16082
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16083
apply(drule nrename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16084
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16085
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16086
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16087
apply(drule nrename_OrL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16088
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16089
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16090
apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16091
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16092
apply(drule nrename_OrR1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16093
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16094
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16095
apply(rule_tac x="Cut <a>.N' (x1).M'a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16096
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16097
apply(simp add: abs_fresh fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16098
apply(rule l_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16099
apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: nrename_fresh_interesting2)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16100
apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16101
apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16102
apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16103
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16104
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16105
(* LOr2 *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16106
apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16107
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16108
apply(drule nrename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16109
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16110
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16111
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16112
apply(drule nrename_OrL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16113
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16114
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16115
apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16116
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16117
apply(drule nrename_OrR2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16118
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16119
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16120
apply(rule_tac x="Cut <a>.N' (x2).N'a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16121
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16122
apply(simp add: abs_fresh fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16123
apply(rule l_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16124
apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: nrename_fresh_interesting2)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16125
apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16126
apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16127
apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16128
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16129
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16130
(* ImpL *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16131
apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16132
apply(drule sym) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16133
apply(drule nrename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16134
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16135
apply(simp add: fresh_prod abs_fresh fresh_atm abs_supp fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16136
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16137
apply(drule nrename_ImpL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16138
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16139
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16140
apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16141
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16142
apply(drule nrename_ImpR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16143
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16144
apply(simp add: fresh_prod abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16145
apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16146
apply(rule_tac x="Cut <a>.(Cut <c>.M'a (x).N') (y).N'a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16147
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16148
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16149
apply(rule l_redu.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16150
apply(auto simp add: fresh_atm abs_fresh abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16151
apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting2)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16152
apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting1)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16153
apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting2)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16154
apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting2)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16155
apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting2)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16156
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16157
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16158
lemma nrename_aredu:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16159
  assumes a: "(M[x\<turnstile>n>y]) \<longrightarrow>\<^isub>a M'" "x\<noteq>y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16160
  shows "\<exists>M0. M0[x\<turnstile>n>y]=M' \<and> M \<longrightarrow>\<^isub>a M0"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16161
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16162
apply(nominal_induct M\<equiv>"M[x\<turnstile>n>y]" M' avoiding: M x y rule: a_redu.strong_induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16163
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16164
apply(drule  nrename_lredu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16165
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16166
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16167
apply(drule  nrename_credu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16168
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16169
(* Cut *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16170
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16171
apply(drule nrename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16172
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16173
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16174
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16175
apply(drule_tac x="M'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16176
apply(drule_tac x="xa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16177
apply(drule_tac x="y" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16178
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16179
apply(rule_tac x="Cut <a>.M0 (x).N'" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16180
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16181
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16182
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16183
apply(rule nrename.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16184
apply(drule nrename_fresh_interesting2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16185
apply(simp add: fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16186
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16187
apply(drule nrename_fresh_interesting1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16188
apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16189
apply(simp add: fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16190
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16191
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16192
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16193
apply(drule nrename_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16194
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16195
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16196
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16197
apply(drule_tac x="N'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16198
apply(drule_tac x="xa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16199
apply(drule_tac x="y" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16200
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16201
apply(rule_tac x="Cut <a>.M' (x).M0" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16202
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16203
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16204
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16205
apply(rule nrename.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16206
apply(simp add: fresh_a_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16207
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16208
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16209
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16210
(* NotL *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16211
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16212
apply(frule nrename_NotL_aux)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16213
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16214
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16215
apply(drule nrename_NotL')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16216
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16217
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16218
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16219
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16220
apply(drule_tac x="N'" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16221
apply(drule_tac x="xa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16222
apply(drule_tac x="y" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16223
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16224
apply(rule_tac x="NotL <a>.M0 x" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16225
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16226
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16227
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16228
apply(drule_tac x="N'" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16229
apply(drule_tac x="xa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16230
apply(drule_tac x="x" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16231
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16232
apply(rule_tac x="NotL <a>.M0 xa" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16233
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16234
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16235
(* NotR *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16236
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16237
apply(drule nrename_NotR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16238
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16239
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16240
apply(drule_tac x="N'" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16241
apply(drule_tac x="xa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16242
apply(drule_tac x="y" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16243
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16244
apply(rule_tac x="NotR (x).M0 a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16245
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16246
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16247
(* AndR *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16248
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16249
apply(drule nrename_AndR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16250
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16251
apply(auto simp add: fresh_atm fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16252
apply(auto simp add: fresh_atm fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16253
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16254
apply(drule_tac x="M'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16255
apply(drule_tac x="x" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16256
apply(drule_tac x="y" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16257
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16258
apply(rule_tac x="AndR <a>.M0 <b>.N' c" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16259
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16260
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16261
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16262
apply(rule nrename.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16263
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16264
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16265
apply(auto intro: fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16266
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16267
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16268
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16269
apply(drule nrename_AndR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16270
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16271
apply(auto simp add: fresh_atm fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16272
apply(auto simp add: fresh_atm fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16273
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16274
apply(drule_tac x="N'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16275
apply(drule_tac x="x" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16276
apply(drule_tac x="y" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16277
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16278
apply(rule_tac x="AndR <a>.M' <b>.M0 c" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16279
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16280
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16281
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16282
apply(rule nrename.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16283
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16284
apply(auto intro: fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16285
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16286
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16287
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16288
(* AndL1 *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16289
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16290
apply(frule nrename_AndL1_aux)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16291
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16292
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16293
apply(drule nrename_AndL1')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16294
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16295
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16296
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16297
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16298
apply(drule_tac x="N'" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16299
apply(drule_tac x="xa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16300
apply(drule_tac x="ya" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16301
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16302
apply(rule_tac x="AndL1 (x).M0 y" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16303
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16304
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16305
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16306
apply(drule_tac x="N'" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16307
apply(drule_tac x="xa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16308
apply(drule_tac x="y" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16309
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16310
apply(rule_tac x="AndL1 (x).M0 xa" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16311
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16312
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16313
(* AndL2 *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16314
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16315
apply(frule nrename_AndL2_aux)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16316
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16317
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16318
apply(drule nrename_AndL2')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16319
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16320
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16321
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16322
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16323
apply(drule_tac x="N'" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16324
apply(drule_tac x="xa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16325
apply(drule_tac x="ya" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16326
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16327
apply(rule_tac x="AndL2 (x).M0 y" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16328
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16329
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16330
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16331
apply(drule_tac x="N'" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16332
apply(drule_tac x="xa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16333
apply(drule_tac x="y" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16334
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16335
apply(rule_tac x="AndL2 (x).M0 xa" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16336
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16337
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16338
(* OrL *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16339
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16340
apply(frule nrename_OrL_aux)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16341
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16342
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16343
apply(drule nrename_OrL')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16344
apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16345
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16346
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16347
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16348
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16349
apply(drule_tac x="M'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16350
apply(drule_tac x="xa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16351
apply(drule_tac x="ya" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16352
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16353
apply(rule_tac x="OrL (x).M0 (y).N' z" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16354
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16355
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16356
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16357
apply(rule nrename.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16358
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16359
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16360
apply(auto intro: fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16361
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16362
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16363
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16364
apply(drule_tac x="M'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16365
apply(drule_tac x="xa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16366
apply(drule_tac x="z" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16367
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16368
apply(rule_tac x="OrL (x).M0 (y).N' xa" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16369
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16370
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16371
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16372
apply(rule nrename.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16373
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16374
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16375
apply(auto intro: fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16376
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16377
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16378
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16379
apply(frule nrename_OrL_aux)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16380
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16381
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16382
apply(drule nrename_OrL')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16383
apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16384
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16385
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16386
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16387
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16388
apply(drule_tac x="N'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16389
apply(drule_tac x="xa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16390
apply(drule_tac x="ya" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16391
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16392
apply(rule_tac x="OrL (x).M' (y).M0 z" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16393
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16394
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16395
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16396
apply(rule nrename.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16397
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16398
apply(auto intro: fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16399
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16400
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16401
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16402
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16403
apply(drule_tac x="N'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16404
apply(drule_tac x="xa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16405
apply(drule_tac x="z" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16406
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16407
apply(rule_tac x="OrL (x).M' (y).M0 xa" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16408
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16409
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16410
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16411
apply(rule nrename.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16412
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16413
apply(auto intro: fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16414
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16415
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16416
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16417
(* OrR1 *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16418
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16419
apply(drule nrename_OrR1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16420
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16421
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16422
apply(drule_tac x="N'" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16423
apply(drule_tac x="x" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16424
apply(drule_tac x="y" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16425
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16426
apply(rule_tac x="OrR1 <a>.M0 b" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16427
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16428
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16429
(* OrR2 *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16430
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16431
apply(drule nrename_OrR2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16432
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16433
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16434
apply(drule_tac x="N'" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16435
apply(drule_tac x="x" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16436
apply(drule_tac x="y" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16437
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16438
apply(rule_tac x="OrR2 <a>.M0 b" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16439
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16440
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16441
(* ImpL *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16442
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16443
apply(frule nrename_ImpL_aux)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16444
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16445
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16446
apply(drule nrename_ImpL')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16447
apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16448
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16449
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16450
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16451
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16452
apply(drule_tac x="M'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16453
apply(drule_tac x="xa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16454
apply(drule_tac x="ya" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16455
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16456
apply(rule_tac x="ImpL <a>.M0 (x).N' y" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16457
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16458
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16459
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16460
apply(rule nrename.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16461
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16462
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16463
apply(auto intro: fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16464
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16465
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16466
apply(drule_tac x="M'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16467
apply(drule_tac x="xa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16468
apply(drule_tac x="y" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16469
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16470
apply(rule_tac x="ImpL <a>.M0 (x).N' xa" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16471
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16472
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16473
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16474
apply(rule nrename.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16475
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16476
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16477
apply(auto intro: fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16478
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16479
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16480
apply(frule nrename_ImpL_aux)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16481
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16482
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16483
apply(drule nrename_ImpL')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16484
apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16485
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16486
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16487
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16488
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16489
apply(drule_tac x="N'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16490
apply(drule_tac x="xa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16491
apply(drule_tac x="ya" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16492
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16493
apply(rule_tac x="ImpL <a>.M' (x).M0 y" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16494
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16495
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16496
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16497
apply(rule nrename.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16498
apply(auto intro: fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16499
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16500
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16501
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16502
apply(drule_tac x="N'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16503
apply(drule_tac x="xa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16504
apply(drule_tac x="y" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16505
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16506
apply(rule_tac x="ImpL <a>.M' (x).M0 xa" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16507
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16508
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16509
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16510
apply(rule nrename.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16511
apply(auto intro: fresh_a_redu)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16512
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16513
apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16514
(* ImpR *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16515
apply(drule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16516
apply(drule nrename_ImpR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16517
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16518
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16519
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16520
apply(drule_tac x="N'" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16521
apply(drule_tac x="xa" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16522
apply(drule_tac x="y" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16523
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16524
apply(rule_tac x="ImpR (x).<a>.M0 b" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16525
apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16526
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16527
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16528
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16529
lemma SNa_preserved_renaming2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16530
  assumes a: "SNa N"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16531
  shows "SNa (N[x\<turnstile>n>y])"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16532
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16533
apply(induct rule: SNa_induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16534
apply(case_tac "x=y")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16535
apply(simp add: nrename_id)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16536
apply(rule SNaI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16537
apply(drule nrename_aredu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16538
apply(blast)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16539
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16540
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16541
text {* helper-stuff to set up the induction *}
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16542
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16543
abbreviation
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16544
  SNa_set :: "trm set"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16545
where
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16546
  "SNa_set \<equiv> {M. SNa M}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16547
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16548
abbreviation
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16549
  A_Redu_set :: "(trm\<times>trm) set"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16550
where
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16551
 "A_Redu_set \<equiv> {(N,M)| M N. M \<longrightarrow>\<^isub>a N}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16552
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16553
lemma SNa_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16554
  assumes a: "SNa M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16555
  shows "(\<forall>M. (\<forall>N. M \<longrightarrow>\<^isub>a N \<longrightarrow> P N)\<longrightarrow> P M) \<longrightarrow> P M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16556
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16557
by (induct rule: SNa.induct) (blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16558
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16559
lemma wf_SNa_restricted:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16560
  shows "wf (A_Redu_set \<inter> (UNIV <*> SNa_set))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16561
apply(unfold wf_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16562
apply(intro strip)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16563
apply(case_tac "SNa x")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16564
apply(simp (no_asm_use))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16565
apply(drule_tac P="P" in SNa_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16566
apply(erule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16567
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16568
(* other case *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16569
apply(drule_tac x="x" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16570
apply(erule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16571
apply(fast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16572
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16573
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16574
constdefs
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16575
  SNa_Redu :: "(trm \<times> trm) set"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16576
  "SNa_Redu \<equiv> A_Redu_set \<inter> (UNIV <*> SNa_set)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16577
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16578
lemma wf_SNa_Redu:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16579
  shows "wf SNa_Redu"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16580
apply(unfold SNa_Redu_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16581
apply(rule wf_SNa_restricted)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16582
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16583
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16584
lemma wf_measure_triple:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16585
shows "wf ((measure size) <*lex*> SNa_Redu <*lex*> SNa_Redu)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16586
by (auto intro: wf_SNa_Redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16587
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16588
lemma my_wf_induct_triple: 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16589
 assumes a: " wf(r1 <*lex*> r2 <*lex*> r3)"           
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16590
 and     b: "\<And>x. \<lbrakk>\<And>y. ((fst y,fst (snd y),snd (snd y)),(fst x,fst (snd x), snd (snd x))) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16591
                                    \<in> (r1 <*lex*> r2 <*lex*> r3) \<longrightarrow> P y\<rbrakk> \<Longrightarrow> P x"  
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16592
 shows "P x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16593
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16594
apply(induct x rule: wf_induct_rule)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16595
apply(rule b)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16596
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16597
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16598
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16599
lemma my_wf_induct_triple': 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16600
 assumes a: " wf(r1 <*lex*> r2 <*lex*> r3)"           
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16601
 and    b: "\<And>x1 x2 x3. \<lbrakk>\<And>y1 y2 y3. ((y1,y2,y3),(x1,x2,x3)) \<in> (r1 <*lex*> r2 <*lex*> r3) \<longrightarrow> P (y1,y2,y3)\<rbrakk> 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16602
             \<Longrightarrow> P (x1,x2,x3)"  
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16603
 shows "P (x1,x2,x3)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16604
apply(rule_tac my_wf_induct_triple[OF a])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16605
apply(case_tac x)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16606
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16607
apply(case_tac b)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16608
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16609
apply(rule b)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16610
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16611
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16612
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16613
lemma my_wf_induct_triple'': 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16614
 assumes a: " wf(r1 <*lex*> r2 <*lex*> r3)"           
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16615
 and     b: "\<And>x1 x2 x3. \<lbrakk>\<And>y1 y2 y3. ((y1,y2,y3),(x1,x2,x3)) \<in> (r1 <*lex*> r2 <*lex*> r3) \<longrightarrow> P y1 y2 y3\<rbrakk>
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16616
               \<Longrightarrow> P x1 x2 x3"  
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16617
 shows "P x1 x2 x3"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16618
apply(rule_tac my_wf_induct_triple'[where P="\<lambda>(x1,x2,x3). P x1 x2 x3", simplified])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16619
apply(rule a)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16620
apply(rule b)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16621
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16622
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16623
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16624
lemma excluded_m:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16625
  assumes a: "<a>:M \<in> (\<parallel><B>\<parallel>)" "(x):N \<in> (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16626
  shows "(<a>:M \<in> BINDINGc B (\<parallel>(B)\<parallel>) \<or> (x):N \<in> BINDINGn B (\<parallel><B>\<parallel>))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16627
      \<or>\<not>(<a>:M \<in> BINDINGc B (\<parallel>(B)\<parallel>) \<or> (x):N \<in> BINDINGn B (\<parallel><B>\<parallel>))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16628
by (blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16629
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16630
lemma tricky_subst:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16631
  assumes a1: "b\<sharp>(c,N)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16632
  and     a2: "z\<sharp>(x,P)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16633
  and     a3: "M\<noteq>Ax z b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16634
  shows "(Cut <c>.N (z).M){b:=(x).P} = Cut <c>.N (z).(M{b:=(x).P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16635
using a1 a2 a3
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16636
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16637
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16638
apply(subgoal_tac "Cut <c>.N (z).M = Cut <ca>.([(ca,c)]\<bullet>N) (z).M")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16639
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16640
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16641
apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16642
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16643
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16644
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16645
apply(subgoal_tac "b\<sharp>([(ca,c)]\<bullet>N)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16646
apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16647
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16648
apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16649
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16650
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16651
apply(simp add: alpha fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16652
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16653
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16654
text {* 3rd lemma *}
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16655
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16656
lemma CUT_SNa_aux:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16657
  assumes a1: "<a>:M \<in> (\<parallel><B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16658
  and     a2: "SNa M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16659
  and     a3: "(x):N \<in> (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16660
  and     a4: "SNa N"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16661
  shows   "SNa (Cut <a>.M (x).N)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16662
using a1 a2 a3 a4
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16663
apply(induct B M N arbitrary: a x rule: my_wf_induct_triple''[OF wf_measure_triple])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16664
apply(rule SNaI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16665
apply(drule Cut_a_redu_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16666
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16667
(* left-inner reduction *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16668
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16669
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16670
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16671
apply(drule_tac x="x1" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16672
apply(drule_tac x="M'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16673
apply(drule_tac x="x3" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16674
apply(drule conjunct2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16675
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16676
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16677
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16678
apply(rule disjI1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16679
apply(simp add: SNa_Redu_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16680
apply(drule_tac x="a" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16681
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16682
apply(simp add: CANDs_preserved_single)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16683
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16684
apply(simp add: a_preserves_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16685
apply(drule_tac x="x" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16686
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16687
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16688
(* right-inner reduction *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16689
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16690
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16691
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16692
apply(drule_tac x="x1" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16693
apply(drule_tac x="x2" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16694
apply(drule_tac x="N'" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16695
apply(drule conjunct2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16696
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16697
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16698
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16699
apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16700
apply(simp add: SNa_Redu_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16701
apply(drule_tac x="a" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16702
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16703
apply(simp add: CANDs_preserved_single)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16704
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16705
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16706
apply(drule_tac x="x" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16707
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16708
apply(simp add: CANDs_preserved_single)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16709
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16710
apply(simp add: a_preserves_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16711
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16712
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16713
(******** c-reduction *********)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16714
apply(drule Cut_c_redu_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16715
(* c-left reduction*)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16716
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16717
apply(erule conjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16718
apply(frule_tac B="x1" in fic_CANDS)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16719
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16720
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16721
(* in AXIOMSc *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16722
apply(simp add: AXIOMSc_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16723
apply(erule exE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16724
apply(simp add: ctrm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16725
apply(simp add: alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16726
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16727
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16728
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16729
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16730
apply(subgoal_tac "fic (Ax y b) b")(*A*)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16731
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16732
(*A*)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16733
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16734
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16735
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16736
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16737
apply(subgoal_tac "fic (Ax ([(a,aa)]\<bullet>y) a) a")(*B*)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16738
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16739
(*B*)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16740
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16741
(* in BINDINGc *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16742
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16743
apply(drule BINDINGc_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16744
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16745
(* c-right reduction*)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16746
apply(erule conjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16747
apply(frule_tac B="x1" in fin_CANDS)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16748
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16749
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16750
(* in AXIOMSc *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16751
apply(simp add: AXIOMSn_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16752
apply(erule exE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16753
apply(simp add: ntrm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16754
apply(simp add: alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16755
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16756
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16757
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16758
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16759
apply(subgoal_tac "fin (Ax xa b) xa")(*A*)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16760
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16761
(*A*)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16762
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16763
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16764
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16765
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16766
apply(subgoal_tac "fin (Ax x ([(x,xa)]\<bullet>b)) x")(*B*)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16767
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16768
(*B*)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16769
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16770
(* in BINDINGc *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16771
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16772
apply(drule BINDINGn_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16773
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16774
(*********** l-reductions ************)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16775
apply(drule Cut_l_redu_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16776
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16777
(* ax1 *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16778
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16779
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16780
apply(simp add: SNa_preserved_renaming1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16781
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16782
(* ax2 *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16783
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16784
apply(simp add: SNa_preserved_renaming2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16785
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16786
(* LNot *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16787
apply(erule exE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16788
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16789
apply(frule_tac excluded_m)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16790
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16791
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16792
(* one of them in BINDING *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16793
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16794
apply(drule fin_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16795
apply(drule fic_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16796
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16797
apply(drule BINDINGc_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16798
apply(drule_tac x="x" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16799
apply(drule_tac x="NotL <b>.N' x" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16800
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16801
apply(simp add: better_NotR_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16802
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16803
apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.NotR (y).M'a a' (x).NotL <b>.N' x) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16804
                   =  Cut <c>.NotR (y).M'a c (x).NotL <b>.N' x")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16805
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16806
apply(subgoal_tac "Cut <c>.NotR (y).M'a c (x).NotL <b>.N' x \<longrightarrow>\<^isub>a Cut <b>.N' (y).M'a")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16807
apply(simp only: a_preserves_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16808
apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16809
apply(rule better_LNot_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16810
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16811
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16812
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16813
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16814
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16815
(* other case of in BINDING *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16816
apply(drule fin_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16817
apply(drule fic_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16818
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16819
apply(drule BINDINGn_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16820
apply(drule_tac x="a" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16821
apply(drule_tac x="NotR (y).M'a a" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16822
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16823
apply(simp add: better_NotL_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16824
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16825
apply(subgoal_tac "fresh_fun (\<lambda>x'. Cut <a>.NotR (y).M'a a (x').NotL <b>.N' x') 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16826
                   = Cut <a>.NotR (y).M'a a (c).NotL <b>.N' c")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16827
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16828
apply(subgoal_tac "Cut <a>.NotR (y).M'a a (c).NotL <b>.N' c \<longrightarrow>\<^isub>a Cut <b>.N' (y).M'a")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16829
apply(simp only: a_preserves_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16830
apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16831
apply(rule better_LNot_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16832
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16833
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16834
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16835
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16836
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16837
(* none of them in BINDING *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16838
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16839
apply(erule conjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16840
apply(frule CAND_NotR_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16841
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16842
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16843
apply(frule CAND_NotL_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16844
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16845
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16846
apply(simp only: ty.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16847
apply(drule_tac x="B'" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16848
apply(drule_tac x="N'" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16849
apply(drule_tac x="M'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16850
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16851
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16852
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16853
apply(drule_tac x="b" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16854
apply(rotate_tac 13)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16855
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16856
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16857
apply(rotate_tac 13)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16858
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16859
apply(simp add: CANDs_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16860
apply(drule_tac x="y" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16861
apply(rotate_tac 13)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16862
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16863
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16864
apply(rotate_tac 13)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16865
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16866
apply(simp add: CANDs_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16867
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16868
(* LAnd1 case *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16869
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16870
apply(erule exE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16871
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16872
apply(frule_tac excluded_m)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16873
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16874
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16875
(* one of them in BINDING *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16876
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16877
apply(drule fin_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16878
apply(drule fic_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16879
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16880
apply(drule BINDINGc_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16881
apply(drule_tac x="x" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16882
apply(drule_tac x="AndL1 (y).N' x" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16883
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16884
apply(simp add: better_AndR_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16885
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16886
apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.AndR <b>.M1 <c>.M2 a' (x).AndL1 (y).N' x) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16887
                   = Cut <ca>.AndR <b>.M1 <c>.M2 ca (x).AndL1 (y).N' x")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16888
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16889
apply(subgoal_tac "Cut <ca>.AndR <b>.M1 <c>.M2 ca (x).AndL1 (y).N' x \<longrightarrow>\<^isub>a Cut <b>.M1 (y).N'")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16890
apply(auto intro: a_preserves_SNa)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16891
apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16892
apply(rule better_LAnd1_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16893
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16894
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16895
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16896
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16897
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16898
(* other case of in BINDING *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16899
apply(drule fin_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16900
apply(drule fic_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16901
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16902
apply(drule BINDINGn_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16903
apply(drule_tac x="a" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16904
apply(drule_tac x="AndR <b>.M1 <c>.M2 a" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16905
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16906
apply(simp add: better_AndL1_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16907
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16908
apply(subgoal_tac "fresh_fun (\<lambda>z'. Cut <a>.AndR <b>.M1 <c>.M2 a (z').AndL1 (y).N' z') 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16909
                   = Cut <a>.AndR <b>.M1 <c>.M2 a (ca).AndL1 (y).N' ca")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16910
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16911
apply(subgoal_tac "Cut <a>.AndR <b>.M1 <c>.M2 a (ca).AndL1 (y).N' ca \<longrightarrow>\<^isub>a Cut <b>.M1 (y).N'")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16912
apply(auto intro: a_preserves_SNa)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16913
apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16914
apply(rule better_LAnd1_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16915
apply(simp add: abs_fresh fresh_prod fresh_atm) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16916
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16917
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16918
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16919
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16920
(* none of them in BINDING *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16921
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16922
apply(erule conjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16923
apply(frule CAND_AndR_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16924
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16925
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16926
apply(frule CAND_AndL1_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16927
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16928
apply(erule exE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16929
apply(simp only: ty.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16930
apply(drule_tac x="B1" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16931
apply(drule_tac x="M1" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16932
apply(drule_tac x="N'" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16933
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16934
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16935
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16936
apply(drule_tac x="b" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16937
apply(rotate_tac 14)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16938
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16939
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16940
apply(rotate_tac 14)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16941
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16942
apply(simp add: CANDs_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16943
apply(drule_tac x="y" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16944
apply(rotate_tac 15)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16945
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16946
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16947
apply(rotate_tac 15)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16948
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16949
apply(simp add: CANDs_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16950
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16951
(* LAnd2 case *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16952
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16953
apply(erule exE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16954
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16955
apply(frule_tac excluded_m)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16956
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16957
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16958
(* one of them in BINDING *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16959
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16960
apply(drule fin_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16961
apply(drule fic_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16962
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16963
apply(drule BINDINGc_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16964
apply(drule_tac x="x" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16965
apply(drule_tac x="AndL2 (y).N' x" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16966
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16967
apply(simp add: better_AndR_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16968
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16969
apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.AndR <b>.M1 <c>.M2 a' (x).AndL2 (y).N' x) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16970
                   = Cut <ca>.AndR <b>.M1 <c>.M2 ca (x).AndL2 (y).N' x")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16971
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16972
apply(subgoal_tac "Cut <ca>.AndR <b>.M1 <c>.M2 ca (x).AndL2 (y).N' x \<longrightarrow>\<^isub>a Cut <c>.M2 (y).N'")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16973
apply(auto intro: a_preserves_SNa)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16974
apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16975
apply(rule better_LAnd2_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16976
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16977
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16978
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16979
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16980
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16981
(* other case of in BINDING *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16982
apply(drule fin_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16983
apply(drule fic_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16984
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16985
apply(drule BINDINGn_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16986
apply(drule_tac x="a" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16987
apply(drule_tac x="AndR <b>.M1 <c>.M2 a" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16988
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16989
apply(simp add: better_AndL2_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16990
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16991
apply(subgoal_tac "fresh_fun (\<lambda>z'. Cut <a>.AndR <b>.M1 <c>.M2 a (z').AndL2 (y).N' z') 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16992
                   = Cut <a>.AndR <b>.M1 <c>.M2 a (ca).AndL2 (y).N' ca")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16993
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16994
apply(subgoal_tac "Cut <a>.AndR <b>.M1 <c>.M2 a (ca).AndL2 (y).N' ca \<longrightarrow>\<^isub>a Cut <c>.M2 (y).N'")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16995
apply(auto intro: a_preserves_SNa)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16996
apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16997
apply(rule better_LAnd2_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16998
apply(simp add: abs_fresh fresh_prod fresh_atm) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 16999
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17000
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17001
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17002
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17003
(* none of them in BINDING *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17004
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17005
apply(erule conjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17006
apply(frule CAND_AndR_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17007
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17008
apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17009
apply(frule CAND_AndL2_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17010
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17011
apply(erule exE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17012
apply(simp only: ty.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17013
apply(drule_tac x="B2" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17014
apply(drule_tac x="M2" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17015
apply(drule_tac x="N'" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17016
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17017
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17018
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17019
apply(drule_tac x="c" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17020
apply(rotate_tac 14)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17021
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17022
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17023
apply(rotate_tac 14)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17024
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17025
apply(simp add: CANDs_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17026
apply(drule_tac x="y" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17027
apply(rotate_tac 15)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17028
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17029
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17030
apply(rotate_tac 15)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17031
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17032
apply(simp add: CANDs_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17033
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17034
(* LOr1 case *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17035
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17036
apply(erule exE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17037
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17038
apply(frule_tac excluded_m)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17039
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17040
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17041
(* one of them in BINDING *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17042
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17043
apply(drule fin_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17044
apply(drule fic_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17045
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17046
apply(drule BINDINGc_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17047
apply(drule_tac x="x" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17048
apply(drule_tac x="OrL (z).M1 (y).M2 x" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17049
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17050
apply(simp add: better_OrR1_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17051
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17052
apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.OrR1 <b>.N' a' (x).OrL (z).M1 (y).M2 x) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17053
                   = Cut <c>.OrR1 <b>.N' c (x).OrL (z).M1 (y).M2 x")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17054
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17055
apply(subgoal_tac "Cut <c>.OrR1 <b>.N' c (x).OrL (z).M1 (y).M2 x \<longrightarrow>\<^isub>a Cut <b>.N' (z).M1")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17056
apply(auto intro: a_preserves_SNa)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17057
apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17058
apply(rule better_LOr1_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17059
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17060
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17061
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17062
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17063
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17064
(* other case of in BINDING *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17065
apply(drule fin_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17066
apply(drule fic_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17067
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17068
apply(drule BINDINGn_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17069
apply(drule_tac x="a" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17070
apply(drule_tac x="OrR1 <b>.N' a" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17071
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17072
apply(simp add: better_OrL_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17073
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17074
apply(subgoal_tac "fresh_fun (\<lambda>z'. Cut <a>.OrR1 <b>.N' a (z').OrL (z).M1 (y).M2 z') 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17075
                   = Cut <a>.OrR1 <b>.N' a (c).OrL (z).M1 (y).M2 c")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17076
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17077
apply(subgoal_tac "Cut <a>.OrR1 <b>.N' a (c).OrL (z).M1 (y).M2 c \<longrightarrow>\<^isub>a Cut <b>.N' (z).M1")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17078
apply(auto intro: a_preserves_SNa)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17079
apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17080
apply(rule better_LOr1_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17081
apply(simp add: abs_fresh fresh_prod fresh_atm) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17082
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17083
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17084
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17085
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17086
(* none of them in BINDING *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17087
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17088
apply(erule conjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17089
apply(frule CAND_OrR1_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17090
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17091
apply(erule exE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17092
apply(frule CAND_OrL_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17093
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17094
apply(erule exE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17095
apply(simp only: ty.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17096
apply(drule_tac x="B1" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17097
apply(drule_tac x="N'" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17098
apply(drule_tac x="M1" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17099
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17100
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17101
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17102
apply(drule_tac x="b" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17103
apply(rotate_tac 15)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17104
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17105
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17106
apply(rotate_tac 15)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17107
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17108
apply(simp add: CANDs_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17109
apply(drule_tac x="z" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17110
apply(rotate_tac 15)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17111
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17112
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17113
apply(rotate_tac 15)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17114
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17115
apply(simp add: CANDs_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17116
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17117
(* LOr2 case *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17118
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17119
apply(erule exE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17120
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17121
apply(frule_tac excluded_m)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17122
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17123
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17124
(* one of them in BINDING *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17125
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17126
apply(drule fin_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17127
apply(drule fic_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17128
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17129
apply(drule BINDINGc_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17130
apply(drule_tac x="x" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17131
apply(drule_tac x="OrL (z).M1 (y).M2 x" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17132
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17133
apply(simp add: better_OrR2_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17134
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17135
apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.OrR2 <b>.N' a' (x).OrL (z).M1 (y).M2 x) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17136
                   = Cut <c>.OrR2 <b>.N' c (x).OrL (z).M1 (y).M2 x")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17137
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17138
apply(subgoal_tac "Cut <c>.OrR2 <b>.N' c (x).OrL (z).M1 (y).M2 x \<longrightarrow>\<^isub>a Cut <b>.N' (y).M2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17139
apply(auto intro: a_preserves_SNa)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17140
apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17141
apply(rule better_LOr2_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17142
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17143
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17144
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17145
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17146
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17147
(* other case of in BINDING *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17148
apply(drule fin_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17149
apply(drule fic_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17150
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17151
apply(drule BINDINGn_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17152
apply(drule_tac x="a" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17153
apply(drule_tac x="OrR2 <b>.N' a" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17154
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17155
apply(simp add: better_OrL_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17156
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17157
apply(subgoal_tac "fresh_fun (\<lambda>z'. Cut <a>.OrR2 <b>.N' a (z').OrL (z).M1 (y).M2 z') 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17158
                   = Cut <a>.OrR2 <b>.N' a (c).OrL (z).M1 (y).M2 c")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17159
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17160
apply(subgoal_tac "Cut <a>.OrR2 <b>.N' a (c).OrL (z).M1 (y).M2 c \<longrightarrow>\<^isub>a Cut <b>.N' (y).M2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17161
apply(auto intro: a_preserves_SNa)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17162
apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17163
apply(rule better_LOr2_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17164
apply(simp add: abs_fresh fresh_prod fresh_atm) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17165
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17166
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17167
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17168
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17169
(* none of them in BINDING *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17170
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17171
apply(erule conjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17172
apply(frule CAND_OrR2_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17173
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17174
apply(erule exE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17175
apply(frule CAND_OrL_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17176
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17177
apply(erule exE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17178
apply(simp only: ty.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17179
apply(drule_tac x="B2" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17180
apply(drule_tac x="N'" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17181
apply(drule_tac x="M2" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17182
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17183
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17184
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17185
apply(drule_tac x="b" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17186
apply(rotate_tac 15)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17187
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17188
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17189
apply(rotate_tac 15)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17190
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17191
apply(simp add: CANDs_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17192
apply(drule_tac x="y" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17193
apply(rotate_tac 15)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17194
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17195
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17196
apply(rotate_tac 15)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17197
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17198
apply(simp add: CANDs_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17199
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17200
(* LImp case *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17201
apply(erule exE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17202
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17203
apply(frule_tac excluded_m)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17204
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17205
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17206
(* one of them in BINDING *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17207
apply(erule disjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17208
apply(drule fin_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17209
apply(drule fic_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17210
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17211
apply(drule BINDINGc_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17212
apply(drule_tac x="x" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17213
apply(drule_tac x="ImpL <c>.N1 (y).N2 x" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17214
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17215
apply(simp add: better_ImpR_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17216
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17217
apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.ImpR (z).<b>.M'a a' (x).ImpL <c>.N1 (y).N2 x)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17218
                   = Cut <ca>.ImpR (z).<b>.M'a ca (x).ImpL <c>.N1 (y).N2 x")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17219
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17220
apply(subgoal_tac "Cut <ca>.ImpR (z).<b>.M'a ca (x).ImpL <c>.N1 (y).N2 x \<longrightarrow>\<^isub>a 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17221
                                                          Cut <b>.Cut <c>.N1 (z).M'a (y).N2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17222
apply(auto intro: a_preserves_SNa)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17223
apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17224
apply(rule better_LImp_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17225
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17226
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17227
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17228
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17229
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17230
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17231
(* other case of in BINDING *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17232
apply(drule fin_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17233
apply(drule fic_elims)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17234
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17235
apply(drule BINDINGn_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17236
apply(drule_tac x="a" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17237
apply(drule_tac x="ImpR (z).<b>.M'a a" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17238
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17239
apply(simp add: better_ImpL_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17240
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17241
apply(subgoal_tac "fresh_fun (\<lambda>z'. Cut <a>.ImpR (z).<b>.M'a a (z').ImpL <c>.N1 (y).N2 z')
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17242
                   = Cut <a>.ImpR (z).<b>.M'a a (ca).ImpL <c>.N1 (y).N2 ca")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17243
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17244
apply(subgoal_tac "Cut <a>.ImpR (z).<b>.M'a a (ca).ImpL <c>.N1 (y).N2 ca \<longrightarrow>\<^isub>a 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17245
                                                          Cut <b>.Cut <c>.N1 (z).M'a (y).N2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17246
apply(auto intro: a_preserves_SNa)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17247
apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17248
apply(rule better_LImp_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17249
apply(simp add: abs_fresh fresh_prod fresh_atm) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17250
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17251
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17252
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17253
apply(simp add: abs_fresh abs_supp fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17254
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17255
apply(simp add: abs_fresh abs_supp fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17256
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17257
(* none of them in BINDING *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17258
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17259
apply(erule conjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17260
apply(frule CAND_ImpL_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17261
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17262
apply(erule exE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17263
apply(frule CAND_ImpR_elim) (* check here *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17264
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17265
apply(erule exE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17266
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17267
apply(simp only: ty.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17268
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17269
apply(case_tac "M'a=Ax z b")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17270
(* case Ma = Ax z b *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17271
apply(rule_tac t="Cut <b>.(Cut <c>.N1 (z).M'a) (y).N2" and s="Cut <b>.(M'a{z:=<c>.N1}) (y).N2" in subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17272
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17273
apply(drule_tac x="c" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17274
apply(drule_tac x="N1" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17275
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17276
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17277
apply(drule_tac x="B2" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17278
apply(drule_tac x="M'a{z:=<c>.N1}" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17279
apply(drule_tac x="N2" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17280
apply(drule conjunct1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17281
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17282
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17283
apply(rotate_tac 17)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17284
apply(drule_tac x="b" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17285
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17286
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17287
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17288
apply(simp add: CANDs_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17289
apply(rotate_tac 17)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17290
apply(drule_tac x="y" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17291
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17292
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17293
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17294
apply(simp add: CANDs_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17295
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17296
(* case Ma \<noteq> Ax z b *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17297
apply(subgoal_tac "<b>:Cut <c>.N1 (z).M'a \<in> \<parallel><B2>\<parallel>") (* lemma *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17298
apply(frule_tac meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17299
apply(drule_tac x="B2" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17300
apply(drule_tac x="Cut <c>.N1 (z).M'a" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17301
apply(drule_tac x="N2" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17302
apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17303
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17304
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17305
apply(rotate_tac 20)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17306
apply(drule_tac x="b" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17307
apply(rotate_tac 20)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17308
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17309
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17310
apply(rotate_tac 20)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17311
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17312
apply(simp add: CANDs_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17313
apply(rotate_tac 20)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17314
apply(drule_tac x="y" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17315
apply(rotate_tac 20)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17316
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17317
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17318
apply(rotate_tac 20)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17319
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17320
apply(simp add: CANDs_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17321
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17322
(* lemma *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17323
apply(subgoal_tac "<b>:Cut <c>.N1 (z).M'a \<in> BINDINGc B2 (\<parallel>(B2)\<parallel>)") (* second lemma *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17324
apply(simp add: BINDING_implies_CAND)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17325
(* second lemma *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17326
apply(simp (no_asm) add: BINDINGc_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17327
apply(rule exI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17328
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17329
apply(rule refl)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17330
apply(rule allI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17331
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17332
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17333
apply(rule_tac t="Cut <c>.N1 (z).M'a" and s="Cut <c>.N1 (ca).([(ca,z)]\<bullet>M'a)" in subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17334
apply(simp add: trm.inject alpha fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17335
apply(rule_tac t="(Cut <c>.N1 (ca).([(ca,z)]\<bullet>M'a)){b:=(xa).P}" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17336
           and s="Cut <c>.N1 (ca).(([(ca,z)]\<bullet>M'a){b:=(xa).P})" in subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17337
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17338
apply(rule tricky_subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17339
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17340
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17341
apply(clarify)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17342
apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17343
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17344
apply(drule_tac x="B1" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17345
apply(drule_tac x="N1" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17346
apply(drule_tac x="([(ca,z)]\<bullet>M'a){b:=(xa).P}" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17347
apply(drule conjunct1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17348
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17349
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17350
apply(rotate_tac 19)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17351
apply(drule_tac x="c" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17352
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17353
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17354
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17355
apply(simp add: CANDs_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17356
apply(rotate_tac 19)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17357
apply(drule_tac x="ca" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17358
apply(subgoal_tac "(ca):([(ca,z)]\<bullet>M'a){b:=(xa).P} \<in> \<parallel>(B1)\<parallel>")(*A*)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17359
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17360
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17361
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17362
apply(simp add: CANDs_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17363
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17364
(*A*)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17365
apply(drule_tac x="[(ca,z)]\<bullet>xa" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17366
apply(drule_tac x="[(ca,z)]\<bullet>P" in spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17367
apply(rotate_tac 19)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17368
apply(simp add: fresh_prod fresh_left)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17369
apply(drule mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17370
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17371
apply(auto simp add: calc_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17372
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17373
apply(auto simp add: calc_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17374
apply(drule_tac pi="[(ca,z)]" and x="(xa):P" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17375
apply(simp add: CAND_eqvt_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17376
apply(drule_tac pi="[(ca,z)]" and X="\<parallel>(B1)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17377
apply(simp add: CAND_eqvt_name csubst_eqvt)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17378
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17379
apply(simp add: calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17380
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17381
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17382
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17383
(* HERE *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17384
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17385
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17386
(* parallel substitution *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17387
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17388
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17389
lemma CUT_SNa:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17390
  assumes a1: "<a>:M \<in> (\<parallel><B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17391
  and     a2: "(x):N \<in> (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17392
  shows   "SNa (Cut <a>.M (x).N)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17393
using a1 a2
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17394
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17395
apply(rule CUT_SNa_aux[OF a1])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17396
apply(simp_all add: CANDs_imply_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17397
done 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17398
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17399
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17400
fun 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17401
 findn :: "(name\<times>coname\<times>trm) list\<Rightarrow>name\<Rightarrow>(coname\<times>trm) option"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17402
where
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17403
  "findn [] x = None"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17404
| "findn ((y,c,P)#\<theta>n) x = (if y=x then Some (c,P) else findn \<theta>n x)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17405
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17406
lemma findn_eqvt[eqvt]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17407
  fixes pi1::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17408
  and   pi2::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17409
  shows "(pi1\<bullet>findn \<theta>n x) = findn (pi1\<bullet>\<theta>n) (pi1\<bullet>x)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17410
  and   "(pi2\<bullet>findn \<theta>n x) = findn (pi2\<bullet>\<theta>n) (pi2\<bullet>x)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17411
apply(induct \<theta>n)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17412
apply(auto simp add: perm_bij) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17413
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17414
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17415
lemma findn_fresh:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17416
  assumes a: "x\<sharp>\<theta>n"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17417
  shows "findn \<theta>n x = None"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17418
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17419
apply(induct \<theta>n)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17420
apply(auto simp add: fresh_list_cons fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17421
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17422
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17423
fun 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17424
 findc :: "(coname\<times>name\<times>trm) list\<Rightarrow>coname\<Rightarrow>(name\<times>trm) option"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17425
where
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17426
  "findc [] x = None"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17427
| "findc ((c,y,P)#\<theta>c) a = (if a=c then Some (y,P) else findc \<theta>c a)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17428
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17429
lemma findc_eqvt[eqvt]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17430
  fixes pi1::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17431
  and   pi2::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17432
  shows "(pi1\<bullet>findc \<theta>c a) = findc (pi1\<bullet>\<theta>c) (pi1\<bullet>a)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17433
  and   "(pi2\<bullet>findc \<theta>c a) = findc (pi2\<bullet>\<theta>c) (pi2\<bullet>a)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17434
apply(induct \<theta>c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17435
apply(auto simp add: perm_bij) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17436
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17437
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17438
lemma findc_fresh:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17439
  assumes a: "a\<sharp>\<theta>c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17440
  shows "findc \<theta>c a = None"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17441
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17442
apply(induct \<theta>c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17443
apply(auto simp add: fresh_list_cons fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17444
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17445
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17446
abbreviation 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17447
 nmaps :: "(name\<times>coname\<times>trm) list \<Rightarrow> name \<Rightarrow> (coname\<times>trm) option \<Rightarrow> bool" ("_ nmaps _ to _" [55,55,55] 55) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17448
where
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17449
 "\<theta>n nmaps x to P \<equiv> (findn \<theta>n x) = P"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17450
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17451
abbreviation 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17452
 cmaps :: "(coname\<times>name\<times>trm) list \<Rightarrow> coname \<Rightarrow> (name\<times>trm) option \<Rightarrow> bool" ("_ cmaps _ to _" [55,55,55] 55) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17453
where
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17454
 "\<theta>c cmaps a to P \<equiv> (findc \<theta>c a) = P"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17455
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17456
lemma nmaps_fresh:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17457
  shows "\<theta>n nmaps x to Some (c,P) \<Longrightarrow> a\<sharp>\<theta>n \<Longrightarrow> a\<sharp>(c,P)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17458
apply(induct \<theta>n)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17459
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17460
apply(case_tac "aa=x")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17461
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17462
apply(case_tac "aa=x")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17463
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17464
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17465
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17466
lemma cmaps_fresh:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17467
  shows "\<theta>c cmaps a to Some (y,P) \<Longrightarrow> x\<sharp>\<theta>c \<Longrightarrow> x\<sharp>(y,P)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17468
apply(induct \<theta>c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17469
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17470
apply(case_tac "a=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17471
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17472
apply(case_tac "a=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17473
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17474
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17475
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17476
lemma nmaps_false:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17477
  shows "\<theta>n nmaps x to Some (c,P) \<Longrightarrow> x\<sharp>\<theta>n \<Longrightarrow> False"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17478
apply(induct \<theta>n)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17479
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17480
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17481
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17482
lemma cmaps_false:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17483
  shows "\<theta>c cmaps c to Some (x,P) \<Longrightarrow> c\<sharp>\<theta>c \<Longrightarrow> False"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17484
apply(induct \<theta>c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17485
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17486
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17487
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17488
fun 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17489
 lookupa :: "name\<Rightarrow>coname\<Rightarrow>(coname\<times>name\<times>trm) list\<Rightarrow>trm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17490
where
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17491
  "lookupa x a [] = Ax x a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17492
| "lookupa x a ((c,y,P)#\<theta>c) = (if a=c then Cut <a>.Ax x a (y).P else lookupa x a \<theta>c)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17493
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17494
lemma lookupa_eqvt[eqvt]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17495
  fixes pi1::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17496
  and   pi2::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17497
  shows "(pi1\<bullet>(lookupa x a \<theta>c)) = lookupa (pi1\<bullet>x) (pi1\<bullet>a) (pi1\<bullet>\<theta>c)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17498
  and   "(pi2\<bullet>(lookupa x a \<theta>c)) = lookupa (pi2\<bullet>x) (pi2\<bullet>a) (pi2\<bullet>\<theta>c)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17499
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17500
apply(induct \<theta>c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17501
apply(auto simp add: eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17502
apply(induct \<theta>c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17503
apply(auto simp add: eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17504
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17505
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17506
lemma lookupa_fire:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17507
  assumes a: "\<theta>c cmaps a to Some (y,P)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17508
  shows "(lookupa x a \<theta>c) = Cut <a>.Ax x a (y).P"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17509
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17510
apply(induct \<theta>c arbitrary: x a y P)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17511
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17512
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17513
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17514
fun 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17515
 lookupb :: "name\<Rightarrow>coname\<Rightarrow>(coname\<times>name\<times>trm) list\<Rightarrow>coname\<Rightarrow>trm\<Rightarrow>trm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17516
where
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17517
  "lookupb x a [] c P = Cut <c>.P (x).Ax x a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17518
| "lookupb x a ((d,y,N)#\<theta>c) c P = (if a=d then Cut <c>.P (y).N  else lookupb x a \<theta>c c P)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17519
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17520
lemma lookupb_eqvt[eqvt]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17521
  fixes pi1::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17522
  and   pi2::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17523
  shows "(pi1\<bullet>(lookupb x a \<theta>c c P)) = lookupb (pi1\<bullet>x) (pi1\<bullet>a) (pi1\<bullet>\<theta>c) (pi1\<bullet>c) (pi1\<bullet>P)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17524
  and   "(pi2\<bullet>(lookupb x a \<theta>c c P)) = lookupb (pi2\<bullet>x) (pi2\<bullet>a) (pi2\<bullet>\<theta>c) (pi2\<bullet>c) (pi2\<bullet>P)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17525
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17526
apply(induct \<theta>c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17527
apply(auto simp add: eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17528
apply(induct \<theta>c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17529
apply(auto simp add: eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17530
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17531
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17532
fun 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17533
  lookup :: "name\<Rightarrow>coname\<Rightarrow>(name\<times>coname\<times>trm) list\<Rightarrow>(coname\<times>name\<times>trm) list\<Rightarrow>trm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17534
where
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17535
  "lookup x a [] \<theta>c = lookupa x a \<theta>c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17536
| "lookup x a ((y,c,P)#\<theta>n) \<theta>c = (if x=y then (lookupb x a \<theta>c c P) else lookup x a \<theta>n \<theta>c)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17537
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17538
lemma lookup_eqvt[eqvt]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17539
  fixes pi1::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17540
  and   pi2::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17541
  shows "(pi1\<bullet>(lookup x a \<theta>n \<theta>c)) = lookup (pi1\<bullet>x) (pi1\<bullet>a) (pi1\<bullet>\<theta>n) (pi1\<bullet>\<theta>c)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17542
  and   "(pi2\<bullet>(lookup x a \<theta>n \<theta>c)) = lookup (pi2\<bullet>x) (pi2\<bullet>a) (pi2\<bullet>\<theta>n) (pi2\<bullet>\<theta>c)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17543
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17544
apply(induct \<theta>n)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17545
apply(auto simp add: eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17546
apply(induct \<theta>n)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17547
apply(auto simp add: eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17548
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17549
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17550
fun 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17551
  lookupc :: "name\<Rightarrow>coname\<Rightarrow>(name\<times>coname\<times>trm) list\<Rightarrow>trm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17552
where
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17553
  "lookupc x a [] = Ax x a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17554
| "lookupc x a ((y,c,P)#\<theta>n) = (if x=y then P[c\<turnstile>c>a] else lookupc x a \<theta>n)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17555
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17556
lemma lookupc_eqvt[eqvt]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17557
  fixes pi1::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17558
  and   pi2::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17559
  shows "(pi1\<bullet>(lookupc x a \<theta>n)) = lookupc (pi1\<bullet>x) (pi1\<bullet>a) (pi1\<bullet>\<theta>n)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17560
  and   "(pi2\<bullet>(lookupc x a \<theta>n)) = lookupc (pi2\<bullet>x) (pi2\<bullet>a) (pi2\<bullet>\<theta>n)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17561
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17562
apply(induct \<theta>n)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17563
apply(auto simp add: eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17564
apply(induct \<theta>n)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17565
apply(auto simp add: eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17566
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17567
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17568
fun 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17569
  lookupd :: "name\<Rightarrow>coname\<Rightarrow>(coname\<times>name\<times>trm) list\<Rightarrow>trm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17570
where
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17571
  "lookupd x a [] = Ax x a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17572
| "lookupd x a ((c,y,P)#\<theta>c) = (if a=c then P[y\<turnstile>n>x] else lookupd x a \<theta>c)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17573
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17574
lemma lookupd_eqvt[eqvt]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17575
  fixes pi1::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17576
  and   pi2::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17577
  shows "(pi1\<bullet>(lookupd x a \<theta>n)) = lookupd (pi1\<bullet>x) (pi1\<bullet>a) (pi1\<bullet>\<theta>n)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17578
  and   "(pi2\<bullet>(lookupd x a \<theta>n)) = lookupd (pi2\<bullet>x) (pi2\<bullet>a) (pi2\<bullet>\<theta>n)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17579
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17580
apply(induct \<theta>n)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17581
apply(auto simp add: eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17582
apply(induct \<theta>n)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17583
apply(auto simp add: eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17584
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17585
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17586
lemma lookupa_fresh:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17587
  assumes a: "a\<sharp>\<theta>c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17588
  shows "lookupa y a \<theta>c = Ax y a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17589
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17590
apply(induct \<theta>c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17591
apply(auto simp add: fresh_prod fresh_list_cons fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17592
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17593
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17594
lemma lookupa_csubst:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17595
  assumes a: "a\<sharp>\<theta>c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17596
  shows "Cut <a>.Ax y a (x).P = (lookupa y a \<theta>c){a:=(x).P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17597
using a by (simp add: lookupa_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17598
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17599
lemma lookupa_freshness:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17600
  fixes a::"coname"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17601
  and   x::"name"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17602
  shows "a\<sharp>(\<theta>c,c) \<Longrightarrow> a\<sharp>lookupa y c \<theta>c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17603
  and   "x\<sharp>(\<theta>c,y) \<Longrightarrow> x\<sharp>lookupa y c \<theta>c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17604
apply(induct \<theta>c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17605
apply(auto simp add: fresh_prod fresh_list_cons abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17606
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17607
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17608
lemma lookupa_unicity:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17609
  assumes a: "lookupa x a \<theta>c= Ax y b" "b\<sharp>\<theta>c" "y\<sharp>\<theta>c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17610
  shows "x=y \<and> a=b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17611
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17612
apply(induct \<theta>c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17613
apply(auto simp add: trm.inject fresh_list_cons fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17614
apply(case_tac "a=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17615
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17616
apply(case_tac "a=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17617
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17618
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17619
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17620
lemma lookupb_csubst:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17621
  assumes a: "a\<sharp>(\<theta>c,c,N)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17622
  shows "Cut <c>.N (x).P = (lookupb y a \<theta>c c N){a:=(x).P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17623
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17624
apply(induct \<theta>c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17625
apply(auto simp add: fresh_list_cons fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17626
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17627
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17628
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17629
apply(subgoal_tac "Cut <c>.N (y).Ax y a = Cut <caa>.([(caa,c)]\<bullet>N) (ca).Ax ca a")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17630
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17631
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17632
apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17633
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17634
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17635
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17636
apply(subgoal_tac "a\<sharp>([(caa,c)]\<bullet>N)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17637
apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17638
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17639
apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17640
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17641
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17642
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17643
apply(simp add: alpha fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17644
apply(simp add: alpha calc_atm fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17645
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17646
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17647
lemma lookupb_freshness:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17648
  fixes a::"coname"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17649
  and   x::"name"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17650
  shows "a\<sharp>(\<theta>c,c,b,P) \<Longrightarrow> a\<sharp>lookupb y c \<theta>c b P"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17651
  and   "x\<sharp>(\<theta>c,y,P) \<Longrightarrow> x\<sharp>lookupb y c \<theta>c b P"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17652
apply(induct \<theta>c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17653
apply(auto simp add: fresh_prod fresh_list_cons abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17654
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17655
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17656
lemma lookupb_unicity:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17657
  assumes a: "lookupb x a \<theta>c c P = Ax y b" "b\<sharp>(\<theta>c,c,P)" "y\<sharp>\<theta>c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17658
  shows "x=y \<and> a=b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17659
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17660
apply(induct \<theta>c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17661
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17662
apply(case_tac "a=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17663
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17664
apply(case_tac "a=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17665
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17666
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17667
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17668
lemma lookupb_lookupa:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17669
  assumes a: "x\<sharp>\<theta>c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17670
  shows "lookupb x c \<theta>c a P = (lookupa x c \<theta>c){x:=<a>.P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17671
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17672
apply(induct \<theta>c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17673
apply(auto simp add: fresh_list_cons fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17674
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17675
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17676
apply(subgoal_tac "Cut <c>.Ax x c (aa).b = Cut <ca>.Ax x ca (caa).([(caa,aa)]\<bullet>b)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17677
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17678
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17679
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17680
apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17681
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17682
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17683
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17684
apply(subgoal_tac "x\<sharp>([(caa,aa)]\<bullet>b)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17685
apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17686
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17687
apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17688
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17689
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17690
apply(simp add: alpha calc_atm fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17691
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17692
apply(simp add: alpha calc_atm fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17693
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17694
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17695
lemma lookup_csubst:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17696
  assumes a: "a\<sharp>(\<theta>n,\<theta>c)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17697
  shows "lookup y c \<theta>n ((a,x,P)#\<theta>c) = (lookup y c \<theta>n \<theta>c){a:=(x).P}"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17698
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17699
apply(induct \<theta>n)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17700
apply(auto simp add: fresh_prod fresh_list_cons)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17701
apply(simp add: lookupa_csubst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17702
apply(simp add: lookupa_freshness forget fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17703
apply(rule lookupb_csubst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17704
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17705
apply(auto simp add: lookupb_freshness forget fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17706
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17707
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17708
lemma lookup_fresh:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17709
  assumes a: "x\<sharp>(\<theta>n,\<theta>c)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17710
  shows "lookup x c \<theta>n \<theta>c = lookupa x c \<theta>c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17711
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17712
apply(induct \<theta>n)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17713
apply(auto simp add: fresh_prod fresh_list_cons fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17714
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17715
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17716
lemma lookup_unicity:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17717
  assumes a: "lookup x a \<theta>n \<theta>c= Ax y b" "b\<sharp>(\<theta>c,\<theta>n)" "y\<sharp>(\<theta>c,\<theta>n)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17718
  shows "x=y \<and> a=b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17719
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17720
apply(induct \<theta>n)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17721
apply(auto simp add: trm.inject fresh_list_cons fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17722
apply(drule lookupa_unicity)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17723
apply(simp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17724
apply(drule lookupa_unicity)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17725
apply(simp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17726
apply(case_tac "x=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17727
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17728
apply(drule lookupb_unicity)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17729
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17730
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17731
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17732
apply(case_tac "x=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17733
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17734
apply(drule lookupb_unicity)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17735
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17736
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17737
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17738
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17739
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17740
lemma lookup_freshness:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17741
  fixes a::"coname"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17742
  and   x::"name"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17743
  shows "a\<sharp>(c,\<theta>c,\<theta>n) \<Longrightarrow> a\<sharp>lookup y c \<theta>n \<theta>c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17744
  and   "x\<sharp>(y,\<theta>c,\<theta>n) \<Longrightarrow> x\<sharp>lookup y c \<theta>n \<theta>c"   
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17745
apply(induct \<theta>n)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17746
apply(auto simp add: fresh_prod fresh_list_cons abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17747
apply(simp add: fresh_atm fresh_prod lookupa_freshness)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17748
apply(simp add: fresh_atm fresh_prod lookupa_freshness)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17749
apply(simp add: fresh_atm fresh_prod lookupb_freshness)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17750
apply(simp add: fresh_atm fresh_prod lookupb_freshness)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17751
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17752
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17753
lemma lookupc_freshness:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17754
  fixes a::"coname"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17755
  and   x::"name"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17756
  shows "a\<sharp>(\<theta>c,c) \<Longrightarrow> a\<sharp>lookupc y c \<theta>c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17757
  and   "x\<sharp>(\<theta>c,y) \<Longrightarrow> x\<sharp>lookupc y c \<theta>c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17758
apply(induct \<theta>c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17759
apply(auto simp add: fresh_prod fresh_list_cons abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17760
apply(rule rename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17761
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17762
apply(rule rename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17763
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17764
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17765
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17766
lemma lookupc_fresh:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17767
  assumes a: "y\<sharp>\<theta>n"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17768
  shows "lookupc y a \<theta>n = Ax y a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17769
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17770
apply(induct \<theta>n)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17771
apply(auto simp add: fresh_prod fresh_list_cons fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17772
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17773
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17774
lemma lookupc_nmaps:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17775
  assumes a: "\<theta>n nmaps x to Some (c,P)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17776
  shows "lookupc x a \<theta>n = P[c\<turnstile>c>a]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17777
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17778
apply(induct \<theta>n)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17779
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17780
done 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17781
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17782
lemma lookupc_unicity:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17783
  assumes a: "lookupc y a \<theta>n = Ax x b" "x\<sharp>\<theta>n"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17784
  shows "y=x"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17785
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17786
apply(induct \<theta>n)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17787
apply(auto simp add: trm.inject fresh_list_cons fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17788
apply(case_tac "y=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17789
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17790
apply(subgoal_tac "x\<sharp>(ba[aa\<turnstile>c>a])")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17791
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17792
apply(rule rename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17793
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17794
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17795
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17796
lemma lookupd_fresh:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17797
  assumes a: "a\<sharp>\<theta>c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17798
  shows "lookupd y a \<theta>c = Ax y a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17799
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17800
apply(induct \<theta>c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17801
apply(auto simp add: fresh_prod fresh_list_cons fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17802
done 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17803
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17804
lemma lookupd_unicity:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17805
  assumes a: "lookupd y a \<theta>c = Ax y b" "b\<sharp>\<theta>c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17806
  shows "a=b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17807
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17808
apply(induct \<theta>c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17809
apply(auto simp add: trm.inject fresh_list_cons fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17810
apply(case_tac "a=aa")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17811
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17812
apply(subgoal_tac "b\<sharp>(ba[aa\<turnstile>n>y])")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17813
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17814
apply(rule rename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17815
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17816
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17817
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17818
lemma lookupd_freshness:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17819
  fixes a::"coname"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17820
  and   x::"name"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17821
  shows "a\<sharp>(\<theta>c,c) \<Longrightarrow> a\<sharp>lookupd y c \<theta>c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17822
  and   "x\<sharp>(\<theta>c,y) \<Longrightarrow> x\<sharp>lookupd y c \<theta>c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17823
apply(induct \<theta>c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17824
apply(auto simp add: fresh_prod fresh_list_cons abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17825
apply(rule rename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17826
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17827
apply(rule rename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17828
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17829
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17830
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17831
lemma lookupd_cmaps:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17832
  assumes a: "\<theta>c cmaps a to Some (x,P)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17833
  shows "lookupd y a \<theta>c = P[x\<turnstile>n>y]"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17834
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17835
apply(induct \<theta>c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17836
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17837
done 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17838
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17839
consts
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17840
  stn :: "trm\<Rightarrow>(name\<times>coname\<times>trm) list\<Rightarrow>trm" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17841
  stc :: "trm\<Rightarrow>(coname\<times>name\<times>trm) list\<Rightarrow>trm" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17842
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17843
nominal_primrec (freshness_context: "\<theta>n::(name\<times>coname\<times>trm)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17844
  "stn (Ax x a) \<theta>n = lookupc x a \<theta>n"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17845
  "\<lbrakk>a\<sharp>(N,\<theta>n);x\<sharp>(M,\<theta>n)\<rbrakk> \<Longrightarrow> stn (Cut <a>.M (x).N) \<theta>n = (Cut <a>.M (x).N)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17846
  "x\<sharp>\<theta>n \<Longrightarrow> stn (NotR (x).M a) \<theta>n = (NotR (x).M a)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17847
  "a\<sharp>\<theta>n \<Longrightarrow>stn (NotL <a>.M x) \<theta>n = (NotL <a>.M x)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17848
  "\<lbrakk>a\<sharp>(N,d,b,\<theta>n);b\<sharp>(M,d,a,\<theta>n)\<rbrakk> \<Longrightarrow> stn (AndR <a>.M <b>.N d) \<theta>n = (AndR <a>.M <b>.N d)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17849
  "x\<sharp>(z,\<theta>n) \<Longrightarrow> stn (AndL1 (x).M z) \<theta>n = (AndL1 (x).M z)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17850
  "x\<sharp>(z,\<theta>n) \<Longrightarrow> stn (AndL2 (x).M z) \<theta>n = (AndL2 (x).M z)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17851
  "a\<sharp>(b,\<theta>n) \<Longrightarrow> stn (OrR1 <a>.M b) \<theta>n = (OrR1 <a>.M b)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17852
  "a\<sharp>(b,\<theta>n) \<Longrightarrow> stn (OrR2 <a>.M b) \<theta>n = (OrR2 <a>.M b)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17853
  "\<lbrakk>x\<sharp>(N,z,u,\<theta>n);u\<sharp>(M,z,x,\<theta>n)\<rbrakk> \<Longrightarrow> stn (OrL (x).M (u).N z) \<theta>n = (OrL (x).M (u).N z)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17854
  "\<lbrakk>a\<sharp>(b,\<theta>n);x\<sharp>\<theta>n\<rbrakk> \<Longrightarrow> stn (ImpR (x).<a>.M b) \<theta>n = (ImpR (x).<a>.M b)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17855
  "\<lbrakk>a\<sharp>(N,\<theta>n);x\<sharp>(M,z,\<theta>n)\<rbrakk> \<Longrightarrow> stn (ImpL <a>.M (x).N z) \<theta>n = (ImpL <a>.M (x).N z)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17856
apply(finite_guess)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17857
apply(rule TrueI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17858
apply(simp add: abs_fresh abs_supp fin_supp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17859
apply(fresh_guess)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17860
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17861
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17862
nominal_primrec (freshness_context: "\<theta>c::(coname\<times>name\<times>trm)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17863
  "stc (Ax x a) \<theta>c = lookupd x a \<theta>c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17864
  "\<lbrakk>a\<sharp>(N,\<theta>c);x\<sharp>(M,\<theta>c)\<rbrakk> \<Longrightarrow> stc (Cut <a>.M (x).N) \<theta>c = (Cut <a>.M (x).N)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17865
  "x\<sharp>\<theta>c \<Longrightarrow> stc (NotR (x).M a) \<theta>c = (NotR (x).M a)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17866
  "a\<sharp>\<theta>c \<Longrightarrow> stc (NotL <a>.M x) \<theta>c = (NotL <a>.M x)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17867
  "\<lbrakk>a\<sharp>(N,d,b,\<theta>c);b\<sharp>(M,d,a,\<theta>c)\<rbrakk> \<Longrightarrow> stc (AndR <a>.M <b>.N d) \<theta>c = (AndR <a>.M <b>.N d)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17868
  "x\<sharp>(z,\<theta>c) \<Longrightarrow> stc (AndL1 (x).M z) \<theta>c = (AndL1 (x).M z)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17869
  "x\<sharp>(z,\<theta>c) \<Longrightarrow> stc (AndL2 (x).M z) \<theta>c = (AndL2 (x).M z)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17870
  "a\<sharp>(b,\<theta>c) \<Longrightarrow> stc (OrR1 <a>.M b) \<theta>c = (OrR1 <a>.M b)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17871
  "a\<sharp>(b,\<theta>c) \<Longrightarrow> stc (OrR2 <a>.M b) \<theta>c = (OrR2 <a>.M b)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17872
  "\<lbrakk>x\<sharp>(N,z,u,\<theta>c);u\<sharp>(M,z,x,\<theta>c)\<rbrakk> \<Longrightarrow> stc (OrL (x).M (u).N z) \<theta>c = (OrL (x).M (u).N z)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17873
  "\<lbrakk>a\<sharp>(b,\<theta>c);x\<sharp>\<theta>c\<rbrakk> \<Longrightarrow> stc (ImpR (x).<a>.M b) \<theta>c = (ImpR (x).<a>.M b)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17874
  "\<lbrakk>a\<sharp>(N,\<theta>c);x\<sharp>(M,z,\<theta>c)\<rbrakk> \<Longrightarrow> stc (ImpL <a>.M (x).N z) \<theta>c = (ImpL <a>.M (x).N z)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17875
apply(finite_guess)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17876
apply(rule TrueI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17877
apply(simp add: abs_fresh abs_supp fin_supp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17878
apply(fresh_guess)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17879
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17880
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17881
lemma stn_eqvt[eqvt]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17882
  fixes pi1::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17883
  and   pi2::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17884
  shows "(pi1\<bullet>(stn M \<theta>n)) = stn (pi1\<bullet>M) (pi1\<bullet>\<theta>n)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17885
  and   "(pi2\<bullet>(stn M \<theta>n)) = stn (pi2\<bullet>M) (pi2\<bullet>\<theta>n)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17886
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17887
apply(nominal_induct M avoiding: \<theta>n rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17888
apply(auto simp add: eqvts fresh_bij fresh_prod eq_bij fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17889
apply(nominal_induct M avoiding: \<theta>n rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17890
apply(auto simp add: eqvts fresh_bij fresh_prod eq_bij fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17891
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17892
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17893
lemma stc_eqvt[eqvt]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17894
  fixes pi1::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17895
  and   pi2::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17896
  shows "(pi1\<bullet>(stc M \<theta>c)) = stc (pi1\<bullet>M) (pi1\<bullet>\<theta>c)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17897
  and   "(pi2\<bullet>(stc M \<theta>c)) = stc (pi2\<bullet>M) (pi2\<bullet>\<theta>c)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17898
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17899
apply(nominal_induct M avoiding: \<theta>c rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17900
apply(auto simp add: eqvts fresh_bij fresh_prod eq_bij fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17901
apply(nominal_induct M avoiding: \<theta>c rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17902
apply(auto simp add: eqvts fresh_bij fresh_prod eq_bij fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17903
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17904
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17905
lemma stn_fresh:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17906
  fixes a::"coname"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17907
  and   x::"name"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17908
  shows "a\<sharp>(\<theta>n,M) \<Longrightarrow> a\<sharp>stn M \<theta>n"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17909
  and   "x\<sharp>(\<theta>n,M) \<Longrightarrow> x\<sharp>stn M \<theta>n"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17910
apply(nominal_induct M avoiding: \<theta>n a x rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17911
apply(auto simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17912
apply(rule lookupc_freshness)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17913
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17914
apply(rule lookupc_freshness)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17915
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17916
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17917
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17918
lemma stc_fresh:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17919
  fixes a::"coname"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17920
  and   x::"name"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17921
  shows "a\<sharp>(\<theta>c,M) \<Longrightarrow> a\<sharp>stc M \<theta>c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17922
  and   "x\<sharp>(\<theta>c,M) \<Longrightarrow> x\<sharp>stc M \<theta>c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17923
apply(nominal_induct M avoiding: \<theta>c a x rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17924
apply(auto simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17925
apply(rule lookupd_freshness)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17926
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17927
apply(rule lookupd_freshness)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17928
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17929
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17930
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17931
lemma option_case_eqvt1[eqvt_force]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17932
  fixes pi1::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17933
  and   pi2::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17934
  and   B::"(name\<times>trm) option"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17935
  and   r::"trm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17936
  shows "(pi1\<bullet>(case B of Some (x,P) \<Rightarrow> s x P | None \<Rightarrow> r)) = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17937
              (case (pi1\<bullet>B) of Some (x,P) \<Rightarrow> (pi1\<bullet>s) x P | None \<Rightarrow> (pi1\<bullet>r))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17938
  and   "(pi2\<bullet>(case B of Some (x,P) \<Rightarrow> s x P| None \<Rightarrow> r)) = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17939
              (case (pi2\<bullet>B) of Some (x,P) \<Rightarrow> (pi2\<bullet>s) x P | None \<Rightarrow> (pi2\<bullet>r))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17940
apply(cases "B")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17941
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17942
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17943
apply(cases "B")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17944
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17945
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17946
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17947
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17948
lemma option_case_eqvt2[eqvt_force]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17949
  fixes pi1::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17950
  and   pi2::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17951
  and   B::"(coname\<times>trm) option"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17952
  and   r::"trm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17953
  shows "(pi1\<bullet>(case B of Some (x,P) \<Rightarrow> s x P | None \<Rightarrow> r)) = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17954
              (case (pi1\<bullet>B) of Some (x,P) \<Rightarrow> (pi1\<bullet>s) x P | None \<Rightarrow> (pi1\<bullet>r))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17955
  and   "(pi2\<bullet>(case B of Some (x,P) \<Rightarrow> s x P| None \<Rightarrow> r)) = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17956
              (case (pi2\<bullet>B) of Some (x,P) \<Rightarrow> (pi2\<bullet>s) x P | None \<Rightarrow> (pi2\<bullet>r))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17957
apply(cases "B")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17958
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17959
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17960
apply(cases "B")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17961
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17962
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17963
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17964
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17965
consts
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17966
  psubst :: "(name\<times>coname\<times>trm) list\<Rightarrow>(coname\<times>name\<times>trm) list\<Rightarrow>trm\<Rightarrow>trm" ("_,_<_>" [100,100,100] 100) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17967
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17968
nominal_primrec (freshness_context: "(\<theta>n::(name\<times>coname\<times>trm) list,\<theta>c::(coname\<times>name\<times>trm) list)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17969
  "\<theta>n,\<theta>c<Ax x a> = lookup x a \<theta>n \<theta>c" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17970
  "\<lbrakk>a\<sharp>(N,\<theta>n,\<theta>c);x\<sharp>(M,\<theta>n,\<theta>c)\<rbrakk> \<Longrightarrow> \<theta>n,\<theta>c<Cut <a>.M (x).N> = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17971
   Cut <a>.(if \<exists>x. M=Ax x a then stn M \<theta>n else \<theta>n,\<theta>c<M>) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17972
       (x).(if \<exists>a. N=Ax x a then stc N \<theta>c else \<theta>n,\<theta>c<N>)" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17973
  "x\<sharp>(\<theta>n,\<theta>c) \<Longrightarrow> \<theta>n,\<theta>c<NotR (x).M a> = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17974
  (case (findc \<theta>c a) of 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17975
       Some (u,P) \<Rightarrow> fresh_fun (\<lambda>a'. Cut <a'>.NotR (x).(\<theta>n,\<theta>c<M>) a' (u).P) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17976
     | None \<Rightarrow> NotR (x).(\<theta>n,\<theta>c<M>) a)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17977
  "a\<sharp>(\<theta>n,\<theta>c) \<Longrightarrow> \<theta>n,\<theta>c<NotL <a>.M x> = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17978
  (case (findn \<theta>n x) of 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17979
       Some (c,P) \<Rightarrow> fresh_fun (\<lambda>x'. Cut <c>.P (x').(NotL <a>.(\<theta>n,\<theta>c<M>) x')) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17980
     | None \<Rightarrow> NotL <a>.(\<theta>n,\<theta>c<M>) x)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17981
  "\<lbrakk>a\<sharp>(N,c,\<theta>n,\<theta>c);b\<sharp>(M,c,\<theta>n,\<theta>c);b\<noteq>a\<rbrakk> \<Longrightarrow> (\<theta>n,\<theta>c<AndR <a>.M <b>.N c>) = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17982
  (case (findc \<theta>c c) of 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17983
       Some (x,P) \<Rightarrow> fresh_fun (\<lambda>a'. Cut <a'>.(AndR <a>.(\<theta>n,\<theta>c<M>) <b>.(\<theta>n,\<theta>c<N>) a') (x).P)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17984
     | None \<Rightarrow> AndR <a>.(\<theta>n,\<theta>c<M>) <b>.(\<theta>n,\<theta>c<N>) c)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17985
  "x\<sharp>(z,\<theta>n,\<theta>c) \<Longrightarrow> (\<theta>n,\<theta>c<AndL1 (x).M z>) = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17986
  (case (findn \<theta>n z) of 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17987
       Some (c,P) \<Rightarrow> fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL1 (x).(\<theta>n,\<theta>c<M>) z') 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17988
     | None \<Rightarrow> AndL1 (x).(\<theta>n,\<theta>c<M>) z)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17989
  "x\<sharp>(z,\<theta>n,\<theta>c) \<Longrightarrow> (\<theta>n,\<theta>c<AndL2 (x).M z>) = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17990
  (case (findn \<theta>n z) of 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17991
       Some (c,P) \<Rightarrow> fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL2 (x).(\<theta>n,\<theta>c<M>) z') 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17992
     | None \<Rightarrow> AndL2 (x).(\<theta>n,\<theta>c<M>) z)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17993
  "\<lbrakk>x\<sharp>(N,z,\<theta>n,\<theta>c);u\<sharp>(M,z,\<theta>n,\<theta>c);x\<noteq>u\<rbrakk> \<Longrightarrow> (\<theta>n,\<theta>c<OrL (x).M (u).N z>) =
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17994
  (case (findn \<theta>n z) of  
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17995
       Some (c,P) \<Rightarrow> fresh_fun (\<lambda>z'. Cut <c>.P (z').OrL (x).(\<theta>n,\<theta>c<M>) (u).(\<theta>n,\<theta>c<N>) z') 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17996
     | None \<Rightarrow> OrL (x).(\<theta>n,\<theta>c<M>) (u).(\<theta>n,\<theta>c<N>) z)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17997
  "a\<sharp>(b,\<theta>n,\<theta>c) \<Longrightarrow> (\<theta>n,\<theta>c<OrR1 <a>.M b>) = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17998
  (case (findc \<theta>c b) of
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 17999
       Some (x,P) \<Rightarrow> fresh_fun (\<lambda>a'. Cut <a'>.OrR1 <a>.(\<theta>n,\<theta>c<M>) a' (x).P) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18000
     | None \<Rightarrow> OrR1 <a>.(\<theta>n,\<theta>c<M>) b)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18001
  "a\<sharp>(b,\<theta>n,\<theta>c) \<Longrightarrow> (\<theta>n,\<theta>c<OrR2 <a>.M b>) = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18002
  (case (findc \<theta>c b) of
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18003
       Some (x,P) \<Rightarrow> fresh_fun (\<lambda>a'. Cut <a'>.OrR2 <a>.(\<theta>n,\<theta>c<M>) a' (x).P) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18004
     | None \<Rightarrow> OrR2 <a>.(\<theta>n,\<theta>c<M>) b)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18005
  "\<lbrakk>a\<sharp>(b,\<theta>n,\<theta>c); x\<sharp>(\<theta>n,\<theta>c)\<rbrakk> \<Longrightarrow> (\<theta>n,\<theta>c<ImpR (x).<a>.M b>) = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18006
  (case (findc \<theta>c b) of
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18007
       Some (z,P) \<Rightarrow> fresh_fun (\<lambda>a'. Cut <a'>.ImpR (x).<a>.(\<theta>n,\<theta>c<M>) a' (z).P)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18008
     | None \<Rightarrow> ImpR (x).<a>.(\<theta>n,\<theta>c<M>) b)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18009
  "\<lbrakk>a\<sharp>(N,\<theta>n,\<theta>c); x\<sharp>(z,M,\<theta>n,\<theta>c)\<rbrakk> \<Longrightarrow> (\<theta>n,\<theta>c<ImpL <a>.M (x).N z>) = 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18010
  (case (findn \<theta>n z) of
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18011
       Some (c,P) \<Rightarrow> fresh_fun (\<lambda>z'. Cut <c>.P (z').ImpL <a>.(\<theta>n,\<theta>c<M>) (x).(\<theta>n,\<theta>c<N>) z') 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18012
     | None \<Rightarrow> ImpL <a>.(\<theta>n,\<theta>c<M>) (x).(\<theta>n,\<theta>c<N>) z)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18013
apply(finite_guess)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18014
apply(rule TrueI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18015
apply(simp add: abs_fresh stc_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18016
apply(simp add: abs_fresh stn_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18017
apply(case_tac "findc \<theta>c x3")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18018
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18019
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18020
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18021
apply(thin_tac "x1\<sharp>?p")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18022
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18023
apply(drule cmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18024
apply(auto simp add: fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18025
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18026
apply(case_tac "findn \<theta>n x3")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18027
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18028
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18029
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18030
apply(thin_tac "x1\<sharp>?p")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18031
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18032
apply(drule nmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18033
apply(auto simp add: fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18034
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18035
apply(case_tac "findc \<theta>c x5")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18036
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18037
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18038
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18039
apply(thin_tac "x1\<sharp>?p")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18040
apply(thin_tac "x3\<sharp>?p")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18041
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18042
apply(drule cmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18043
apply(auto simp add: fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18044
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18045
apply(case_tac "findc \<theta>c x5")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18046
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18047
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18048
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18049
apply(thin_tac "x1\<sharp>?p")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18050
apply(thin_tac "x3\<sharp>?p")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18051
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18052
apply(drule_tac x="x3" in cmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18053
apply(auto simp add: fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18054
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18055
apply(case_tac "findn \<theta>n x3")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18056
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18057
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18058
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18059
apply(thin_tac "x1\<sharp>?p")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18060
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18061
apply(drule nmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18062
apply(auto simp add: fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18063
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18064
apply(case_tac "findn \<theta>n x3")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18065
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18066
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18067
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18068
apply(thin_tac "x1\<sharp>?p")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18069
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18070
apply(drule nmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18071
apply(auto simp add: fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18072
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18073
apply(case_tac "findc \<theta>c x3")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18074
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18075
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18076
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18077
apply(thin_tac "x1\<sharp>?p")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18078
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18079
apply(drule cmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18080
apply(auto simp add: fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18081
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18082
apply(case_tac "findc \<theta>c x3")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18083
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18084
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18085
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18086
apply(thin_tac "x1\<sharp>?p")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18087
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18088
apply(drule cmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18089
apply(auto simp add: fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18090
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18091
apply(case_tac "findn \<theta>n x5")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18092
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18093
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18094
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18095
apply(thin_tac "x1\<sharp>?p")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18096
apply(thin_tac "x3\<sharp>?p")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18097
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18098
apply(drule nmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18099
apply(auto simp add: fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18100
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18101
apply(case_tac "findn \<theta>n x5")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18102
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18103
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18104
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18105
apply(thin_tac "x1\<sharp>?p")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18106
apply(thin_tac "x3\<sharp>?p")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18107
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18108
apply(drule_tac a="x3" in nmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18109
apply(auto simp add: fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18110
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18111
apply(case_tac "findc \<theta>c x4")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18112
apply(simp add: abs_fresh abs_supp fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18113
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18114
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18115
apply(thin_tac "x1\<sharp>?p")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18116
apply(thin_tac "x2\<sharp>?p")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18117
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18118
apply(drule cmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18119
apply(auto simp add: fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18120
apply(simp add: abs_fresh fresh_prod fresh_atm abs_supp fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18121
apply(case_tac "findc \<theta>c x4")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18122
apply(simp add: abs_fresh abs_supp fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18123
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18124
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18125
apply(thin_tac "x1\<sharp>?p")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18126
apply(thin_tac "x2\<sharp>?p")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18127
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18128
apply(drule_tac x="x2" in cmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18129
apply(auto simp add: fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18130
apply(simp add: abs_fresh fresh_prod fresh_atm abs_supp fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18131
apply(case_tac "findn \<theta>n x5")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18132
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18133
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18134
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18135
apply(thin_tac "x1\<sharp>?p")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18136
apply(thin_tac "x3\<sharp>?p")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18137
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18138
apply(drule nmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18139
apply(auto simp add: fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18140
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18141
apply(case_tac "findn \<theta>n x5")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18142
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18143
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18144
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18145
apply(thin_tac "x1\<sharp>?p")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18146
apply(thin_tac "x3\<sharp>?p")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18147
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18148
apply(drule_tac a="x3" in nmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18149
apply(auto simp add: fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18150
apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18151
apply(fresh_guess)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18152
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18153
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18154
lemma case_cong:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18155
  assumes a: "B1=B2" "x1=x2" "y1=y2"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18156
  shows "(case B1 of None \<Rightarrow> x1 | Some (x,P) \<Rightarrow> y1 x P) = (case B2 of None \<Rightarrow> x2 | Some (x,P) \<Rightarrow> y2 x P)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18157
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18158
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18159
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18160
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18161
lemma find_maps:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18162
  shows "\<theta>c cmaps a to (findc \<theta>c a)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18163
  and   "\<theta>n nmaps x to (findn \<theta>n x)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18164
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18165
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18166
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18167
lemma psubst_eqvt[eqvt]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18168
  fixes pi1::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18169
  and   pi2::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18170
  shows "pi1\<bullet>(\<theta>n,\<theta>c<M>) = (pi1\<bullet>\<theta>n),(pi1\<bullet>\<theta>c)<(pi1\<bullet>M)>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18171
  and   "pi2\<bullet>(\<theta>n,\<theta>c<M>) = (pi2\<bullet>\<theta>n),(pi2\<bullet>\<theta>c)<(pi2\<bullet>M)>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18172
apply(nominal_induct M avoiding: \<theta>n \<theta>c rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18173
apply(auto simp add: eq_bij fresh_bij eqvts perm_pi_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18174
apply(rule case_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18175
apply(rule find_maps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18176
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18177
apply(perm_simp add: eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18178
apply(rule case_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18179
apply(rule find_maps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18180
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18181
apply(perm_simp add: eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18182
apply(rule case_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18183
apply(rule find_maps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18184
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18185
apply(perm_simp add: eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18186
apply(rule case_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18187
apply(rule find_maps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18188
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18189
apply(perm_simp add: eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18190
apply(rule case_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18191
apply(rule find_maps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18192
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18193
apply(perm_simp add: eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18194
apply(rule case_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18195
apply(rule find_maps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18196
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18197
apply(perm_simp add: eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18198
apply(rule case_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18199
apply(rule find_maps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18200
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18201
apply(perm_simp add: eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18202
apply(rule case_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18203
apply(rule find_maps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18204
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18205
apply(perm_simp add: eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18206
apply(rule case_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18207
apply(rule find_maps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18208
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18209
apply(perm_simp add: eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18210
apply(rule case_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18211
apply(rule find_maps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18212
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18213
apply(perm_simp add: eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18214
apply(rule case_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18215
apply(rule find_maps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18216
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18217
apply(perm_simp add: eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18218
apply(rule case_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18219
apply(rule find_maps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18220
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18221
apply(perm_simp add: eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18222
apply(rule case_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18223
apply(rule find_maps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18224
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18225
apply(perm_simp add: eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18226
apply(rule case_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18227
apply(rule find_maps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18228
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18229
apply(perm_simp add: eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18230
apply(rule case_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18231
apply(rule find_maps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18232
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18233
apply(perm_simp add: eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18234
apply(rule case_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18235
apply(rule find_maps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18236
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18237
apply(perm_simp add: eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18238
apply(rule case_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18239
apply(rule find_maps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18240
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18241
apply(perm_simp add: eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18242
apply(rule case_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18243
apply(rule find_maps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18244
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18245
apply(perm_simp add: eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18246
apply(rule case_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18247
apply(rule find_maps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18248
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18249
apply(perm_simp add: eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18250
apply(rule case_cong)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18251
apply(rule find_maps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18252
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18253
apply(perm_simp add: eqvts)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18254
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18255
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18256
lemma ax_psubst:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18257
  assumes a: "\<theta>n,\<theta>c<M> = Ax x a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18258
  and     b: "a\<sharp>(\<theta>n,\<theta>c)" "x\<sharp>(\<theta>n,\<theta>c)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18259
  shows "M = Ax x a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18260
using a b
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18261
apply(nominal_induct M avoiding: \<theta>n \<theta>c rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18262
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18263
apply(drule lookup_unicity)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18264
apply(simp)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18265
apply(case_tac "findc \<theta>c coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18266
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18267
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18268
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18269
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18270
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18271
apply(case_tac "findn \<theta>n name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18272
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18273
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18274
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18275
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18276
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18277
apply(case_tac "findc \<theta>c coname3")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18278
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18279
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18280
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18281
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18282
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18283
apply(case_tac "findn \<theta>n name2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18284
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18285
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18286
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18287
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18288
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18289
apply(case_tac "findn \<theta>n name2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18290
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18291
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18292
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18293
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18294
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18295
apply(case_tac "findc \<theta>c coname2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18296
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18297
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18298
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18299
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18300
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18301
apply(case_tac "findc \<theta>c coname2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18302
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18303
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18304
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18305
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18306
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18307
apply(case_tac "findn \<theta>n name3")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18308
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18309
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18310
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18311
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18312
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18313
apply(case_tac "findc \<theta>c coname2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18314
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18315
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18316
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18317
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18318
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18319
apply(case_tac "findn \<theta>n name2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18320
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18321
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18322
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18323
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18324
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18325
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18326
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18327
lemma better_Cut_substc1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18328
  assumes a: "a\<sharp>(P,b)" "b\<sharp>N" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18329
  shows "(Cut <a>.M (x).N){b:=(y).P} = Cut <a>.(M{b:=(y).P}) (x).N"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18330
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18331
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18332
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18333
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18334
apply(subgoal_tac "Cut <a>.M (x).N = Cut <c>.([(c,a)]\<bullet>M) (ca).([(ca,x)]\<bullet>N)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18335
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18336
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18337
apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18338
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18339
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18340
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18341
apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18342
apply(simp add: calc_atm fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18343
apply(subgoal_tac"b\<sharp>([(ca,x)]\<bullet>N)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18344
apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18345
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18346
apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18347
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18348
apply(simp add: fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18349
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18350
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18351
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18352
apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18353
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18354
apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18355
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18356
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18357
lemma better_Cut_substc2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18358
  assumes a: "x\<sharp>(y,P)" "b\<sharp>(a,M)" "N\<noteq>Ax x b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18359
  shows "(Cut <a>.M (x).N){b:=(y).P} = Cut <a>.M (x).(N{b:=(y).P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18360
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18361
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18362
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18363
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18364
apply(subgoal_tac "Cut <a>.M (x).N = Cut <c>.([(c,a)]\<bullet>M) (ca).([(ca,x)]\<bullet>N)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18365
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18366
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18367
apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18368
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18369
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18370
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18371
apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18372
apply(simp add: calc_atm fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18373
apply(subgoal_tac"b\<sharp>([(c,a)]\<bullet>M)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18374
apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18375
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18376
apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18377
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18378
apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18379
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18380
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18381
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18382
apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18383
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18384
apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18385
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18386
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18387
lemma better_Cut_substn1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18388
  assumes a: "y\<sharp>(x,N)" "a\<sharp>(b,P)" "M\<noteq>Ax y a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18389
  shows "(Cut <a>.M (x).N){y:=<b>.P} = Cut <a>.(M{y:=<b>.P}) (x).N"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18390
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18391
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18392
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18393
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18394
apply(subgoal_tac "Cut <a>.M (x).N = Cut <c>.([(c,a)]\<bullet>M) (ca).([(ca,x)]\<bullet>N)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18395
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18396
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18397
apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18398
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18399
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18400
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18401
apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18402
apply(simp add: calc_atm fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18403
apply(subgoal_tac"y\<sharp>([(ca,x)]\<bullet>N)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18404
apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18405
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18406
apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18407
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18408
apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18409
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18410
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18411
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18412
apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18413
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18414
apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18415
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18416
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18417
lemma better_Cut_substn2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18418
  assumes a: "x\<sharp>(P,y)" "y\<sharp>M" 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18419
  shows "(Cut <a>.M (x).N){y:=<b>.P} = Cut <a>.M (x).(N{y:=<b>.P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18420
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18421
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18422
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18423
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18424
apply(subgoal_tac "Cut <a>.M (x).N = Cut <c>.([(c,a)]\<bullet>M) (ca).([(ca,x)]\<bullet>N)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18425
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18426
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18427
apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18428
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18429
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18430
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18431
apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18432
apply(simp add: calc_atm fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18433
apply(subgoal_tac"y\<sharp>([(c,a)]\<bullet>M)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18434
apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18435
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18436
apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18437
apply(perm_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18438
apply(simp add: fresh_left calc_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18439
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18440
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18441
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18442
apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18443
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18444
apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18445
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18446
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18447
lemma psubst_fresh_name:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18448
  fixes x::"name"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18449
  assumes a: "x\<sharp>\<theta>n" "x\<sharp>\<theta>c" "x\<sharp>M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18450
  shows "x\<sharp>\<theta>n,\<theta>c<M>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18451
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18452
apply(nominal_induct M avoiding: x \<theta>n \<theta>c rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18453
apply(simp add: lookup_freshness)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18454
apply(auto simp add: abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18455
apply(simp add: lookupc_freshness)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18456
apply(simp add: lookupc_freshness)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18457
apply(simp add: lookupc_freshness)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18458
apply(simp add: lookupd_freshness)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18459
apply(simp add: lookupd_freshness)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18460
apply(simp add: lookupc_freshness)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18461
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18462
apply(case_tac "findc \<theta>c coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18463
apply(auto simp add: abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18464
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18465
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18466
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18467
apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18468
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18469
apply(case_tac "findn \<theta>n name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18470
apply(auto simp add: abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18471
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18472
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18473
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18474
apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18475
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18476
apply(case_tac "findc \<theta>c coname3")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18477
apply(auto simp add: abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18478
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18479
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18480
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18481
apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18482
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18483
apply(case_tac "findn \<theta>n name2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18484
apply(auto simp add: abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18485
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18486
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18487
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18488
apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18489
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18490
apply(case_tac "findn \<theta>n name2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18491
apply(auto simp add: abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18492
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18493
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18494
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18495
apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18496
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18497
apply(case_tac "findc \<theta>c coname2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18498
apply(auto simp add: abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18499
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18500
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18501
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18502
apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18503
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18504
apply(case_tac "findc \<theta>c coname2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18505
apply(auto simp add: abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18506
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18507
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18508
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18509
apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18510
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18511
apply(case_tac "findn \<theta>n name3")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18512
apply(auto simp add: abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18513
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18514
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18515
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18516
apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18517
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18518
apply(case_tac "findc \<theta>c coname2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18519
apply(auto simp add: abs_fresh abs_supp fin_supp)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18520
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18521
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18522
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18523
apply(simp add: abs_fresh abs_supp fin_supp fresh_prod fresh_atm cmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18524
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18525
apply(case_tac "findn \<theta>n name2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18526
apply(auto simp add: abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18527
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18528
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18529
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18530
apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18531
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18532
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18533
lemma psubst_fresh_coname:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18534
  fixes a::"coname"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18535
  assumes a: "a\<sharp>\<theta>n" "a\<sharp>\<theta>c" "a\<sharp>M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18536
  shows "a\<sharp>\<theta>n,\<theta>c<M>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18537
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18538
apply(nominal_induct M avoiding: a \<theta>n \<theta>c rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18539
apply(simp add: lookup_freshness)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18540
apply(auto simp add: abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18541
apply(simp add: lookupd_freshness)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18542
apply(simp add: lookupd_freshness)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18543
apply(simp add: lookupc_freshness)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18544
apply(simp add: lookupd_freshness)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18545
apply(simp add: lookupc_freshness)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18546
apply(simp add: lookupd_freshness)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18547
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18548
apply(case_tac "findc \<theta>c coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18549
apply(auto simp add: abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18550
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18551
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18552
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18553
apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18554
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18555
apply(case_tac "findn \<theta>n name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18556
apply(auto simp add: abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18557
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18558
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18559
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18560
apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18561
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18562
apply(case_tac "findc \<theta>c coname3")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18563
apply(auto simp add: abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18564
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18565
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18566
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18567
apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18568
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18569
apply(case_tac "findn \<theta>n name2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18570
apply(auto simp add: abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18571
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18572
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18573
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18574
apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18575
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18576
apply(case_tac "findn \<theta>n name2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18577
apply(auto simp add: abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18578
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18579
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18580
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18581
apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18582
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18583
apply(case_tac "findc \<theta>c coname2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18584
apply(auto simp add: abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18585
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18586
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18587
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18588
apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18589
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18590
apply(case_tac "findc \<theta>c coname2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18591
apply(auto simp add: abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18592
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18593
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18594
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18595
apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18596
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18597
apply(case_tac "findn \<theta>n name3")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18598
apply(auto simp add: abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18599
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18600
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18601
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18602
apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18603
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18604
apply(case_tac "findc \<theta>c coname2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18605
apply(auto simp add: abs_fresh abs_supp fin_supp)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18606
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18607
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18608
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18609
apply(simp add: abs_fresh abs_supp fin_supp fresh_prod fresh_atm cmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18610
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18611
apply(case_tac "findn \<theta>n name2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18612
apply(auto simp add: abs_fresh)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18613
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18614
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18615
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18616
apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18617
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18618
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18619
lemma psubst_csubst:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18620
  assumes a: "a\<sharp>(\<theta>n,\<theta>c)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18621
  shows "\<theta>n,((a,x,P)#\<theta>c)<M> = ((\<theta>n,\<theta>c<M>){a:=(x).P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18622
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18623
apply(nominal_induct M avoiding: a x \<theta>n \<theta>c P rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18624
apply(auto simp add: fresh_list_cons fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18625
apply(simp add: lookup_csubst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18626
apply(simp add: fresh_list_cons fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18627
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18628
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18629
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18630
apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18631
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18632
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18633
apply(simp add: lookupd_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18634
apply(subgoal_tac "a\<sharp>lookupc xa coname \<theta>n")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18635
apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18636
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18637
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18638
apply(simp add: alpha nrename_swap fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18639
apply(rule lookupc_freshness)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18640
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18641
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18642
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18643
apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18644
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18645
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18646
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18647
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18648
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18649
apply(simp add: lookupd_unicity)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18650
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18651
apply(subgoal_tac "a\<sharp>lookupc xa coname \<theta>n")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18652
apply(subgoal_tac "a\<sharp>lookupd name aa \<theta>c")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18653
apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18654
apply(rule lookupd_freshness)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18655
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18656
apply(rule lookupc_freshness)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18657
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18658
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18659
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18660
apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18661
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18662
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18663
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18664
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18665
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18666
apply(drule ax_psubst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18667
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18668
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18669
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18670
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18671
apply(subgoal_tac "a\<sharp>lookupc xa coname \<theta>n")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18672
apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18673
apply(rule lookupc_freshness)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18674
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18675
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18676
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18677
apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18678
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18679
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18680
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18681
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18682
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18683
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18684
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18685
apply(simp add: alpha)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18686
apply(simp add: alpha nrename_swap fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18687
apply(simp add: lookupd_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18688
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18689
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18690
apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18691
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18692
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18693
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18694
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18695
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18696
apply(simp add: lookupd_unicity)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18697
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18698
apply(subgoal_tac "a\<sharp>lookupd name aa \<theta>c")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18699
apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18700
apply(rule lookupd_freshness)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18701
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18702
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18703
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18704
apply(rule better_Cut_substc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18705
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18706
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18707
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18708
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18709
apply(drule ax_psubst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18710
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18711
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18712
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18713
(* NotR *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18714
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18715
apply(case_tac "findc \<theta>c coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18716
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18717
apply(auto simp add: fresh_list_cons fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18718
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18719
apply(auto simp add: fresh_list_cons fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18720
apply(drule cmaps_false)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18721
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18722
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18723
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18724
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18725
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18726
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18727
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18728
apply(rule better_Cut_substc1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18729
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18730
apply(simp add: cmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18731
apply(auto simp add: fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18732
(* NotL *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18733
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18734
apply(case_tac "findn \<theta>n name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18735
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18736
apply(auto simp add: fresh_list_cons fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18737
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18738
apply(auto simp add: fresh_list_cons fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18739
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18740
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18741
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18742
apply(drule_tac a="a" in nmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18743
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18744
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18745
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18746
apply(rule better_Cut_substc2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18747
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18748
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18749
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18750
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18751
(* AndR *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18752
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18753
apply(case_tac "findc \<theta>c coname3")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18754
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18755
apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18756
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18757
apply(auto simp add: fresh_list_cons fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18758
apply(drule cmaps_false)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18759
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18760
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18761
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18762
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18763
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18764
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18765
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18766
apply(rule better_Cut_substc1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18767
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18768
apply(simp add: cmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18769
apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18770
(* AndL1 *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18771
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18772
apply(case_tac "findn \<theta>n name2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18773
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18774
apply(auto simp add: fresh_list_cons fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18775
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18776
apply(auto simp add: fresh_list_cons fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18777
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18778
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18779
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18780
apply(drule_tac a="a" in nmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18781
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18782
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18783
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18784
apply(rule better_Cut_substc2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18785
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18786
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18787
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18788
apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18789
(* AndL2 *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18790
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18791
apply(case_tac "findn \<theta>n name2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18792
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18793
apply(auto simp add: fresh_list_cons fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18794
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18795
apply(auto simp add: fresh_list_cons fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18796
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18797
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18798
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18799
apply(drule_tac a="a" in nmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18800
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18801
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18802
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18803
apply(rule better_Cut_substc2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18804
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18805
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18806
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18807
apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18808
(* OrR1 *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18809
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18810
apply(case_tac "findc \<theta>c coname2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18811
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18812
apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18813
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18814
apply(auto simp add: fresh_list_cons fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18815
apply(drule cmaps_false)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18816
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18817
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18818
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18819
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18820
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18821
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18822
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18823
apply(rule better_Cut_substc1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18824
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18825
apply(simp add: cmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18826
apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18827
(* OrR2 *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18828
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18829
apply(case_tac "findc \<theta>c coname2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18830
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18831
apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18832
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18833
apply(auto simp add: fresh_list_cons fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18834
apply(drule cmaps_false)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18835
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18836
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18837
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18838
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18839
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18840
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18841
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18842
apply(rule better_Cut_substc1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18843
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18844
apply(simp add: cmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18845
apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18846
(* OrL *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18847
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18848
apply(case_tac "findn \<theta>n name3")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18849
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18850
apply(auto simp add: fresh_list_cons psubst_fresh_name fresh_atm fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18851
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18852
apply(auto simp add: fresh_list_cons fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18853
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18854
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18855
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18856
apply(drule_tac a="a" in nmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18857
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18858
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18859
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18860
apply(rule better_Cut_substc2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18861
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18862
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18863
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18864
apply(auto simp add:  psubst_fresh_name fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18865
(* ImpR *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18866
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18867
apply(case_tac "findc \<theta>c coname2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18868
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18869
apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18870
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18871
apply(auto simp add: fresh_list_cons fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18872
apply(drule cmaps_false)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18873
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18874
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18875
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18876
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18877
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18878
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18879
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18880
apply(rule better_Cut_substc1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18881
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18882
apply(simp add: cmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18883
apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18884
(* ImpL *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18885
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18886
apply(case_tac "findn \<theta>n name2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18887
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18888
apply(auto simp add: fresh_list_cons psubst_fresh_coname psubst_fresh_name fresh_atm fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18889
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18890
apply(auto simp add: fresh_list_cons fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18891
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18892
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18893
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18894
apply(simp add: abs_fresh subst_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18895
apply(drule_tac a="a" in nmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18896
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18897
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18898
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18899
apply(rule better_Cut_substc2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18900
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18901
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18902
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18903
apply(auto simp add: psubst_fresh_coname psubst_fresh_name fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18904
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18905
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18906
lemma psubst_nsubst:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18907
  assumes a: "x\<sharp>(\<theta>n,\<theta>c)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18908
  shows "((x,a,P)#\<theta>n),\<theta>c<M> = ((\<theta>n,\<theta>c<M>){x:=<a>.P})"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18909
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18910
apply(nominal_induct M avoiding: a x \<theta>n \<theta>c P rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18911
apply(auto simp add: fresh_list_cons fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18912
apply(simp add: lookup_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18913
apply(rule lookupb_lookupa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18914
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18915
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18916
apply(rule forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18917
apply(rule lookup_freshness)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18918
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18919
apply(auto simp add: lookupc_freshness fresh_list_cons fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18920
apply(simp add: lookupc_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18921
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18922
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18923
apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18924
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18925
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18926
apply(simp add: lookupd_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18927
apply(subgoal_tac "x\<sharp>lookupd name aa \<theta>c")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18928
apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18929
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18930
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18931
apply(simp add: alpha crename_swap fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18932
apply(rule lookupd_freshness)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18933
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18934
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18935
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18936
apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18937
apply(simp add: abs_fresh) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18938
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18939
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18940
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18941
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18942
apply(simp add: lookupc_unicity)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18943
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18944
apply(subgoal_tac "x\<sharp>lookupc xa coname \<theta>n")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18945
apply(subgoal_tac "x\<sharp>lookupd name aa \<theta>c")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18946
apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18947
apply(rule lookupd_freshness)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18948
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18949
apply(rule lookupc_freshness)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18950
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18951
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18952
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18953
apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18954
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18955
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18956
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18957
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18958
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18959
apply(simp add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18960
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18961
apply(simp add: alpha crename_swap fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18962
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18963
apply(simp add: lookupc_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18964
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18965
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18966
apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18967
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18968
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18969
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18970
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18971
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18972
apply(simp add: lookupc_unicity)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18973
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18974
apply(subgoal_tac "x\<sharp>lookupc xa coname \<theta>n")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18975
apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18976
apply(rule lookupc_freshness)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18977
apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18978
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18979
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18980
apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18981
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18982
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18983
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18984
apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18985
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18986
apply(drule ax_psubst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18987
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18988
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18989
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18990
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18991
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18992
apply(subgoal_tac "x\<sharp>lookupd name aa \<theta>c")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18993
apply(simp add: forget)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18994
apply(rule lookupd_freshness)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18995
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18996
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18997
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18998
apply(rule better_Cut_substn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 18999
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19000
apply(simp add: abs_fresh fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19001
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19002
apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19003
apply(drule ax_psubst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19004
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19005
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19006
apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19007
(* NotR *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19008
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19009
apply(case_tac "findc \<theta>c coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19010
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19011
apply(auto simp add: fresh_list_cons fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19012
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19013
apply(auto simp add: fresh_list_cons fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19014
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19015
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19016
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19017
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19018
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19019
apply(rule better_Cut_substn1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19020
apply(simp add: cmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19021
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19022
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19023
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19024
(* NotL *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19025
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19026
apply(case_tac "findn \<theta>n name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19027
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19028
apply(auto simp add: fresh_list_cons fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19029
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19030
apply(auto simp add: fresh_list_cons fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19031
apply(drule nmaps_false)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19032
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19033
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19034
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19035
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19036
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19037
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19038
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19039
apply(rule better_Cut_substn2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19040
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19041
apply(simp add: nmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19042
apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19043
(* AndR *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19044
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19045
apply(case_tac "findc \<theta>c coname3")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19046
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19047
apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19048
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19049
apply(auto simp add: fresh_list_cons fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19050
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19051
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19052
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19053
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19054
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19055
apply(rule better_Cut_substn1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19056
apply(simp add: cmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19057
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19058
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19059
apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19060
(* AndL1 *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19061
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19062
apply(case_tac "findn \<theta>n name2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19063
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19064
apply(auto simp add: fresh_list_cons fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19065
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19066
apply(auto simp add: fresh_list_cons fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19067
apply(drule nmaps_false)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19068
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19069
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19070
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19071
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19072
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19073
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19074
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19075
apply(rule better_Cut_substn2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19076
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19077
apply(simp add: nmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19078
apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19079
(* AndL2 *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19080
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19081
apply(case_tac "findn \<theta>n name2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19082
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19083
apply(auto simp add: fresh_list_cons fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19084
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19085
apply(auto simp add: fresh_list_cons fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19086
apply(drule nmaps_false)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19087
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19088
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19089
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19090
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19091
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19092
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19093
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19094
apply(rule better_Cut_substn2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19095
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19096
apply(simp add: nmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19097
apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19098
(* OrR1 *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19099
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19100
apply(case_tac "findc \<theta>c coname2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19101
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19102
apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19103
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19104
apply(auto simp add: fresh_list_cons fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19105
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19106
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19107
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19108
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19109
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19110
apply(rule better_Cut_substn1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19111
apply(simp add: cmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19112
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19113
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19114
apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19115
(* OrR2 *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19116
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19117
apply(case_tac "findc \<theta>c coname2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19118
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19119
apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19120
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19121
apply(auto simp add: fresh_list_cons fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19122
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19123
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19124
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19125
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19126
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19127
apply(rule better_Cut_substn1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19128
apply(simp add: cmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19129
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19130
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19131
apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19132
(* OrL *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19133
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19134
apply(case_tac "findn \<theta>n name3")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19135
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19136
apply(auto simp add: fresh_list_cons psubst_fresh_name fresh_atm fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19137
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19138
apply(auto simp add: fresh_list_cons fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19139
apply(drule nmaps_false)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19140
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19141
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19142
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19143
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19144
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19145
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19146
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19147
apply(rule better_Cut_substn2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19148
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19149
apply(simp add: nmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19150
apply(auto simp add:  psubst_fresh_name fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19151
(* ImpR *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19152
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19153
apply(case_tac "findc \<theta>c coname2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19154
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19155
apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19156
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19157
apply(auto simp add: fresh_list_cons fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19158
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19159
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19160
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19161
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19162
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19163
apply(rule better_Cut_substn1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19164
apply(simp add: cmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19165
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19166
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19167
apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19168
(* ImpL *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19169
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19170
apply(case_tac "findn \<theta>n name2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19171
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19172
apply(auto simp add: fresh_list_cons psubst_fresh_coname psubst_fresh_name fresh_atm fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19173
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19174
apply(auto simp add: fresh_list_cons fresh_prod)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19175
apply(drule nmaps_false)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19176
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19177
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19178
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19179
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19180
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19181
apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19182
apply(rule trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19183
apply(rule better_Cut_substn2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19184
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19185
apply(simp add: nmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19186
apply(auto simp add: psubst_fresh_coname psubst_fresh_name fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19187
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19188
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19189
definition 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19190
  ncloses :: "(name\<times>coname\<times>trm) list\<Rightarrow>(name\<times>ty) list \<Rightarrow> bool" ("_ ncloses _" [55,55] 55) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19191
where
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19192
  "\<theta>n ncloses \<Gamma> \<equiv> \<forall>x B. ((x,B) \<in> set \<Gamma> \<longrightarrow> (\<exists>c P. \<theta>n nmaps x to Some (c,P) \<and> <c>:P \<in> (\<parallel><B>\<parallel>)))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19193
  
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19194
definition 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19195
  ccloses :: "(coname\<times>name\<times>trm) list\<Rightarrow>(coname\<times>ty) list \<Rightarrow> bool" ("_ ccloses _" [55,55] 55) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19196
where
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19197
  "\<theta>c ccloses \<Delta> \<equiv> \<forall>a B. ((a,B) \<in> set \<Delta> \<longrightarrow> (\<exists>x P. \<theta>c cmaps a to Some (x,P) \<and> (x):P \<in> (\<parallel>(B)\<parallel>)))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19198
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19199
lemma ncloses_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19200
  assumes a: "(x,B) \<in> set \<Gamma>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19201
  and     b: "\<theta>n ncloses \<Gamma>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19202
  shows "\<exists>c P. \<theta>n nmaps x to Some (c,P) \<and> <c>:P \<in> (\<parallel><B>\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19203
using a b by (auto simp add: ncloses_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19204
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19205
lemma ccloses_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19206
  assumes a: "(a,B) \<in> set \<Delta>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19207
  and     b: "\<theta>c ccloses \<Delta>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19208
  shows "\<exists>x P. \<theta>c cmaps a to Some (x,P) \<and> (x):P \<in> (\<parallel>(B)\<parallel>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19209
using a b by (auto simp add: ccloses_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19210
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19211
lemma ncloses_subset:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19212
  assumes a: "\<theta>n ncloses \<Gamma>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19213
  and     b: "set \<Gamma>' \<subseteq> set \<Gamma>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19214
  shows "\<theta>n ncloses \<Gamma>'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19215
using a b by (auto  simp add: ncloses_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19216
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19217
lemma ccloses_subset:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19218
  assumes a: "\<theta>c ccloses \<Delta>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19219
  and     b: "set \<Delta>' \<subseteq> set \<Delta>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19220
  shows "\<theta>c ccloses \<Delta>'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19221
using a b by (auto  simp add: ccloses_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19222
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19223
lemma validc_fresh:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19224
  fixes a::"coname"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19225
  and   \<Delta>::"(coname\<times>ty) list"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19226
  assumes a: "a\<sharp>\<Delta>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19227
  shows "\<not>(\<exists>B. (a,B)\<in>set \<Delta>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19228
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19229
apply(induct \<Delta>)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19230
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19231
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19232
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19233
lemma validn_fresh:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19234
  fixes x::"name"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19235
  and   \<Gamma>::"(name\<times>ty) list"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19236
  assumes a: "x\<sharp>\<Gamma>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19237
  shows "\<not>(\<exists>B. (x,B)\<in>set \<Gamma>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19238
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19239
apply(induct \<Gamma>)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19240
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19241
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19242
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19243
lemma ccloses_extend:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19244
  assumes a: "\<theta>c ccloses \<Delta>" "a\<sharp>\<Delta>" "a\<sharp>\<theta>c" "(x):P\<in>\<parallel>(B)\<parallel>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19245
  shows "(a,x,P)#\<theta>c ccloses (a,B)#\<Delta>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19246
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19247
apply(simp add: ccloses_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19248
apply(drule validc_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19249
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19250
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19251
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19252
lemma ncloses_extend:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19253
  assumes a: "\<theta>n ncloses \<Gamma>" "x\<sharp>\<Gamma>" "x\<sharp>\<theta>n" "<a>:P\<in>\<parallel><B>\<parallel>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19254
  shows "(x,a,P)#\<theta>n ncloses (x,B)#\<Gamma>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19255
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19256
apply(simp add: ncloses_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19257
apply(drule validn_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19258
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19259
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19260
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19261
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
 19262
text {* typing relation *}
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
 19263
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
 19264
inductive2
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
 19265
   typing :: "ctxtn \<Rightarrow> trm \<Rightarrow> ctxtc \<Rightarrow> bool" ("_ \<turnstile> _ \<turnstile> _" [100,100,100] 100)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22232
diff changeset
 19266
where
23092
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19267
  TAx:    "\<lbrakk>validn \<Gamma>;validc \<Delta>; (x,B)\<in>set \<Gamma>; (a,B)\<in>set \<Delta>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Ax x a \<turnstile> \<Delta>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19268
| TNotR:  "\<lbrakk>x\<sharp>\<Gamma>; ((x,B)#\<Gamma>) \<turnstile> M \<turnstile> \<Delta>; set \<Delta>' = {(a,NOT B)}\<union>set \<Delta>; validc \<Delta>'\<rbrakk> 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19269
           \<Longrightarrow> \<Gamma> \<turnstile> NotR (x).M a \<turnstile> \<Delta>'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19270
| TNotL:  "\<lbrakk>a\<sharp>\<Delta>; \<Gamma> \<turnstile> M \<turnstile> ((a,B)#\<Delta>); set \<Gamma>' = {(x,NOT B)} \<union> set \<Gamma>; validn \<Gamma>'\<rbrakk>  
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19271
           \<Longrightarrow> \<Gamma>' \<turnstile> NotL <a>.M x \<turnstile> \<Delta>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19272
| TAndL1: "\<lbrakk>x\<sharp>(\<Gamma>,y); ((x,B1)#\<Gamma>) \<turnstile> M \<turnstile> \<Delta>; set \<Gamma>' = {(y,B1 AND B2)} \<union> set \<Gamma>; validn \<Gamma>'\<rbrakk> 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19273
           \<Longrightarrow> \<Gamma>' \<turnstile> AndL1 (x).M y \<turnstile> \<Delta>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19274
| TAndL2: "\<lbrakk>x\<sharp>(\<Gamma>,y); ((x,B2)#\<Gamma>) \<turnstile> M \<turnstile> \<Delta>; set \<Gamma>' = {(y,B1 AND B2)} \<union> set \<Gamma>; validn \<Gamma>'\<rbrakk> 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19275
           \<Longrightarrow> \<Gamma>' \<turnstile> AndL2 (x).M y \<turnstile> \<Delta>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19276
| TAndR:  "\<lbrakk>a\<sharp>(\<Delta>,N,c); b\<sharp>(\<Delta>,M,c); a\<noteq>b; \<Gamma> \<turnstile> M \<turnstile> ((a,B)#\<Delta>); \<Gamma> \<turnstile> N \<turnstile> ((b,C)#\<Delta>); 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19277
           set \<Delta>' = {(c,B AND C)}\<union>set \<Delta>; validc \<Delta>'\<rbrakk> 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19278
           \<Longrightarrow> \<Gamma> \<turnstile> AndR <a>.M <b>.N c \<turnstile> \<Delta>'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19279
| TOrL:   "\<lbrakk>x\<sharp>(\<Gamma>,N,z); y\<sharp>(\<Gamma>,M,z); x\<noteq>y; ((x,B)#\<Gamma>) \<turnstile> M \<turnstile> \<Delta>; ((y,C)#\<Gamma>) \<turnstile> N \<turnstile> \<Delta>;
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19280
           set \<Gamma>' = {(z,B OR C)} \<union> set \<Gamma>; validn \<Gamma>'\<rbrakk> 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19281
           \<Longrightarrow> \<Gamma>' \<turnstile> OrL (x).M (y).N z \<turnstile> \<Delta>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19282
| TOrR1:  "\<lbrakk>a\<sharp>(\<Delta>,b); \<Gamma> \<turnstile> M \<turnstile> ((a,B1)#\<Delta>); set \<Delta>' = {(b,B1 OR B2)}\<union>set \<Delta>; validc \<Delta>'\<rbrakk> 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19283
           \<Longrightarrow> \<Gamma> \<turnstile> OrR1 <a>.M b \<turnstile> \<Delta>'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19284
| TOrR2:  "\<lbrakk>a\<sharp>(\<Delta>,b); \<Gamma> \<turnstile> M \<turnstile> ((a,B2)#\<Delta>); set \<Delta>' = {(b,B1 OR B2)}\<union>set \<Delta>; validc \<Delta>'\<rbrakk> 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19285
           \<Longrightarrow> \<Gamma> \<turnstile> OrR2 <a>.M b \<turnstile> \<Delta>'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19286
| TImpL:  "\<lbrakk>a\<sharp>(\<Delta>,N); x\<sharp>(\<Gamma>,M,y); \<Gamma> \<turnstile> M \<turnstile> ((a,B)#\<Delta>); ((x,C)#\<Gamma>) \<turnstile> N \<turnstile> \<Delta>;
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19287
           set \<Gamma>' = {(y,B IMP C)} \<union> set \<Gamma>; validn \<Gamma>'\<rbrakk> 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19288
           \<Longrightarrow> \<Gamma>' \<turnstile> ImpL <a>.M (x).N y \<turnstile> \<Delta>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19289
| TImpR:  "\<lbrakk>a\<sharp>(\<Delta>,b); x\<sharp>\<Gamma>; ((x,B)#\<Gamma>) \<turnstile> M \<turnstile> ((a,C)#\<Delta>); set \<Delta>' = {(b,B IMP C)}\<union>set \<Delta>; validc \<Delta>'\<rbrakk> 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19290
           \<Longrightarrow> \<Gamma> \<turnstile> ImpR (x).<a>.M b \<turnstile> \<Delta>'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19291
| TCut:   "\<lbrakk>a\<sharp>(\<Delta>,N); x\<sharp>(\<Gamma>,M); \<Gamma> \<turnstile> M \<turnstile> ((a,B)#\<Delta>); ((x,B)#\<Gamma>) \<turnstile> N \<turnstile> \<Delta>\<rbrakk> 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19292
           \<Longrightarrow> \<Gamma> \<turnstile> Cut <a>.M (x).N \<turnstile> \<Delta>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19293
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19294
equivariance typing
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19295
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19296
lemma fresh_set_member:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19297
  fixes x::"name"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19298
  and   a::"coname"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19299
  shows "x\<sharp>L \<Longrightarrow> e\<in>set L \<Longrightarrow> x\<sharp>e"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19300
  and   "a\<sharp>L \<Longrightarrow> e\<in>set L \<Longrightarrow> a\<sharp>e"   
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19301
by (induct L) (auto simp add: fresh_list_cons) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19302
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19303
lemma fresh_subset:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19304
  fixes x::"name"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19305
  and   a::"coname"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19306
  shows "x\<sharp>L \<Longrightarrow> set L' \<subseteq> set L \<Longrightarrow> x\<sharp>L'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19307
  and   "a\<sharp>L \<Longrightarrow> set L' \<subseteq> set L \<Longrightarrow> a\<sharp>L'"   
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19308
apply(induct L' arbitrary: L) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19309
apply(auto simp add: fresh_list_cons fresh_list_nil intro: fresh_set_member)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19310
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19311
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19312
lemma fresh_subset_ext:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19313
  fixes x::"name"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19314
  and   a::"coname"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19315
  shows "x\<sharp>L \<Longrightarrow> x\<sharp>e \<Longrightarrow> set L' \<subseteq> set (e#L) \<Longrightarrow> x\<sharp>L'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19316
  and   "a\<sharp>L \<Longrightarrow> a\<sharp>e \<Longrightarrow> set L' \<subseteq> set (e#L) \<Longrightarrow> a\<sharp>L'"   
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19317
apply(induct L' arbitrary: L) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19318
apply(auto simp add: fresh_list_cons fresh_list_nil intro: fresh_set_member)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19319
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19320
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19321
lemma fresh_under_insert:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19322
  fixes x::"name"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19323
  and   a::"coname"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19324
  and   \<Gamma>::"ctxtn"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19325
  and   \<Delta>::"ctxtc"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19326
  shows "x\<sharp>\<Gamma> \<Longrightarrow> x\<noteq>y \<Longrightarrow> set \<Gamma>' = insert (y,B) (set \<Gamma>) \<Longrightarrow> x\<sharp>\<Gamma>'"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19327
  and   "a\<sharp>\<Delta> \<Longrightarrow> a\<noteq>c \<Longrightarrow> set \<Delta>' = insert (c,B) (set \<Delta>) \<Longrightarrow> a\<sharp>\<Delta>'"   
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19328
apply(rule fresh_subset_ext(1))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19329
apply(auto simp add: fresh_prod fresh_atm fresh_ty)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19330
apply(rule fresh_subset_ext(2))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19331
apply(auto simp add: fresh_prod fresh_atm fresh_ty)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19332
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19333
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19334
nominal_inductive typing
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19335
  apply (simp_all add: abs_fresh fresh_atm fresh_list_cons fresh_prod fresh_ty fresh_ctxt 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19336
                       fresh_list_append abs_supp fin_supp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19337
  apply(auto intro: fresh_under_insert)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19338
  done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19339
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19340
lemma validn_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19341
  assumes a: "validn ((x,B)#\<Gamma>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19342
  shows "validn \<Gamma> \<and> x\<sharp>\<Gamma>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19343
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19344
apply(erule_tac validn.cases)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19345
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19346
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19347
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19348
lemma validc_elim:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19349
  assumes a: "validc ((a,B)#\<Delta>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19350
  shows "validc \<Delta> \<and> a\<sharp>\<Delta>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19351
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19352
apply(erule_tac validc.cases)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19353
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19354
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19355
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19356
lemma context_fresh:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19357
  fixes x::"name"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19358
  and   a::"coname"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19359
  shows "x\<sharp>\<Gamma> \<Longrightarrow> \<not>(\<exists>B. (x,B)\<in>set \<Gamma>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19360
  and   "a\<sharp>\<Delta> \<Longrightarrow> \<not>(\<exists>B. (a,B)\<in>set \<Delta>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19361
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19362
apply(induct \<Gamma>)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19363
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19364
apply(induct \<Delta>)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19365
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19366
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19367
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19368
lemma typing_implies_valid:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19369
  assumes a: "\<Gamma> \<turnstile> M \<turnstile> \<Delta>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19370
  shows "validn \<Gamma> \<and> validc \<Delta>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19371
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19372
apply(nominal_induct rule: typing.strong_induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19373
apply(auto dest: validn_elim validc_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19374
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19375
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19376
lemma ty_perm:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19377
  fixes pi1::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19378
  and   pi2::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19379
  and   B::"ty"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19380
  shows "pi1\<bullet>B=B" and "pi2\<bullet>B=B"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19381
apply(nominal_induct B rule: ty.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19382
apply(auto simp add: perm_string)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19383
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19384
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19385
lemma ctxt_perm:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19386
  fixes pi1::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19387
  and   pi2::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19388
  and   \<Gamma>::"ctxtn"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19389
  and   \<Delta>::"ctxtc"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19390
  shows "pi2\<bullet>\<Gamma>=\<Gamma>" and "pi1\<bullet>\<Delta>=\<Delta>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19391
apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19392
apply(induct \<Gamma>)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19393
apply(auto simp add: calc_atm ty_perm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19394
apply(induct \<Delta>)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19395
apply(auto simp add: calc_atm ty_perm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19396
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19397
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19398
lemma typing_Ax_elim1: 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19399
  assumes a: "\<Gamma> \<turnstile> Ax x a \<turnstile> ((a,B)#\<Delta>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19400
  shows "(x,B)\<in>set \<Gamma>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19401
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19402
apply(erule_tac typing.cases)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19403
apply(simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19404
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19405
apply(auto dest: validc_elim context_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19406
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19407
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19408
lemma typing_Ax_elim2: 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19409
  assumes a: "((x,B)#\<Gamma>) \<turnstile> Ax x a \<turnstile> \<Delta>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19410
  shows "(a,B)\<in>set \<Delta>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19411
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19412
apply(erule_tac typing.cases)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19413
apply(simp_all add: trm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19414
apply(auto  dest!: validn_elim context_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19415
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19416
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19417
lemma psubst_Ax_aux: 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19418
  assumes a: "\<theta>c cmaps a to Some (y,N)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19419
  shows "lookupb x a \<theta>c c P = Cut <c>.P (y).N"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19420
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19421
apply(induct \<theta>c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19422
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19423
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19424
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19425
lemma psubst_Ax:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19426
  assumes a: "\<theta>n nmaps x to Some (c,P)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19427
  and     b: "\<theta>c cmaps a to Some (y,N)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19428
  shows "\<theta>n,\<theta>c<Ax x a> = Cut <c>.P (y).N"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19429
using a b
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19430
apply(induct \<theta>n)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19431
apply(auto simp add: psubst_Ax_aux)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19432
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19433
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19434
lemma psubst_Cut:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19435
  assumes a: "\<forall>x. M\<noteq>Ax x c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19436
  and     b: "\<forall>a. N\<noteq>Ax x a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19437
  and     c: "c\<sharp>(\<theta>n,\<theta>c,N)" "x\<sharp>(\<theta>n,\<theta>c,M)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19438
  shows "\<theta>n,\<theta>c<Cut <c>.M (x).N> = Cut <c>.(\<theta>n,\<theta>c<M>) (x).(\<theta>n,\<theta>c<N>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19439
using a b c
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19440
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19441
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19442
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19443
lemma all_CAND: 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19444
  assumes a: "\<Gamma> \<turnstile> M \<turnstile> \<Delta>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19445
  and     b: "\<theta>n ncloses \<Gamma>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19446
  and     c: "\<theta>c ccloses \<Delta>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19447
  shows "SNa (\<theta>n,\<theta>c<M>)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19448
using a b c
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19449
proof(nominal_induct avoiding: \<theta>n \<theta>c rule: typing.strong_induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19450
  case (TAx \<Gamma> \<Delta> x B a \<theta>n \<theta>c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19451
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19452
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19453
    apply(drule ncloses_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19454
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19455
    apply(drule ccloses_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19456
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19457
    apply(erule exE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19458
    apply(erule conjE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19459
    apply(rule_tac s="Cut <c>.P (xa).Pa" and t="\<theta>n,\<theta>c<Ax x a>" in subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19460
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19461
    apply(simp only: psubst_Ax)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19462
    apply(simp add: CUT_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19463
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19464
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19465
  case (TNotR x \<Gamma> B M \<Delta> \<Delta>' a \<theta>n \<theta>c) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19466
  then show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19467
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19468
    apply(subgoal_tac "(a,NOT B) \<in> set \<Delta>'")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19469
    apply(drule ccloses_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19470
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19471
    apply(erule exE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19472
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19473
    apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19474
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19475
    apply(rule_tac B="NOT B" in CUT_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19476
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19477
    apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19478
    apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19479
    apply(rule_tac x="c" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19480
    apply(rule_tac x="x" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19481
    apply(rule_tac x="\<theta>n,\<theta>c<M>" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19482
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19483
    apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19484
    apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19485
    apply(rule psubst_fresh_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19486
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19487
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19488
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19489
    apply(rule BINDING_implies_CAND)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19490
    apply(unfold BINDINGn_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19491
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19492
    apply(rule_tac x="x" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19493
    apply(rule_tac x="\<theta>n,\<theta>c<M>" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19494
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19495
    apply(rule allI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19496
    apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19497
    apply(simp add: psubst_nsubst[symmetric])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19498
    apply(drule_tac x="(x,aa,Pa)#\<theta>n" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19499
    apply(drule_tac x="\<theta>c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19500
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19501
    apply(rule ncloses_extend)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19502
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19503
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19504
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19505
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19506
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19507
    apply(rule ccloses_subset)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19508
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19509
    apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19510
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19511
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19512
    apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19513
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19514
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19515
  case (TNotL a \<Delta> \<Gamma> M B \<Gamma>' x \<theta>n \<theta>c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19516
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19517
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19518
    apply(subgoal_tac "(x,NOT B) \<in> set \<Gamma>'")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19519
    apply(drule ncloses_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19520
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19521
    apply(erule exE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19522
    apply(simp del: NEGc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19523
    apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19524
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19525
    apply(rule_tac B="NOT B" in CUT_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19526
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19527
    apply(rule NEG_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19528
    apply(simp (no_asm))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19529
    apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19530
    apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19531
    apply(rule_tac x="a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19532
    apply(rule_tac x="ca" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19533
    apply(rule_tac x="\<theta>n,\<theta>c<M>" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19534
    apply(simp del: NEGc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19535
    apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19536
    apply(rule fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19537
    apply(rule psubst_fresh_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19538
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19539
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19540
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19541
    apply(rule BINDING_implies_CAND)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19542
    apply(unfold BINDINGc_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19543
    apply(simp (no_asm))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19544
    apply(rule_tac x="a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19545
    apply(rule_tac x="\<theta>n,\<theta>c<M>" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19546
    apply(simp (no_asm))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19547
    apply(rule allI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19548
    apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19549
    apply(simp del: NEGc.simps add: psubst_csubst[symmetric])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19550
    apply(drule_tac x="\<theta>n" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19551
    apply(drule_tac x="(a,xa,Pa)#\<theta>c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19552
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19553
    apply(rule ncloses_subset)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19554
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19555
    apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19556
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19557
    apply(rule ccloses_extend)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19558
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19559
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19560
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19561
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19562
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19563
    apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19564
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19565
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19566
  case (TAndL1 x \<Gamma> y B1 M \<Delta> \<Gamma>' B2 \<theta>n \<theta>c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19567
  then show ?case     
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19568
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19569
    apply(subgoal_tac "(y,B1 AND B2) \<in> set \<Gamma>'")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19570
    apply(drule ncloses_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19571
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19572
    apply(erule exE)+ 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19573
    apply(simp del: NEGc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19574
    apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19575
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19576
    apply(rule_tac B="B1 AND B2" in CUT_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19577
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19578
    apply(rule NEG_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19579
    apply(simp (no_asm))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19580
    apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19581
    apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19582
    apply(rule disjI1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19583
    apply(rule_tac x="x" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19584
    apply(rule_tac x="ca" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19585
    apply(rule_tac x="\<theta>n,\<theta>c<M>" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19586
    apply(simp del: NEGc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19587
    apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19588
    apply(rule fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19589
    apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19590
    apply(rule psubst_fresh_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19591
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19592
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19593
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19594
    apply(rule BINDING_implies_CAND)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19595
    apply(unfold BINDINGn_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19596
    apply(simp (no_asm))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19597
    apply(rule_tac x="x" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19598
    apply(rule_tac x="\<theta>n,\<theta>c<M>" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19599
    apply(simp (no_asm))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19600
    apply(rule allI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19601
    apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19602
    apply(simp del: NEGc.simps add: psubst_nsubst[symmetric])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19603
    apply(drule_tac x="(x,a,Pa)#\<theta>n" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19604
    apply(drule_tac x="\<theta>c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19605
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19606
    apply(rule ncloses_extend)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19607
    apply(rule ncloses_subset)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19608
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19609
    apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19610
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19611
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19612
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19613
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19614
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19615
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19616
    apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19617
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19618
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19619
  case (TAndL2 x \<Gamma> y B2 M \<Delta> \<Gamma>' B1 \<theta>n \<theta>c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19620
  then show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19621
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19622
    apply(subgoal_tac "(y,B1 AND B2) \<in> set \<Gamma>'")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19623
    apply(drule ncloses_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19624
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19625
    apply(erule exE)+ 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19626
    apply(simp del: NEGc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19627
    apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19628
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19629
    apply(rule_tac B="B1 AND B2" in CUT_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19630
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19631
    apply(rule NEG_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19632
    apply(simp (no_asm))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19633
    apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19634
    apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19635
    apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19636
    apply(rule_tac x="x" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19637
    apply(rule_tac x="ca" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19638
    apply(rule_tac x="\<theta>n,\<theta>c<M>" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19639
    apply(simp del: NEGc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19640
    apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19641
    apply(rule fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19642
    apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19643
    apply(rule psubst_fresh_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19644
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19645
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19646
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19647
    apply(rule BINDING_implies_CAND)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19648
    apply(unfold BINDINGn_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19649
    apply(simp (no_asm))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19650
    apply(rule_tac x="x" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19651
    apply(rule_tac x="\<theta>n,\<theta>c<M>" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19652
    apply(simp (no_asm))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19653
    apply(rule allI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19654
    apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19655
    apply(simp del: NEGc.simps add: psubst_nsubst[symmetric])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19656
    apply(drule_tac x="(x,a,Pa)#\<theta>n" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19657
    apply(drule_tac x="\<theta>c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19658
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19659
    apply(rule ncloses_extend)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19660
    apply(rule ncloses_subset)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19661
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19662
    apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19663
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19664
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19665
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19666
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19667
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19668
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19669
    apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19670
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19671
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19672
  case (TAndR a \<Delta> N c b M \<Gamma> B C \<Delta>' \<theta>n \<theta>c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19673
  then show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19674
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19675
    apply(subgoal_tac "(c,B AND C) \<in> set \<Delta>'")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19676
    apply(drule ccloses_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19677
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19678
    apply(erule exE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19679
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19680
    apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19681
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19682
    apply(rule_tac B="B AND C" in CUT_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19683
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19684
    apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19685
    apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19686
    apply(rule_tac x="ca" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19687
    apply(rule_tac x="a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19688
    apply(rule_tac x="b" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19689
    apply(rule_tac x="\<theta>n,\<theta>c<M>" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19690
    apply(rule_tac x="\<theta>n,\<theta>c<N>" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19691
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19692
    apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19693
    apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19694
    apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19695
    apply(rule psubst_fresh_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19696
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19697
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19698
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19699
    apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19700
    apply(rule psubst_fresh_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19701
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19702
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19703
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19704
    apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19705
    apply(rule BINDING_implies_CAND)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19706
    apply(unfold BINDINGc_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19707
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19708
    apply(rule_tac x="a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19709
    apply(rule_tac x="\<theta>n,\<theta>c<M>" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19710
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19711
    apply(rule allI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19712
    apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19713
    apply(simp add: psubst_csubst[symmetric])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19714
    apply(drule_tac x="\<theta>n" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19715
    apply(drule_tac x="(a,xa,Pa)#\<theta>c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19716
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19717
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19718
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19719
    apply(rule ccloses_extend)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19720
    apply(rule ccloses_subset)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19721
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19722
    apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19723
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19724
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19725
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19726
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19727
    apply(rule BINDING_implies_CAND)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19728
    apply(unfold BINDINGc_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19729
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19730
    apply(rule_tac x="b" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19731
    apply(rule_tac x="\<theta>n,\<theta>c<N>" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19732
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19733
    apply(rule allI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19734
    apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19735
    apply(simp add: psubst_csubst[symmetric])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19736
    apply(rotate_tac 14)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19737
    apply(drule_tac x="\<theta>n" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19738
    apply(drule_tac x="(b,xa,Pa)#\<theta>c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19739
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19740
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19741
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19742
    apply(rule ccloses_extend)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19743
    apply(rule ccloses_subset)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19744
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19745
    apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19746
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19747
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19748
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19749
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19750
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19751
    apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19752
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19753
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19754
  case (TOrL x \<Gamma> N z y M B \<Delta> C \<Gamma>' \<theta>n \<theta>c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19755
  then show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19756
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19757
    apply(subgoal_tac "(z,B OR C) \<in> set \<Gamma>'")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19758
    apply(drule ncloses_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19759
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19760
    apply(erule exE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19761
    apply(simp del: NEGc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19762
    apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19763
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19764
    apply(rule_tac B="B OR C" in CUT_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19765
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19766
    apply(rule NEG_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19767
    apply(simp (no_asm))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19768
    apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19769
    apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19770
    apply(rule_tac x="x" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19771
    apply(rule_tac x="y" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19772
    apply(rule_tac x="ca" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19773
    apply(rule_tac x="\<theta>n,\<theta>c<M>" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19774
    apply(rule_tac x="\<theta>n,\<theta>c<N>" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19775
    apply(simp del: NEGc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19776
    apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19777
    apply(rule fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19778
    apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19779
    apply(rule psubst_fresh_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19780
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19781
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19782
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19783
    apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19784
    apply(rule psubst_fresh_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19785
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19786
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19787
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19788
    apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19789
    apply(rule BINDING_implies_CAND)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19790
    apply(unfold BINDINGn_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19791
    apply(simp del: NEGc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19792
    apply(rule_tac x="x" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19793
    apply(rule_tac x="\<theta>n,\<theta>c<M>" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19794
    apply(simp del: NEGc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19795
    apply(rule allI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19796
    apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19797
    apply(simp del: NEGc.simps add: psubst_nsubst[symmetric])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19798
    apply(drule_tac x="(x,a,Pa)#\<theta>n" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19799
    apply(drule_tac x="\<theta>c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19800
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19801
    apply(rule ncloses_extend)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19802
    apply(rule ncloses_subset)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19803
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19804
    apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19805
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19806
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19807
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19808
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19809
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19810
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19811
    apply(rule BINDING_implies_CAND)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19812
    apply(unfold BINDINGn_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19813
    apply(simp del: NEGc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19814
    apply(rule_tac x="y" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19815
    apply(rule_tac x="\<theta>n,\<theta>c<N>" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19816
    apply(simp del: NEGc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19817
    apply(rule allI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19818
    apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19819
    apply(simp del: NEGc.simps add: psubst_nsubst[symmetric])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19820
    apply(rotate_tac 14)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19821
    apply(drule_tac x="(y,a,Pa)#\<theta>n" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19822
    apply(drule_tac x="\<theta>c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19823
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19824
    apply(rule ncloses_extend)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19825
    apply(rule ncloses_subset)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19826
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19827
    apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19828
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19829
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19830
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19831
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19832
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19833
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19834
    apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19835
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19836
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19837
  case (TOrR1 a \<Delta> b \<Gamma> M B1 \<Delta>' B2 \<theta>n \<theta>c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19838
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19839
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19840
    apply(subgoal_tac "(b,B1 OR B2) \<in> set \<Delta>'")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19841
    apply(drule ccloses_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19842
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19843
    apply(erule exE)+ 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19844
    apply(simp del: NEGc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19845
    apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19846
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19847
    apply(rule_tac B="B1 OR B2" in CUT_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19848
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19849
    apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19850
    apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19851
    apply(rule disjI1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19852
    apply(rule_tac x="a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19853
    apply(rule_tac x="c" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19854
    apply(rule_tac x="\<theta>n,\<theta>c<M>" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19855
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19856
    apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19857
    apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19858
    apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19859
    apply(rule psubst_fresh_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19860
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19861
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19862
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19863
    apply(rule BINDING_implies_CAND)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19864
    apply(unfold BINDINGc_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19865
    apply(simp (no_asm))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19866
    apply(rule_tac x="a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19867
    apply(rule_tac x="\<theta>n,\<theta>c<M>" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19868
    apply(simp (no_asm))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19869
    apply(rule allI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19870
    apply(rule impI)    
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19871
    apply(simp del: NEGc.simps add: psubst_csubst[symmetric])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19872
    apply(drule_tac x="\<theta>n" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19873
    apply(drule_tac x="(a,xa,Pa)#\<theta>c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19874
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19875
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19876
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19877
    apply(rule ccloses_extend)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19878
    apply(rule ccloses_subset)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19879
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19880
    apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19881
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19882
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19883
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19884
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19885
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19886
    apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19887
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19888
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19889
  case (TOrR2 a \<Delta> b \<Gamma> M B2 \<Delta>' B1 \<theta>n \<theta>c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19890
  then show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19891
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19892
    apply(subgoal_tac "(b,B1 OR B2) \<in> set \<Delta>'")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19893
    apply(drule ccloses_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19894
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19895
    apply(erule exE)+ 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19896
    apply(simp del: NEGc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19897
    apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19898
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19899
    apply(rule_tac B="B1 OR B2" in CUT_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19900
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19901
    apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19902
    apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19903
    apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19904
    apply(rule_tac x="a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19905
    apply(rule_tac x="c" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19906
    apply(rule_tac x="\<theta>n,\<theta>c<M>" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19907
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19908
    apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19909
    apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19910
    apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19911
    apply(rule psubst_fresh_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19912
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19913
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19914
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19915
    apply(rule BINDING_implies_CAND)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19916
    apply(unfold BINDINGc_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19917
    apply(simp (no_asm))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19918
    apply(rule_tac x="a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19919
    apply(rule_tac x="\<theta>n,\<theta>c<M>" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19920
    apply(simp (no_asm))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19921
    apply(rule allI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19922
    apply(rule impI)    
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19923
    apply(simp del: NEGc.simps add: psubst_csubst[symmetric])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19924
    apply(drule_tac x="\<theta>n" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19925
    apply(drule_tac x="(a,xa,Pa)#\<theta>c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19926
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19927
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19928
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19929
    apply(rule ccloses_extend)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19930
    apply(rule ccloses_subset)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19931
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19932
    apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19933
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19934
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19935
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19936
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19937
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19938
    apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19939
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19940
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19941
  case (TImpL a \<Delta> N x \<Gamma> M y B C \<Gamma>' \<theta>n \<theta>c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19942
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19943
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19944
    apply(subgoal_tac "(y,B IMP C) \<in> set \<Gamma>'")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19945
    apply(drule ncloses_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19946
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19947
    apply(erule exE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19948
    apply(simp del: NEGc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19949
    apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19950
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19951
    apply(rule_tac B="B IMP C" in CUT_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19952
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19953
    apply(rule NEG_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19954
    apply(simp (no_asm))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19955
    apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19956
    apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19957
    apply(rule_tac x="x" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19958
    apply(rule_tac x="a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19959
    apply(rule_tac x="ca" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19960
    apply(rule_tac x="\<theta>n,\<theta>c<M>" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19961
    apply(rule_tac x="\<theta>n,\<theta>c<N>" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19962
    apply(simp del: NEGc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19963
    apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19964
    apply(rule fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19965
    apply(rule psubst_fresh_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19966
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19967
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19968
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19969
    apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19970
    apply(rule psubst_fresh_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19971
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19972
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19973
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19974
    apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19975
    apply(rule BINDING_implies_CAND)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19976
    apply(unfold BINDINGc_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19977
    apply(simp del: NEGc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19978
    apply(rule_tac x="a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19979
    apply(rule_tac x="\<theta>n,\<theta>c<M>" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19980
    apply(simp del: NEGc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19981
    apply(rule allI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19982
    apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19983
    apply(simp del: NEGc.simps add: psubst_csubst[symmetric])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19984
    apply(drule_tac x="\<theta>n" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19985
    apply(drule_tac x="(a,xa,Pa)#\<theta>c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19986
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19987
    apply(rule ncloses_subset)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19988
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19989
    apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19990
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19991
    apply(rule ccloses_extend)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19992
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19993
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19994
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19995
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19996
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19997
    apply(rule BINDING_implies_CAND)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19998
    apply(unfold BINDINGn_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 19999
    apply(simp del: NEGc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20000
    apply(rule_tac x="x" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20001
    apply(rule_tac x="\<theta>n,\<theta>c<N>" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20002
    apply(simp del: NEGc.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20003
    apply(rule allI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20004
    apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20005
    apply(simp del: NEGc.simps add: psubst_nsubst[symmetric])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20006
    apply(rotate_tac 12)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20007
    apply(drule_tac x="(x,aa,Pa)#\<theta>n" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20008
    apply(drule_tac x="\<theta>c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20009
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20010
    apply(rule ncloses_extend)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20011
    apply(rule ncloses_subset)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20012
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20013
    apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20014
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20015
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20016
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20017
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20018
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20019
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20020
    apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20021
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20022
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20023
  case (TImpR a \<Delta> b x \<Gamma> B M C \<Delta>' \<theta>n \<theta>c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20024
  then show ?case
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20025
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20026
    apply(subgoal_tac "(b,B IMP C) \<in> set \<Delta>'")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20027
    apply(drule ccloses_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20028
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20029
    apply(erule exE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20030
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20031
    apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20032
    apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20033
    apply(rule_tac B="B IMP C" in CUT_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20034
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20035
    apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20036
    apply(rule disjI2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20037
    apply(rule_tac x="x" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20038
    apply(rule_tac x="a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20039
    apply(rule_tac x="c" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20040
    apply(rule_tac x="\<theta>n,\<theta>c<M>" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20041
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20042
    apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20043
    apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20044
    apply(simp add: abs_fresh fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20045
    apply(rule psubst_fresh_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20046
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20047
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20048
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20049
    apply(rule conjI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20050
    apply(rule allI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20051
    apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20052
    apply(simp add: psubst_csubst[symmetric])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20053
    apply(rule BINDING_implies_CAND)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20054
    apply(unfold BINDINGn_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20055
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20056
    apply(rule_tac x="x" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20057
    apply(rule_tac x="\<theta>n,((a,z,Pa)#\<theta>c)<M>" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20058
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20059
    apply(rule allI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20060
    apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20061
    apply(rule_tac t="\<theta>n,((a,z,Pa)#\<theta>c)<M>{x:=<aa>.Pb}" and 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20062
                   s="((x,aa,Pb)#\<theta>n),((a,z,Pa)#\<theta>c)<M>" in subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20063
    apply(rule psubst_nsubst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20064
    apply(simp add: fresh_prod fresh_atm fresh_list_cons)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20065
    apply(drule_tac x="(x,aa,Pb)#\<theta>n" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20066
    apply(drule_tac x="(a,z,Pa)#\<theta>c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20067
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20068
    apply(rule ncloses_extend)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20069
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20070
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20071
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20072
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20073
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20074
    apply(rule ccloses_extend)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20075
    apply(rule ccloses_subset)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20076
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20077
    apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20078
    apply(auto intro: fresh_subset simp del: NEGc.simps)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20079
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20080
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20081
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20082
    apply(rule allI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20083
    apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20084
    apply(simp add: psubst_nsubst[symmetric])
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20085
    apply(rule BINDING_implies_CAND)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20086
    apply(unfold BINDINGc_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20087
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20088
    apply(rule_tac x="a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20089
    apply(rule_tac x="((x,ca,Q)#\<theta>n),\<theta>c<M>" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20090
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20091
    apply(rule allI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20092
    apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20093
    apply(rule_tac t="((x,ca,Q)#\<theta>n),\<theta>c<M>{a:=(xaa).Pa}" and 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20094
                   s="((x,ca,Q)#\<theta>n),((a,xaa,Pa)#\<theta>c)<M>" in subst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20095
    apply(rule psubst_csubst)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20096
    apply(simp add: fresh_prod fresh_atm fresh_list_cons)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20097
    apply(drule_tac x="(x,ca,Q)#\<theta>n" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20098
    apply(drule_tac x="(a,xaa,Pa)#\<theta>c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20099
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20100
    apply(rule ncloses_extend)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20101
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20102
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20103
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20104
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20105
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20106
    apply(rule ccloses_extend)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20107
    apply(rule ccloses_subset)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20108
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20109
    apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20110
    apply(auto intro: fresh_subset simp del: NEGc.simps)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20111
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20112
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20113
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20114
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20115
    apply(blast)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20116
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20117
next
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20118
  case (TCut a \<Delta> N x \<Gamma> M B \<theta>n \<theta>c)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20119
  then show ?case 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20120
    apply -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20121
    apply(case_tac "\<forall>y. M\<noteq>Ax y a")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20122
    apply(case_tac "\<forall>c. N\<noteq>Ax x c")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20123
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20124
    apply(rule_tac B="B" in CUT_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20125
    apply(rule BINDING_implies_CAND)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20126
    apply(unfold BINDINGc_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20127
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20128
    apply(rule_tac x="a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20129
    apply(rule_tac x="\<theta>n,\<theta>c<M>" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20130
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20131
    apply(rule allI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20132
    apply(rule allI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20133
    apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20134
    apply(simp add: psubst_csubst[symmetric]) (*?*)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20135
    apply(drule_tac x="\<theta>n" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20136
    apply(drule_tac x="(a,xa,P)#\<theta>c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20137
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20138
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20139
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20140
    apply(rule ccloses_extend) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20141
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20142
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20143
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20144
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20145
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20146
    apply(rule BINDING_implies_CAND)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20147
    apply(unfold BINDINGn_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20148
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20149
    apply(rule_tac x="x" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20150
    apply(rule_tac x="\<theta>n,\<theta>c<N>" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20151
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20152
    apply(rule allI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20153
    apply(rule allI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20154
    apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20155
    apply(simp add: psubst_nsubst[symmetric]) (*?*)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20156
    apply(rotate_tac 11)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20157
    apply(drule_tac x="(x,aa,P)#\<theta>n" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20158
    apply(drule_tac x="\<theta>c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20159
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20160
    apply(rule ncloses_extend)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20161
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20162
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20163
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20164
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20165
    apply(drule_tac meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20166
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20167
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20168
    (* cases at least one axiom *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20169
    apply(simp (no_asm_use))
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20170
    apply(erule exE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20171
    apply(simp del: psubst.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20172
    apply(drule typing_Ax_elim2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20173
    apply(auto simp add: trm.inject)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20174
    apply(rule_tac B="B" in CUT_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20175
    (* left term *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20176
    apply(rule BINDING_implies_CAND)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20177
    apply(unfold BINDINGc_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20178
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20179
    apply(rule_tac x="a" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20180
    apply(rule_tac x="\<theta>n,\<theta>c<M>" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20181
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20182
    apply(rule allI)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20183
    apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20184
    apply(drule_tac x="\<theta>n" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20185
    apply(drule_tac x="(a,xa,P)#\<theta>c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20186
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20187
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20188
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20189
    apply(rule ccloses_extend) 
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20190
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20191
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20192
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20193
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20194
    apply(simp add: psubst_csubst[symmetric]) (*?*)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20195
    (* right term -axiom *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20196
    apply(drule ccloses_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20197
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20198
    apply(erule exE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20199
    apply(erule conjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20200
    apply(frule_tac y="x" in lookupd_cmaps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20201
    apply(drule cmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20202
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20203
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20204
    apply(subgoal_tac "(x):P[xa\<turnstile>n>x] = (xa):P")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20205
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20206
    apply(simp add: ntrm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20207
    apply(simp add: alpha fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20208
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20209
    apply(rule nrename_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20210
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20211
    (* M is axiom *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20212
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20213
    apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20214
    (* both are axioms *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20215
    apply(rule_tac B="B" in CUT_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20216
    apply(drule typing_Ax_elim1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20217
    apply(drule ncloses_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20218
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20219
    apply(erule exE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20220
    apply(erule conjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20221
    apply(frule_tac a="a" in lookupc_nmaps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20222
    apply(drule_tac a="a" in nmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20223
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20224
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20225
    apply(subgoal_tac "<a>:P[c\<turnstile>c>a] = <c>:P")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20226
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20227
    apply(simp add: ctrm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20228
    apply(simp add: alpha fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20229
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20230
    apply(rule crename_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20231
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20232
    apply(drule typing_Ax_elim2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20233
    apply(drule ccloses_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20234
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20235
    apply(erule exE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20236
    apply(erule conjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20237
    apply(frule_tac y="x" in lookupd_cmaps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20238
    apply(drule cmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20239
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20240
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20241
    apply(subgoal_tac "(x):P[xa\<turnstile>n>x] = (xa):P")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20242
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20243
    apply(simp add: ntrm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20244
    apply(simp add: alpha fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20245
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20246
    apply(rule nrename_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20247
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20248
    (* N is not axioms *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20249
    apply(rule_tac B="B" in CUT_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20250
    (* left term *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20251
    apply(drule typing_Ax_elim1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20252
    apply(drule ncloses_elim)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20253
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20254
    apply(erule exE)+
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20255
    apply(erule conjE)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20256
    apply(frule_tac a="a" in lookupc_nmaps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20257
    apply(drule_tac a="a" in nmaps_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20258
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20259
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20260
    apply(subgoal_tac "<a>:P[c\<turnstile>c>a] = <c>:P")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20261
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20262
    apply(simp add: ctrm.inject)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20263
    apply(simp add: alpha fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20264
    apply(rule sym)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20265
    apply(rule crename_swap)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20266
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20267
    apply(rule BINDING_implies_CAND)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20268
    apply(unfold BINDINGn_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20269
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20270
    apply(rule_tac x="x" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20271
    apply(rule_tac x="\<theta>n,\<theta>c<N>" in exI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20272
    apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20273
    apply(rule allI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20274
    apply(rule allI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20275
    apply(rule impI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20276
    apply(simp add: psubst_nsubst[symmetric]) (*?*)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20277
    apply(rotate_tac 10)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20278
    apply(drule_tac x="(x,aa,P)#\<theta>n" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20279
    apply(drule_tac x="\<theta>c" in meta_spec)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20280
    apply(drule meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20281
    apply(rule ncloses_extend)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20282
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20283
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20284
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20285
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20286
    apply(drule_tac meta_mp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20287
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20288
    apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20289
    done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20290
qed
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20291
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20292
consts
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20293
  "idn" :: "(name\<times>ty) list\<Rightarrow>coname\<Rightarrow>(name\<times>coname\<times>trm) list"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20294
primrec
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20295
  "idn [] a   = []"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20296
  "idn (p#\<Gamma>) a = ((fst p),a,Ax (fst p) a)#(idn \<Gamma> a)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20297
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20298
consts
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20299
  "idc" :: "(coname\<times>ty) list\<Rightarrow>name\<Rightarrow>(coname\<times>name\<times>trm) list"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20300
primrec
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20301
  "idc [] x    = []"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20302
  "idc (p#\<Delta>) x = ((fst p),x,Ax x (fst p))#(idc \<Delta> x)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20303
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20304
lemma idn_eqvt[eqvt]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20305
  fixes pi1::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20306
  and   pi2::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20307
  shows "(pi1\<bullet>(idn \<Gamma> a)) = idn (pi1\<bullet>\<Gamma>) (pi1\<bullet>a)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20308
  and   "(pi2\<bullet>(idn \<Gamma> a)) = idn (pi2\<bullet>\<Gamma>) (pi2\<bullet>a)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20309
apply(induct \<Gamma>)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20310
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20311
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20312
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20313
lemma idc_eqvt[eqvt]:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20314
  fixes pi1::"name prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20315
  and   pi2::"coname prm"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20316
  shows "(pi1\<bullet>(idc \<Delta> x)) = idc (pi1\<bullet>\<Delta>) (pi1\<bullet>x)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20317
  and   "(pi2\<bullet>(idc \<Delta> x)) = idc (pi2\<bullet>\<Delta>) (pi2\<bullet>x)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20318
apply(induct \<Delta>)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20319
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20320
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20321
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20322
lemma ccloses_id:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20323
  shows "(idc \<Delta> x) ccloses \<Delta>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20324
apply(induct \<Delta>)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20325
apply(auto simp add: ccloses_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20326
apply(rule Ax_in_CANDs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20327
apply(rule Ax_in_CANDs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20328
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20329
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20330
lemma ncloses_id:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20331
  shows "(idn \<Gamma> a) ncloses \<Gamma>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20332
apply(induct \<Gamma>)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20333
apply(auto simp add: ncloses_def)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20334
apply(rule Ax_in_CANDs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20335
apply(rule Ax_in_CANDs)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20336
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20337
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20338
lemma fresh_idn:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20339
  fixes x::"name"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20340
  and   a::"coname"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20341
  shows "x\<sharp>\<Gamma> \<Longrightarrow> x\<sharp>idn \<Gamma> a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20342
  and   "a\<sharp>(\<Gamma>,b) \<Longrightarrow> a\<sharp>idn \<Gamma> b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20343
apply(induct \<Gamma>)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20344
apply(auto simp add: fresh_list_cons fresh_list_nil fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20345
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20346
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20347
lemma fresh_idc:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20348
  fixes x::"name"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20349
  and   a::"coname"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20350
  shows "x\<sharp>(\<Delta>,y) \<Longrightarrow> x\<sharp>idc \<Delta> y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20351
  and   "a\<sharp>\<Delta>  \<Longrightarrow> a\<sharp>idc \<Delta> y"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20352
apply(induct \<Delta>)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20353
apply(auto simp add: fresh_list_cons fresh_list_nil fresh_atm fresh_prod)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20354
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20355
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20356
lemma idc_cmaps:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20357
  assumes a: "idc \<Delta> y cmaps b to Some (x,M)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20358
  shows "M=Ax x b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20359
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20360
apply(induct \<Delta>)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20361
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20362
apply(case_tac "b=a")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20363
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20364
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20365
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20366
lemma idn_nmaps:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20367
  assumes a: "idn \<Gamma> a nmaps x to Some (b,M)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20368
  shows "M=Ax x b"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20369
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20370
apply(induct \<Gamma>)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20371
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20372
apply(case_tac "aa=x")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20373
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20374
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20375
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20376
lemma lookup1:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20377
  assumes a: "x\<sharp>(idn \<Gamma> b)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20378
  shows "lookup x a (idn \<Gamma> b) \<theta>c = lookupa x a \<theta>c"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20379
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20380
apply(induct \<Gamma>)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20381
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20382
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20383
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20384
lemma lookup2:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20385
  assumes a: "\<not>(x\<sharp>(idn \<Gamma> b))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20386
  shows "lookup x a (idn \<Gamma> b) \<theta>c = lookupb x a \<theta>c b (Ax x b)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20387
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20388
apply(induct \<Gamma>)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20389
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20390
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20391
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20392
lemma lookup3:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20393
  assumes a: "a\<sharp>(idc \<Delta> y)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20394
  shows "lookupa x a (idc \<Delta> y) = Ax x a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20395
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20396
apply(induct \<Delta>)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20397
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20398
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20399
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20400
lemma lookup4:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20401
  assumes a: "\<not>(a\<sharp>(idc \<Delta> y))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20402
  shows "lookupa x a (idc \<Delta> y) = Cut <a>.(Ax x a) (y).Ax y a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20403
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20404
apply(induct \<Delta>)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20405
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20406
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20407
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20408
lemma lookup5:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20409
  assumes a: "a\<sharp>(idc \<Delta> y)"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20410
  shows "lookupb x a (idc \<Delta> y) c P = Cut <c>.P (x).Ax x a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20411
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20412
apply(induct \<Delta>)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20413
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20414
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20415
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20416
lemma lookup6:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20417
  assumes a: "\<not>(a\<sharp>(idc \<Delta> y))"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20418
  shows "lookupb x a (idc \<Delta> y) c P = Cut <c>.P (y).Ax y a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20419
using a
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20420
apply(induct \<Delta>)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20421
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20422
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20423
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20424
lemma lookup7:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20425
  shows "lookupc x a (idn \<Gamma> b) = Ax x a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20426
apply(induct \<Gamma>)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20427
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20428
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20429
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20430
lemma lookup8:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20431
  shows "lookupd x a (idc \<Delta> y) = Ax x a"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20432
apply(induct \<Delta>)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20433
apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20434
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20435
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20436
lemma id_redu:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20437
  shows "(idn \<Gamma> x),(idc \<Delta> a)<M> \<longrightarrow>\<^isub>a* M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20438
apply(nominal_induct M avoiding: \<Gamma> \<Delta> x a rule: trm.induct)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20439
apply(auto)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20440
(* Ax *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20441
apply(case_tac "name\<sharp>(idn \<Gamma> x)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20442
apply(simp add: lookup1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20443
apply(case_tac "coname\<sharp>(idc \<Delta> a)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20444
apply(simp add: lookup3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20445
apply(simp add: lookup4)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20446
apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20447
apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20448
apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20449
apply(rule better_LAxR_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20450
apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20451
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20452
apply(simp add: lookup2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20453
apply(case_tac "coname\<sharp>(idc \<Delta> a)")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20454
apply(simp add: lookup5)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20455
apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20456
apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20457
apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20458
apply(rule better_LAxR_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20459
apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20460
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20461
apply(simp add: lookup6)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20462
apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20463
apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20464
apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20465
apply(rule better_LAxR_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20466
apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20467
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20468
(* Cut *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20469
apply(auto simp add: fresh_idn fresh_idc psubst_fresh_name psubst_fresh_coname fresh_atm fresh_prod )[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20470
apply(simp add: lookup7 lookup8)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20471
apply(simp add: lookup7 lookup8)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20472
apply(simp add: a_star_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20473
apply(simp add: lookup7 lookup8)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20474
apply(simp add: a_star_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20475
apply(simp add: a_star_Cut)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20476
(* NotR *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20477
apply(simp add: fresh_idn fresh_idc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20478
apply(case_tac "findc (idc \<Delta> a) coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20479
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20480
apply(simp add: a_star_NotR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20481
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20482
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20483
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20484
apply(drule idc_cmaps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20485
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20486
apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20487
apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20488
apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20489
apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20490
apply(rule better_LAxR_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20491
apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20492
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20493
apply(simp add: crename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20494
apply(simp add: a_star_NotR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20495
apply(rule psubst_fresh_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20496
apply(rule fresh_idn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20497
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20498
apply(rule fresh_idc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20499
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20500
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20501
(* NotL *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20502
apply(simp add: fresh_idn fresh_idc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20503
apply(case_tac "findn (idn \<Gamma> x) name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20504
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20505
apply(simp add: a_star_NotL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20506
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20507
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20508
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20509
apply(drule idn_nmaps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20510
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20511
apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20512
apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20513
apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20514
apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20515
apply(rule better_LAxL_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20516
apply(rule fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20517
apply(assumption)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20518
apply(simp add: nrename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20519
apply(simp add: a_star_NotL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20520
apply(rule psubst_fresh_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20521
apply(rule fresh_idn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20522
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20523
apply(rule fresh_idc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20524
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20525
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20526
(* AndR *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20527
apply(simp add: fresh_idn fresh_idc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20528
apply(case_tac "findc (idc \<Delta> a) coname3")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20529
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20530
apply(simp add: a_star_AndR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20531
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20532
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20533
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20534
apply(drule idc_cmaps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20535
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20536
apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm1>")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20537
apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm2>")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20538
apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20539
apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20540
apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20541
apply(rule better_LAxR_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20542
apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20543
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20544
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20545
apply(auto simp add: fresh_idn fresh_idc psubst_fresh_name crename_fresh fresh_atm fresh_prod )[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20546
apply(rule aux3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20547
apply(rule crename.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20548
apply(auto simp add: fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20549
apply(rule psubst_fresh_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20550
apply(rule fresh_idn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20551
apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20552
apply(rule fresh_idc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20553
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20554
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20555
apply(auto simp add: fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20556
apply(rule psubst_fresh_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20557
apply(rule fresh_idn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20558
apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20559
apply(rule fresh_idc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20560
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20561
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20562
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20563
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20564
apply(simp add: crename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20565
apply(simp add: a_star_AndR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20566
apply(rule psubst_fresh_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20567
apply(rule fresh_idn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20568
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20569
apply(rule fresh_idc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20570
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20571
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20572
apply(rule psubst_fresh_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20573
apply(rule fresh_idn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20574
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20575
apply(rule fresh_idc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20576
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20577
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20578
(* AndL1 *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20579
apply(simp add: fresh_idn fresh_idc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20580
apply(case_tac "findn (idn \<Gamma> x) name2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20581
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20582
apply(simp add: a_star_AndL1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20583
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20584
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20585
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20586
apply(drule idn_nmaps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20587
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20588
apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20589
apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20590
apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20591
apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20592
apply(rule better_LAxL_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20593
apply(rule fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20594
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20595
apply(rule aux3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20596
apply(rule nrename.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20597
apply(auto simp add: fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20598
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20599
apply(simp add: nrename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20600
apply(simp add: a_star_AndL1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20601
apply(rule psubst_fresh_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20602
apply(rule fresh_idn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20603
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20604
apply(rule fresh_idc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20605
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20606
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20607
(* AndL2 *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20608
apply(simp add: fresh_idn fresh_idc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20609
apply(case_tac "findn (idn \<Gamma> x) name2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20610
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20611
apply(simp add: a_star_AndL2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20612
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20613
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20614
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20615
apply(drule idn_nmaps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20616
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20617
apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20618
apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20619
apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20620
apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20621
apply(rule better_LAxL_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20622
apply(rule fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20623
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20624
apply(rule aux3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20625
apply(rule nrename.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20626
apply(auto simp add: fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20627
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20628
apply(simp add: nrename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20629
apply(simp add: a_star_AndL2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20630
apply(rule psubst_fresh_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20631
apply(rule fresh_idn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20632
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20633
apply(rule fresh_idc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20634
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20635
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20636
(* OrR1 *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20637
apply(simp add: fresh_idn fresh_idc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20638
apply(case_tac "findc (idc \<Delta> a) coname2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20639
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20640
apply(simp add: a_star_OrR1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20641
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20642
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20643
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20644
apply(drule idc_cmaps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20645
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20646
apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20647
apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20648
apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20649
apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20650
apply(rule better_LAxR_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20651
apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20652
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20653
apply(rule aux3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20654
apply(rule crename.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20655
apply(auto simp add: fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20656
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20657
apply(simp add: crename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20658
apply(simp add: a_star_OrR1)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20659
apply(rule psubst_fresh_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20660
apply(rule fresh_idn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20661
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20662
apply(rule fresh_idc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20663
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20664
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20665
(* OrR2 *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20666
apply(simp add: fresh_idn fresh_idc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20667
apply(case_tac "findc (idc \<Delta> a) coname2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20668
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20669
apply(simp add: a_star_OrR2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20670
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20671
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20672
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20673
apply(drule idc_cmaps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20674
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20675
apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20676
apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20677
apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20678
apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20679
apply(rule better_LAxR_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20680
apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20681
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20682
apply(rule aux3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20683
apply(rule crename.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20684
apply(auto simp add: fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20685
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20686
apply(simp add: crename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20687
apply(simp add: a_star_OrR2)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20688
apply(rule psubst_fresh_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20689
apply(rule fresh_idn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20690
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20691
apply(rule fresh_idc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20692
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20693
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20694
(* OrL *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20695
apply(simp add: fresh_idn fresh_idc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20696
apply(case_tac "findn (idn \<Gamma> x) name3")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20697
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20698
apply(simp add: a_star_OrL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20699
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20700
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20701
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20702
apply(drule idn_nmaps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20703
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20704
apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm1>")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20705
apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm2>")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20706
apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20707
apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20708
apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20709
apply(rule better_LAxL_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20710
apply(rule fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20711
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20712
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20713
apply(rule aux3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20714
apply(rule nrename.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20715
apply(auto simp add: fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20716
apply(rule psubst_fresh_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20717
apply(rule fresh_idn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20718
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20719
apply(rule fresh_idc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20720
apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20721
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20722
apply(auto simp add: fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20723
apply(rule psubst_fresh_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20724
apply(rule fresh_idn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20725
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20726
apply(rule fresh_idc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20727
apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20728
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20729
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20730
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20731
apply(simp add: nrename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20732
apply(simp add: a_star_OrL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20733
apply(rule psubst_fresh_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20734
apply(rule fresh_idn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20735
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20736
apply(rule fresh_idc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20737
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20738
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20739
apply(rule psubst_fresh_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20740
apply(rule fresh_idn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20741
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20742
apply(rule fresh_idc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20743
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20744
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20745
(* ImpR *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20746
apply(simp add: fresh_idn fresh_idc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20747
apply(case_tac "findc (idc \<Delta> a) coname2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20748
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20749
apply(simp add: a_star_ImpR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20750
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20751
apply(generate_fresh "coname")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20752
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20753
apply(drule idc_cmaps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20754
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20755
apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20756
apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20757
apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20758
apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20759
apply(rule better_LAxR_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20760
apply(rule fic.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20761
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20762
apply(rule aux3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20763
apply(rule crename.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20764
apply(auto simp add: fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20765
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20766
apply(simp add: crename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20767
apply(simp add: a_star_ImpR)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20768
apply(rule psubst_fresh_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20769
apply(rule fresh_idn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20770
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20771
apply(rule fresh_idc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20772
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20773
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20774
(* ImpL *)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20775
apply(simp add: fresh_idn fresh_idc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20776
apply(case_tac "findn (idn \<Gamma> x) name2")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20777
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20778
apply(simp add: a_star_ImpL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20779
apply(auto)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20780
apply(generate_fresh "name")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20781
apply(fresh_fun_simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20782
apply(drule idn_nmaps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20783
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20784
apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm1>")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20785
apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm2>")
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20786
apply(rule a_star_trans)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20787
apply(rule a_starI)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20788
apply(rule al_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20789
apply(rule better_LAxL_intro)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20790
apply(rule fin.intros)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20791
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20792
apply(simp add: abs_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20793
apply(rule aux3)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20794
apply(rule nrename.simps)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20795
apply(auto simp add: fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20796
apply(rule psubst_fresh_coname)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20797
apply(rule fresh_idn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20798
apply(simp add: fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20799
apply(rule fresh_idc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20800
apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20801
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20802
apply(auto simp add: fresh_prod fresh_atm)[1]
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20803
apply(rule psubst_fresh_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20804
apply(rule fresh_idn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20805
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20806
apply(rule fresh_idc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20807
apply(simp add: fresh_prod fresh_atm)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20808
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20809
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20810
apply(simp add: nrename_fresh)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20811
apply(simp add: a_star_ImpL)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20812
apply(rule psubst_fresh_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20813
apply(rule fresh_idn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20814
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20815
apply(rule fresh_idc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20816
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20817
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20818
apply(rule psubst_fresh_name)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20819
apply(rule fresh_idn)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20820
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20821
apply(rule fresh_idc)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20822
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20823
apply(simp)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20824
done
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20825
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20826
theorem ALL_SNa:
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20827
  assumes a: "\<Gamma> \<turnstile> M \<turnstile> \<Delta>"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20828
  shows "SNa M"
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20829
proof -
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20830
  have "(idc \<Delta> x) ccloses \<Delta>" by (simp add: ccloses_id)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20831
  moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20832
  have "(idn \<Gamma> a) ncloses \<Gamma>" by (simp add: ncloses_id)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20833
  ultimately have "SNa ((idn \<Gamma> a),(idc \<Delta> x)<M>)" using a by (simp add: all_CAND)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20834
  moreover
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20835
  have "((idn \<Gamma> a),(idc \<Delta> x)<M>) \<longrightarrow>\<^isub>a* M" by (simp add: id_redu)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20836
  ultimately show "SNa M" by (simp add: a_star_preserves_SNa)
f3615235dc4d formalisation of my PhD (the result was correct, but the proof needed several corrections)
urbanc
parents: 22542
diff changeset
 20837
qed
19500
188d4e44c1a6 Removed broken proof.
berghofe
parents: 19477
diff changeset
 20838
188d4e44c1a6 Removed broken proof.
berghofe
parents: 19477
diff changeset
 20839
end
188d4e44c1a6 Removed broken proof.
berghofe
parents: 19477
diff changeset
 20840