src/HOL/Nominal/Examples/CR_Takahashi.thy
author urbanc
Wed, 02 May 2007 01:42:23 +0200
changeset 22829 f1db55c7534d
parent 22821 15b2e7ec1f3b
child 22833 51a16a718820
permissions -rw-r--r--
tuned some proofs and changed variable names in some definitions of Nominal.thy
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
22821
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
     1
(* $Id$ *)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
     2
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
     3
theory CR_Takahashi
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
     4
imports Lam_Funs
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
     5
begin
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
     6
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
     7
text {* Authors: Mathilde Arnaud and Christian Urban
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
     8
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
     9
        The Church-Rosser proof from a paper by Masako Takahashi.
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
    10
        This formalisation follows with some very slight exceptions 
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
    11
        the one  given by Randy Pollack in his paper:
22821
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    12
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
    13
             Polishing Up the Tait-Martin Löf Proof of the 
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
    14
             Church-Rosser Theorem (1995).
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
    15
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
    16
  *}
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
    17
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
    18
section {* Lemmas about Capture-Avoiding Substitution *}
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
    19
 
22821
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    20
lemma forget:
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    21
  assumes asm: "x\<sharp>L"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    22
  shows "L[x::=P] = L"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    23
  using asm 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    24
by (nominal_induct L avoiding: x P rule: lam.induct)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    25
   (auto simp add: abs_fresh fresh_atm)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    26
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    27
lemma fresh_fact:
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    28
  fixes z::"name"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    29
  assumes asms: "z\<sharp>N" "z\<sharp>L"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    30
  shows "z\<sharp>(N[y::=L])"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    31
  using asms 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    32
by (nominal_induct N avoiding: z y L rule: lam.induct)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    33
   (auto simp add: abs_fresh fresh_atm)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    34
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    35
lemma fresh_fact': 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    36
  fixes a::"name"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    37
  assumes a: "a\<sharp>t2"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    38
  shows "a\<sharp>t1[a::=t2]"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    39
using a 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    40
by (nominal_induct t1 avoiding: a t2 rule: lam.induct)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    41
   (auto simp add: abs_fresh fresh_atm)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    42
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    43
lemma substitution_lemma:  
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    44
  assumes asm: "x\<noteq>y" "x\<sharp>L"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    45
  shows "M[x::=N][y::=L] = M[y::=L][x::=N[y::=L]]"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    46
  using asm 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    47
by (nominal_induct M avoiding: x y N L rule: lam.induct)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    48
   (auto simp add: fresh_fact forget)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    49
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
    50
lemma subst_rename: 
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
    51
  assumes a: "c\<sharp>t1"
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
    52
  shows "t1[a::=t2] = ([(c,a)]\<bullet>t1)[c::=t2]"
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
    53
using a
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
    54
by (nominal_induct t1 avoiding: a c t2 rule: lam.induct)
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
    55
   (auto simp add: calc_atm fresh_atm abs_fresh)
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
    56
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
    57
22821
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    58
section {* Beta Reduction *}
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    59
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    60
inductive2
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    61
  "Beta" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^isub>\<beta> _" [80,80] 80)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    62
where
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    63
    b1[intro]: "s1\<longrightarrow>\<^isub>\<beta>s2 \<Longrightarrow> (App s1 t)\<longrightarrow>\<^isub>\<beta>(App s2 t)"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    64
  | b2[intro]: "s1\<longrightarrow>\<^isub>\<beta>s2 \<Longrightarrow> (App t s1)\<longrightarrow>\<^isub>\<beta>(App t s2)"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    65
  | b3[intro]: "s1\<longrightarrow>\<^isub>\<beta>s2 \<Longrightarrow> (Lam [a].s1)\<longrightarrow>\<^isub>\<beta> (Lam [a].s2)"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    66
  | b4[intro]: "a\<sharp>s2 \<Longrightarrow> (App (Lam [a].s1) s2)\<longrightarrow>\<^isub>\<beta>(s1[a::=s2])"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    67
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    68
equivariance Beta
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    69
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    70
nominal_inductive Beta
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    71
  by (simp_all add: abs_fresh fresh_fact')
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    72
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
    73
section {* Transitive Closure of Beta *}
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
    74
22821
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    75
inductive2
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    76
  "Beta_star"  :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^isub>\<beta>\<^sup>* _" [80,80] 80)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    77
where
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
    78
    bs1[intro,simp]: "M \<longrightarrow>\<^isub>\<beta>\<^sup>* M"
22821
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    79
  | bs2[intro]: "\<lbrakk>M1\<longrightarrow>\<^isub>\<beta>\<^sup>* M2; M2 \<longrightarrow>\<^isub>\<beta> M3\<rbrakk> \<Longrightarrow> M1 \<longrightarrow>\<^isub>\<beta>\<^sup>* M3"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    80
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    81
equivariance Beta_star
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    82
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    83
lemma beta_star_trans:
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    84
  assumes a1: "M1\<longrightarrow>\<^isub>\<beta>\<^sup>* M2"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    85
  and     a2: "M2\<longrightarrow>\<^isub>\<beta>\<^sup>* M3"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    86
  shows "M1 \<longrightarrow>\<^isub>\<beta>\<^sup>* M3"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    87
using a2 a1
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    88
by (induct) (auto)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    89
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    90
section {* One-Reduction *}
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    91
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    92
inductive2
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    93
  One :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^isub>1 _" [80,80] 80)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    94
where
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    95
    o1[intro!]:      "M\<longrightarrow>\<^isub>1M"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    96
  | o2[simp,intro!]: "\<lbrakk>t1\<longrightarrow>\<^isub>1t2;s1\<longrightarrow>\<^isub>1s2\<rbrakk> \<Longrightarrow> (App t1 s1)\<longrightarrow>\<^isub>1(App t2 s2)"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    97
  | o3[simp,intro!]: "s1\<longrightarrow>\<^isub>1s2 \<Longrightarrow> (Lam [a].s1)\<longrightarrow>\<^isub>1(Lam [a].s2)"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    98
  | o4[simp,intro!]: "\<lbrakk>a\<sharp>(s1,s2); s1\<longrightarrow>\<^isub>1s2;t1\<longrightarrow>\<^isub>1t2\<rbrakk> \<Longrightarrow> (App (Lam [a].t1) s1)\<longrightarrow>\<^isub>1(t2[a::=s2])"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
    99
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   100
equivariance One
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   101
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   102
nominal_inductive One
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   103
  by (simp_all add: abs_fresh fresh_fact')
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   104
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   105
lemma one_subst_aux:
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   106
  assumes a: "N\<longrightarrow>\<^isub>1N'"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   107
  shows "M[x::=N] \<longrightarrow>\<^isub>1 M[x::=N']"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   108
using a
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   109
by (nominal_induct M avoiding: x N N' rule: lam.induct)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   110
   (auto simp add: fresh_prod fresh_atm)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   111
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   112
lemma one_subst: 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   113
  assumes a: "M\<longrightarrow>\<^isub>1M'" 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   114
  and     b: "N\<longrightarrow>\<^isub>1N'"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   115
  shows "M[x::=N]\<longrightarrow>\<^isub>1M'[x::=N']" 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   116
using a b
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   117
by (nominal_induct M M' avoiding: N N' x rule: One.strong_induct)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   118
   (auto simp add: one_subst_aux substitution_lemma fresh_atm fresh_fact)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   119
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
   120
section {* Transitive Closure of One *}
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
   121
22821
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   122
inductive2
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   123
  "One_star"  :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^isub>1\<^sup>* _" [80,80] 80)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   124
where
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   125
    os1[intro, simp]: "M \<longrightarrow>\<^isub>1\<^sup>* M"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   126
  | os2[intro]: "\<lbrakk>M1\<longrightarrow>\<^isub>1\<^sup>* M2; M2 \<longrightarrow>\<^isub>1 M3\<rbrakk> \<Longrightarrow> M1 \<longrightarrow>\<^isub>1\<^sup>* M3"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   127
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   128
equivariance One_star 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   129
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   130
lemma one_star_trans:
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   131
  assumes a1: "M1\<longrightarrow>\<^isub>1\<^sup>* M2" 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   132
  and     a2: "M2\<longrightarrow>\<^isub>1\<^sup>* M3"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   133
  shows "M1\<longrightarrow>\<^isub>1\<^sup>* M3"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   134
using a2 a1
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   135
by (induct) (auto)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   136
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   137
lemma one_fresh_preserv:
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   138
  fixes a :: "name"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   139
  assumes a: "t\<longrightarrow>\<^isub>1s"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   140
  and     b: "a\<sharp>t"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   141
  shows "a\<sharp>s"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   142
using a b
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   143
by (nominal_induct avoiding: a rule: One.strong_induct)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   144
   (auto simp add: abs_fresh fresh_atm fresh_fact)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   145
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
   146
section {* Elimination Rules for One *}
22821
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   147
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   148
lemma one_var:
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   149
  assumes a: "Var x \<longrightarrow>\<^isub>1 t"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   150
  shows "t = Var x"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   151
using a
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
   152
by (erule_tac One.cases) (simp_all) 
22821
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   153
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   154
lemma one_abs: 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   155
  assumes a: "(Lam [a].t)\<longrightarrow>\<^isub>1t'"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   156
  shows "\<exists>t''. t'=Lam [a].t'' \<and> t\<longrightarrow>\<^isub>1t''"
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
   157
using a
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
   158
  apply(erule_tac One.cases)
22821
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   159
  apply(auto simp add: lam.inject alpha)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   160
  apply(rule_tac x="[(a,aa)]\<bullet>s2" in exI)
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
   161
  apply(perm_simp add: fresh_left calc_atm)
22821
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   162
  apply(simp add: One.eqvt)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   163
  apply(simp add: one_fresh_preserv)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   164
done  
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   165
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   166
lemma one_app: 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   167
  assumes a: "App t1 t2 \<longrightarrow>\<^isub>1 t'"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   168
  shows "(\<exists>s1 s2. t' = App s1 s2 \<and> t1 \<longrightarrow>\<^isub>1 s1 \<and> t2 \<longrightarrow>\<^isub>1 s2) \<or> 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   169
         (\<exists>a s s1 s2. t1 = Lam [a].s \<and> a\<sharp>(t2,s2) \<and> t' = s1[a::=s2] \<and> s \<longrightarrow>\<^isub>1 s1 \<and> t2 \<longrightarrow>\<^isub>1 s2)" 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   170
  using a
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
   171
  apply(erule_tac One.cases)
22821
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   172
  apply(auto simp add: lam.distinct lam.inject)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   173
  done
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   174
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   175
lemma one_red: 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   176
  assumes a: "App (Lam [a].t1) t2 \<longrightarrow>\<^isub>1 M"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   177
  shows "(\<exists>s1 s2. M = App (Lam [a].s1) s2 \<and> t1 \<longrightarrow>\<^isub>1 s1 \<and> t2 \<longrightarrow>\<^isub>1 s2) \<or> 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   178
         (\<exists>s1 s2. M = s1[a::=s2] \<and> t1 \<longrightarrow>\<^isub>1 s1 \<and> t2 \<longrightarrow>\<^isub>1 s2)" 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   179
  using a
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
   180
  apply(erule_tac One.cases)
22821
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   181
  apply(simp_all add: lam.inject)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   182
  apply(force)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   183
  apply(erule conjE)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   184
  apply(drule sym[of "Lam [a].t1"])
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   185
  apply(simp)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   186
  apply(drule one_abs)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   187
  apply(erule exE)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   188
  apply(simp)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   189
  apply(force simp add: alpha)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   190
  apply(erule conjE)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   191
  apply(simp add: lam.inject alpha)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   192
  apply(erule disjE)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   193
  apply(simp)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   194
  apply(force)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   195
  apply(simp)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   196
  apply(rule disjI2)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   197
  apply(rule_tac x="[(a,aa)]\<bullet>t2a" in exI)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   198
  apply(rule_tac x="s2" in exI)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   199
  apply(auto)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   200
  apply(subgoal_tac "a\<sharp>t2a")(*A*)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   201
  apply(simp add: subst_rename)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   202
  (*A*)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   203
  apply(force intro: one_fresh_preserv)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   204
  apply(simp add: One.eqvt)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   205
  done
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   206
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
   207
text {* Complete Development Reduction *}
22821
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   208
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   209
inductive2
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   210
  cd1 :: "lam \<Rightarrow> lam \<Rightarrow> bool" (" _ >c _" [80,80]80)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   211
where
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   212
  cd1v[intro!]:      "Var x >c Var x"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   213
  | cd1l[simp,intro!]: "s1 >c s2 \<Longrightarrow> Lam [a].s1 >c Lam[a].s2"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   214
  | cd1a[simp,intro!]: "\<lbrakk>\<not>(\<exists> a s. s1 = Lam [a].s); s1 >c s2; t1 >c t2\<rbrakk> \<Longrightarrow> App s1 t1 >c App s2 t2"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   215
  | cd1r[simp,intro!]: "\<lbrakk>a\<sharp>(s1,s2); s1 >c s2; t1 >c t2\<rbrakk> \<Longrightarrow> App (Lam [a].t1) s1 >c (t2[a::=s2])"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   216
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   217
(* FIXME: needs to be in nominal_inductive *)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   218
declare perm_pi_simp[eqvt_force]
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   219
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   220
equivariance cd1
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   221
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   222
nominal_inductive cd1
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   223
  by (simp_all add: abs_fresh fresh_fact')
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   224
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   225
lemma better_cd1r_intro[intro]:
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   226
  assumes a: "s1 >c s2"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   227
  and     b: "t1 >c t2"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   228
  shows "App (Lam [a].t1) s1 >c (t2[a::=s2])"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   229
proof -
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   230
  obtain c::"name" where fs: "c\<sharp>(a,t1,s1,t2,s2)" by (rule exists_fresh, rule fin_supp,blast)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   231
  have eq1: "Lam [a].t1 = Lam [c].([(c,a)]\<bullet>t1)" using fs
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   232
    by (rule_tac sym, auto simp add: lam.inject alpha fresh_prod fresh_atm)
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
   233
  have "App (Lam [a].t1) s1 = App (Lam [c].([(c,a)]\<bullet>t1)) s1" using eq1 by simp
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
   234
  also have "\<dots> >c  ([(c,a)]\<bullet>t2)[c::=s2]" using fs a b by (simp_all add: cd1.eqvt)
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
   235
  also have "\<dots> = t2[a::=s2]" using fs  by (rule_tac subst_rename[symmetric], simp)
22821
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   236
  finally show "App (Lam [a].t1) s1 >c (t2[a::=s2])" by simp
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   237
qed
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   238
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   239
lemma cd1_fresh_preserve:
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   240
  fixes a::"name"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   241
  assumes a: "a\<sharp>s1" 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   242
  and     b: "s1 >c s2"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   243
  shows "a\<sharp>s2"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   244
using b a
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
   245
 by (induct) (auto simp add: abs_fresh fresh_fact fresh_fact')
22821
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   246
  
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   247
lemma cd1_lam:
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   248
  assumes a: "Lam [a].t >c t'"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   249
  shows "\<exists>s. t'=Lam [a].s \<and> t >c s"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   250
using a
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   251
apply -
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   252
apply(erule cd1.cases)
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
   253
apply(simp_all add: lam.inject alpha)
22821
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   254
apply(auto)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   255
apply(rule_tac x="[(a,aa)]\<bullet>s2" in exI)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   256
apply(perm_simp add: fresh_left cd1.eqvt cd1_fresh_preserve)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   257
done
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   258
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   259
lemma develop_existence:
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   260
  shows "\<exists>M'. M >c M'"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   261
by (nominal_induct M rule: lam.induct)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   262
   (auto dest!: cd1_lam)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   263
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   264
lemma triangle:
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   265
  assumes a: "M >c M'"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   266
  and     b: "M \<longrightarrow>\<^isub>1 M''"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   267
  shows "M'' \<longrightarrow>\<^isub>1 M'"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   268
using a b
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   269
by (nominal_induct avoiding: M'' rule: cd1.strong_induct)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   270
   (auto dest!: one_var one_app one_abs one_red intro: one_subst)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   271
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   272
lemma diamond:
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   273
  assumes a: "M1 \<longrightarrow>\<^isub>1 M2"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   274
  and     b: "M1 \<longrightarrow>\<^isub>1 M3"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   275
  shows "\<exists>M4. M2 \<longrightarrow>\<^isub>1 M4 \<and> M3 \<longrightarrow>\<^isub>1 M4"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   276
proof -
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   277
  obtain Mc where c: "M1 >c Mc" using develop_existence by blast
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   278
  have "M2 \<longrightarrow>\<^isub>1 Mc" using a c by (simp add: triangle)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   279
  moreover
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   280
  have "M3 \<longrightarrow>\<^isub>1 Mc" using b c by (simp add: triangle)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   281
  ultimately show "\<exists>M4. M2 \<longrightarrow>\<^isub>1 M4 \<and> M3 \<longrightarrow>\<^isub>1 M4" by blast
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   282
qed
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   283
  
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   284
lemma one_lam_cong: 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   285
  assumes a: "t1\<longrightarrow>\<^isub>\<beta>\<^sup>*t2" 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   286
  shows "(Lam [a].t1)\<longrightarrow>\<^isub>\<beta>\<^sup>*(Lam [a].t2)"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   287
  using a
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
   288
by (induct) (blast)+
22821
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   289
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   290
lemma one_app_congL: 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   291
  assumes a: "t1\<longrightarrow>\<^isub>\<beta>\<^sup>*t2" 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   292
  shows "App t1 s\<longrightarrow>\<^isub>\<beta>\<^sup>* App t2 s"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   293
  using a
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
   294
by (induct) (blast)+
22821
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   295
  
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   296
lemma one_app_congR: 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   297
  assumes a: "t1\<longrightarrow>\<^isub>\<beta>\<^sup>*t2" 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   298
  shows "App s t1 \<longrightarrow>\<^isub>\<beta>\<^sup>* App s t2"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   299
using a
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
   300
by (induct) (blast)+
22821
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   301
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   302
lemma one_app_cong: 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   303
  assumes a1: "t1\<longrightarrow>\<^isub>\<beta>\<^sup>*t2" 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   304
  and     a2: "s1\<longrightarrow>\<^isub>\<beta>\<^sup>*s2" 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   305
  shows "App t1 s1\<longrightarrow>\<^isub>\<beta>\<^sup>* App t2 s2"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   306
proof -
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   307
  have "App t1 s1 \<longrightarrow>\<^isub>\<beta>\<^sup>* App t2 s1" using a1 by (rule one_app_congL)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   308
  moreover
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   309
  have "App t2 s1 \<longrightarrow>\<^isub>\<beta>\<^sup>* App t2 s2" using a2 by (rule one_app_congR)
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22821
diff changeset
   310
  ultimately show "App t1 s1\<longrightarrow>\<^isub>\<beta>\<^sup>* App t2 s2" by (rule beta_star_trans)
22821
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   311
qed
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   312
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   313
lemma one_beta_star: 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   314
  assumes a: "(t1\<longrightarrow>\<^isub>1t2)" 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   315
  shows "(t1\<longrightarrow>\<^isub>\<beta>\<^sup>*t2)"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   316
  using a
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   317
proof(nominal_induct rule: One.strong_induct)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   318
  case (o4 a s1 s2 t1 t2)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   319
  have vc: "a\<sharp>s1" "a\<sharp>s2" by fact
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   320
  have a1: "t1\<longrightarrow>\<^isub>\<beta>\<^sup>*t2" and a2: "s1\<longrightarrow>\<^isub>\<beta>\<^sup>*s2" by fact
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   321
  have c1: "(App (Lam [a].t2) s2) \<longrightarrow>\<^isub>\<beta> (t2 [a::= s2])" using vc by (simp add: b4)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   322
  from a1 a2 have c2: "App (Lam [a].t1 ) s1 \<longrightarrow>\<^isub>\<beta>\<^sup>* App (Lam [a].t2 ) s2" 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   323
    by (blast intro!: one_app_cong one_lam_cong)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   324
  show ?case using c2 c1 by (blast intro: beta_star_trans)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   325
qed (auto intro!: one_app_cong one_lam_cong)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   326
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   327
lemma one_star_lam_cong: 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   328
  assumes a: "t1\<longrightarrow>\<^isub>1\<^sup>*t2" 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   329
  shows "(Lam  [a].t1)\<longrightarrow>\<^isub>1\<^sup>* (Lam [a].t2)"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   330
  using a
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   331
by (induct) (auto intro: one_star_trans)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   332
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   333
lemma one_star_app_congL: 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   334
  assumes a: "t1\<longrightarrow>\<^isub>1\<^sup>*t2" 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   335
  shows "App t1 s\<longrightarrow>\<^isub>1\<^sup>* App t2 s"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   336
  using a
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   337
by (induct) (auto intro: one_star_trans)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   338
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   339
lemma one_star_app_congR: 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   340
  assumes a: "t1\<longrightarrow>\<^isub>1\<^sup>*t2" 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   341
  shows "App s t1 \<longrightarrow>\<^isub>1\<^sup>* App s t2"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   342
  using a
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   343
by (induct) (auto intro: one_star_trans)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   344
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   345
lemma beta_one_star: 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   346
  assumes a: "t1\<longrightarrow>\<^isub>\<beta>t2" 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   347
  shows "t1\<longrightarrow>\<^isub>1\<^sup>*t2"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   348
  using a
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   349
by (induct)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   350
   (auto intro!: one_star_app_congL one_star_app_congR one_star_lam_cong)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   351
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   352
lemma rectangle_for_one:
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   353
  assumes a: "t\<longrightarrow>\<^isub>1\<^sup>*t1" 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   354
  and     b: "t\<longrightarrow>\<^isub>1t2"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   355
  shows "\<exists>t3. t1\<longrightarrow>\<^isub>1t3 \<and> t2\<longrightarrow>\<^isub>1\<^sup>*t3"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   356
  using a b
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   357
proof (induct arbitrary: t2)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   358
  case os1 thus ?case by force
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   359
next
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   360
  case (os2 t s1 s2 t2)  
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   361
  have b: "s1 \<longrightarrow>\<^isub>1 s2" by fact
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   362
  have h: "\<And>t2. t \<longrightarrow>\<^isub>1 t2 \<Longrightarrow> (\<exists>t3. s1 \<longrightarrow>\<^isub>1 t3 \<and> t2 \<longrightarrow>\<^isub>1\<^sup>* t3)" by fact
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   363
  have c: "t \<longrightarrow>\<^isub>1 t2" by fact
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   364
  show "\<exists>t3. s2 \<longrightarrow>\<^isub>1 t3 \<and>  t2 \<longrightarrow>\<^isub>1\<^sup>* t3" 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   365
  proof -
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   366
    from c h have "\<exists>t3. s1 \<longrightarrow>\<^isub>1 t3 \<and> t2 \<longrightarrow>\<^isub>1\<^sup>* t3" by blast
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   367
    then obtain t3 where c1: "s1 \<longrightarrow>\<^isub>1 t3" and c2: "t2 \<longrightarrow>\<^isub>1\<^sup>* t3" by blast
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   368
    have "\<exists>t4. s2 \<longrightarrow>\<^isub>1 t4 \<and> t3 \<longrightarrow>\<^isub>1 t4" using b c1 by (blast intro: diamond)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   369
    thus ?thesis using c2 by (blast intro: one_star_trans)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   370
  qed
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   371
qed
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   372
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   373
lemma cr_for_one_star: 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   374
  assumes a: "t\<longrightarrow>\<^isub>1\<^sup>*t2"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   375
      and b: "t\<longrightarrow>\<^isub>1\<^sup>*t1"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   376
    shows "\<exists>t3. t1\<longrightarrow>\<^isub>1\<^sup>*t3\<and>t2\<longrightarrow>\<^isub>1\<^sup>*t3"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   377
using a b
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   378
proof (induct arbitrary: t1)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   379
  case (os1 t) then show ?case by force
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   380
next 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   381
  case (os2 t s1 s2 t1)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   382
  have c: "t \<longrightarrow>\<^isub>1\<^sup>* s1" by fact
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   383
  have c': "t \<longrightarrow>\<^isub>1\<^sup>* t1" by fact
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   384
  have d: "s1 \<longrightarrow>\<^isub>1 s2" by fact
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   385
  have "t \<longrightarrow>\<^isub>1\<^sup>* t1 \<Longrightarrow> (\<exists>t3.  t1 \<longrightarrow>\<^isub>1\<^sup>* t3 \<and> s1 \<longrightarrow>\<^isub>1\<^sup>* t3)" by fact
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   386
  then obtain t3 where f1: "t1 \<longrightarrow>\<^isub>1\<^sup>* t3"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   387
                   and f2: "s1 \<longrightarrow>\<^isub>1\<^sup>* t3" using c' by blast
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   388
  from rectangle_for_one d f2 have "\<exists>t4. t3\<longrightarrow>\<^isub>1t4 \<and> s2\<longrightarrow>\<^isub>1\<^sup>*t4" by blast
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   389
  then obtain t4 where g1: "t3\<longrightarrow>\<^isub>1t4"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   390
                   and g2: "s2\<longrightarrow>\<^isub>1\<^sup>*t4" by blast
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   391
  have "t1\<longrightarrow>\<^isub>1\<^sup>*t4" using f1 g1 by (blast intro: one_star_trans)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   392
  thus ?case using g2 by blast
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   393
qed
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   394
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   395
lemma beta_star_and_one_star: 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   396
  shows "(M1\<longrightarrow>\<^isub>1\<^sup>*M2) = (M1\<longrightarrow>\<^isub>\<beta>\<^sup>*M2)"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   397
proof
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   398
  assume "M1 \<longrightarrow>\<^isub>1\<^sup>* M2"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   399
  then show "M1\<longrightarrow>\<^isub>\<beta>\<^sup>*M2"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   400
  proof induct
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   401
    case (os1 M1) thus "M1\<longrightarrow>\<^isub>\<beta>\<^sup>*M1" by simp
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   402
  next
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   403
    case (os2 M1 M2 M3)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   404
    have "M2\<longrightarrow>\<^isub>1M3" by fact
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   405
    then have "M2\<longrightarrow>\<^isub>\<beta>\<^sup>*M3" by (rule one_beta_star)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   406
    moreover have "M1\<longrightarrow>\<^isub>\<beta>\<^sup>*M2" by fact
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   407
    ultimately show "M1\<longrightarrow>\<^isub>\<beta>\<^sup>*M3" by (auto intro: beta_star_trans)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   408
  qed
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   409
next
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   410
  assume "M1 \<longrightarrow>\<^isub>\<beta>\<^sup>* M2" 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   411
  then show "M1\<longrightarrow>\<^isub>1\<^sup>*M2"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   412
  proof induct
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   413
    case (bs1 M1) thus  "M1\<longrightarrow>\<^isub>1\<^sup>*M1" by simp
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   414
  next
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   415
    case (bs2 M1 M2 M3) 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   416
    have "M2\<longrightarrow>\<^isub>\<beta>M3" by fact
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   417
    then have "M2\<longrightarrow>\<^isub>1\<^sup>*M3" by (rule beta_one_star)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   418
    moreover have "M1\<longrightarrow>\<^isub>1\<^sup>*M2" by fact
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   419
    ultimately show "M1\<longrightarrow>\<^isub>1\<^sup>*M3" by (auto intro: one_star_trans)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   420
  qed
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   421
qed
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   422
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   423
lemma cr_for_beta_star: 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   424
  assumes a1: "t\<longrightarrow>\<^isub>\<beta>\<^sup>*t1" 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   425
  and     a2: "t\<longrightarrow>\<^isub>\<beta>\<^sup>*t2" 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   426
  shows "\<exists>t3. t1\<longrightarrow>\<^isub>\<beta>\<^sup>*t3\<and>t2\<longrightarrow>\<^isub>\<beta>\<^sup>*t3"
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   427
proof -
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   428
  from a1 have "t\<longrightarrow>\<^isub>1\<^sup>*t1" by (simp only: beta_star_and_one_star)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   429
  moreover
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   430
  from a2 have "t\<longrightarrow>\<^isub>1\<^sup>*t2" by (simp only: beta_star_and_one_star)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   431
  ultimately have "\<exists>t3. t1\<longrightarrow>\<^isub>1\<^sup>*t3 \<and> t2\<longrightarrow>\<^isub>1\<^sup>*t3" by (blast intro: cr_for_one_star) 
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   432
  then obtain t3 where "t1\<longrightarrow>\<^isub>1\<^sup>*t3" and "t2\<longrightarrow>\<^isub>1\<^sup>*t3" by (blast intro: cr_for_one_star)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   433
  hence "t1\<longrightarrow>\<^isub>\<beta>\<^sup>*t3" and "t2\<longrightarrow>\<^isub>\<beta>\<^sup>*t3" by (simp_all only: beta_star_and_one_star)
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   434
  then show "\<exists>t3. t1\<longrightarrow>\<^isub>\<beta>\<^sup>*t3\<and>t2\<longrightarrow>\<^isub>\<beta>\<^sup>*t3" by blast
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   435
qed
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   436
15b2e7ec1f3b alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff changeset
   437
end