src/HOL/Nominal/Examples/CR.thy
author urbanc
Thu, 01 Dec 2005 05:20:13 +0100
changeset 18312 c68296902ddb
parent 18303 b18fabea0fd0
child 18344 95083a68cbbb
permissions -rw-r--r--
cleaned up further the proofs (diamond still needs work); changed "fresh:" to "avoiding:"
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
18269
3f36e2165e51 some small tuning
urbanc
parents: 18106
diff changeset
     1
(* $Id$ *)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
     2
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
     3
theory cr
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
     4
imports lam_substs
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
     5
begin
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
     6
18269
3f36e2165e51 some small tuning
urbanc
parents: 18106
diff changeset
     7
text {* The Church-Rosser proof from Barendregt's book *}
3f36e2165e51 some small tuning
urbanc
parents: 18106
diff changeset
     8
18312
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
     9
lemma forget: 
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    10
  assumes a: "a\<sharp>t1"
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    11
  shows "t1[a::=t2] = t1"
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    12
using a
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    13
proof (nominal_induct t1 avoiding: a t2 rule: lam_induct)
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    14
  case (Var b) 
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    15
  thus ?case by (simp add: fresh_atm)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    16
next 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    17
  case App
18312
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    18
  thus ?case by simp
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    19
next
18312
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    20
  case (Lam c t)
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    21
  have ih: "\<And>c t2. c\<sharp>t \<Longrightarrow>  t[c::=t2] = t" by fact
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
    22
  have a: "c\<sharp>t2" by fact
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
    23
  have "c\<sharp>a" by fact
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
    24
  hence b: "a\<noteq>c" by (simp add: fresh_atm)
18312
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    25
  have "a\<sharp>Lam [c].t" by fact 
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    26
  hence "a\<sharp>t" using b by (simp add: abs_fresh)
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    27
  hence "t[a::=t2] = t" using ih by simp
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    28
  thus "(Lam [c].t)[a::=t2] = Lam [c].t" using a b by simp
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    29
qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    30
18312
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    31
lemma forget: 
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    32
  assumes a: "a\<sharp>t1"
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    33
  shows "t1[a::=t2] = t1"
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    34
  using a
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    35
apply (nominal_induct t1 avoiding: a t2 rule: lam_induct)
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
    36
apply(auto simp add: abs_fresh fresh_atm)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    37
done
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    38
18312
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    39
lemma fresh_fact: 
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    40
  fixes   b :: "name"
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    41
  and    a  :: "name"
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    42
  and    t1 :: "lam"
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    43
  and    t2 :: "lam"
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    44
  assumes a: "a\<sharp>t1"
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    45
  and     b: "a\<sharp>t2"
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    46
  shows "a\<sharp>(t1[b::=t2])"
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    47
using a b
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    48
proof (nominal_induct t1 avoiding: a b t2 rule: lam_induct)
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    49
  case (Var c) 
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    50
  thus ?case by simp
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    51
next
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    52
  case App thus ?case by simp 
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    53
next
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    54
  case (Lam c t)
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    55
  have ih: "\<And>(a::name) b t2. a\<sharp>t \<Longrightarrow> a\<sharp>t2 \<Longrightarrow> a\<sharp>(t[b::=t2])" by fact
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    56
  have fr: "c\<sharp>a" "c\<sharp>b" "c\<sharp>t2" by fact+
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    57
  hence fr': "c\<noteq>a" by (simp add: fresh_atm) 
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    58
  have a1: "a\<sharp>t2" by fact
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    59
  have a2: "a\<sharp>Lam [c].t" by fact
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    60
  hence "a\<sharp>t" using fr' by (simp add: abs_fresh)
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    61
  hence "a\<sharp>t[b::=t2]" using a1 ih by simp
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    62
  thus "a\<sharp>(Lam [c].t)[b::=t2]" using fr  by (simp add: abs_fresh)
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    63
qed
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    64
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    65
lemma fresh_fact: 
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    66
  fixes   b :: "name"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    67
  and    a  :: "name"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    68
  and    t1 :: "lam"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    69
  and    t2 :: "lam" 
18312
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    70
  assumes a: "a\<sharp>t1"
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    71
  and     b: "a\<sharp>t2"
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    72
  shows "a\<sharp>(t1[b::=t2])"
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    73
using a b
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    74
apply(nominal_induct t1 avoiding: a b t2 rule: lam_induct)
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    75
apply(auto simp add: abs_fresh fresh_atm)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    76
done
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    77
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    78
lemma subs_lemma:  
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    79
  fixes x::"name"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    80
  and   y::"name"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    81
  and   L::"lam"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    82
  and   M::"lam"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    83
  and   N::"lam"
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
    84
  assumes a: "x\<noteq>y"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
    85
  and     b: "x\<sharp>L"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
    86
  shows "M[x::=N][y::=L] = M[y::=L][x::=N[y::=L]]"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
    87
using a b
18312
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
    88
proof (nominal_induct M avoiding: x y N L rule: lam_induct)
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
    89
  case (Var z) (* case 1: Variables*)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
    90
  have "x\<noteq>y" by fact
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
    91
  have "x\<sharp>L" by fact
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
    92
  show "Var z[x::=N][y::=L] = Var z[y::=L][x::=N[y::=L]]" (is "?LHS = ?RHS")
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
    93
  proof -
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
    94
    { (*Case 1.1*)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
    95
      assume  "z=x"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
    96
      have "(1)": "?LHS = N[y::=L]" using `z=x` by simp
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
    97
      have "(2)": "?RHS = N[y::=L]" using `z=x` `x\<noteq>y` by simp
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
    98
      from "(1)" "(2)" have "?LHS = ?RHS"  by simp
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
    99
    }
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   100
    moreover 
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   101
    { (*Case 1.2*)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   102
      assume "z\<noteq>x" and "z=y"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   103
      have "(1)": "?LHS = L"               using `z\<noteq>x` `z=y` by force
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   104
      have "(2)": "?RHS = L[x::=N[y::=L]]" using `z=y` by force
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   105
      have "(3)": "L[x::=N[y::=L]] = L"    using `x\<sharp>L` by (simp add: forget)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   106
      from "(1)" "(2)" "(3)" have "?LHS = ?RHS" by simp
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   107
    }
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   108
    moreover 
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   109
    { (*Case 1.3*)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   110
      assume "z\<noteq>x" and "z\<noteq>y"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   111
      have "(1)": "?LHS = Var z" using `z\<noteq>x` `z\<noteq>y` by force
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   112
      have "(2)": "?RHS = Var z" using `z\<noteq>x` `z\<noteq>y` by force
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   113
      from "(1)" "(2)" have "?LHS = ?RHS" by simp
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   114
    }
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   115
    ultimately show "?LHS = ?RHS" by blast
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   116
  qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   117
next
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   118
  case (Lam z M1) (* case 2: lambdas *)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   119
  have ih: "\<And>x y N L. x\<noteq>y \<Longrightarrow> x\<sharp>L \<Longrightarrow> M1[x::=N][y::=L] = M1[y::=L][x::=N[y::=L]]" by fact
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   120
  have "x\<noteq>y" by fact
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   121
  have "x\<sharp>L" by fact
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   122
  have "z\<sharp>x" "z\<sharp>y" "z\<sharp>N" "z\<sharp>L" by fact
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   123
  hence "z\<sharp>N[y::=L]" by (simp add: fresh_fact)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   124
  show "(Lam [z].M1)[x::=N][y::=L] = (Lam [z].M1)[y::=L][x::=N[y::=L]]" (is "?LHS=?RHS") 
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   125
  proof -
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   126
    have "?LHS = Lam [z].(M1[x::=N][y::=L])" using `z\<sharp>x` `z\<sharp>y` `z\<sharp>N` `z\<sharp>L` by simp
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   127
    also from ih have "\<dots> = Lam [z].(M1[y::=L][x::=N[y::=L]])" using `x\<noteq>y` `x\<sharp>L` by simp
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   128
    also have "\<dots> = (Lam [z].(M1[y::=L]))[x::=N[y::=L]]" using `z\<sharp>x` `z\<sharp>N[y::=L]` by simp
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   129
    also have "\<dots> = ?RHS" using  `z\<sharp>y` `z\<sharp>L` by simp
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   130
    finally show "?LHS = ?RHS" .
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   131
  qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   132
next
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   133
  case (App M1 M2) (* case 3: applications *)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   134
  thus ?case by simp
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   135
qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   136
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   137
lemma subs_lemma:  
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   138
  fixes x::"name"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   139
  and   y::"name"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   140
  and   L::"lam"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   141
  and   M::"lam"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   142
  and   N::"lam"
18312
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   143
  assumes a: "x\<noteq>y"
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   144
  and     b: "x\<sharp>L"
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   145
  shows "M[x::=N][y::=L] = M[y::=L][x::=N[y::=L]]"
18312
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   146
using a b
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   147
by (nominal_induct M avoiding: x y N L rule: lam_induct)
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   148
   (auto simp add: fresh_fact forget)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   149
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   150
lemma subst_rename[rule_format]: 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   151
  fixes  c  :: "name"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   152
  and    a  :: "name"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   153
  and    t1 :: "lam"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   154
  and    t2 :: "lam"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   155
  shows "c\<sharp>t1 \<longrightarrow> (t1[a::=t2] = ([(c,a)]\<bullet>t1)[c::=t2])"
18312
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   156
proof (nominal_induct t1 avoiding: a c t2 rule: lam_induct)
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   157
  case (Var b)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   158
  show "c\<sharp>(Var b) \<longrightarrow> (Var b)[a::=t2] = ([(c,a)]\<bullet>(Var b))[c::=t2]" by (simp add: calc_atm fresh_atm)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   159
next
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   160
  case App thus ?case by force
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   161
next
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   162
  case (Lam b s)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   163
  have i: "\<And>a c t2. c\<sharp>s \<longrightarrow> (s[a::=t2] = ([(c,a)]\<bullet>s)[c::=t2])" by fact
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   164
  have f: "b\<sharp>a" "b\<sharp>c" "b\<sharp>t2" by fact
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   165
  from f have a:"b\<noteq>c" and b: "b\<noteq>a" and c: "b\<sharp>t2" by (simp add: fresh_atm)+
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   166
  show "c\<sharp>Lam [b].s \<longrightarrow> (Lam [b].s)[a::=t2] = ([(c,a)]\<bullet>(Lam [b].s))[c::=t2]" (is "_ \<longrightarrow> ?LHS = ?RHS")
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   167
  proof
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   168
    assume "c\<sharp>Lam [b].s"
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   169
    hence "c\<sharp>s" using a by (simp add: abs_fresh)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   170
    hence d: "s[a::=t2] = ([(c,a)]\<bullet>s)[c::=t2]" using i by simp
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   171
    have    "?LHS = Lam [b].(s[a::=t2])" using b c by simp
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   172
    also have "\<dots> = Lam [b].(([(c,a)]\<bullet>s)[c::=t2])" using d by simp
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   173
    also have "\<dots> = (Lam [b].([(c,a)]\<bullet>s))[c::=t2]" using a c by simp
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   174
    also have "\<dots> = ?RHS" using a b by (simp add: calc_atm)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   175
    finally show "?LHS = ?RHS" by simp
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   176
  qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   177
qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   178
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   179
lemma subst_rename[rule_format]: 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   180
  fixes  c  :: "name"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   181
  and    a  :: "name"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   182
  and    t1 :: "lam"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   183
  and    t2 :: "lam"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   184
  shows "c\<sharp>t1 \<longrightarrow> (t1[a::=t2] = ([(c,a)]\<bullet>t1)[c::=t2])"
18312
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   185
apply(nominal_induct t1 avoiding: a c t2 rule: lam_induct)
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   186
apply(auto simp add: calc_atm fresh_atm abs_fresh)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   187
done
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   188
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   189
section {* Beta Reduction *}
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   190
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   191
consts
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   192
  Beta :: "(lam\<times>lam) set"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   193
syntax 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   194
  "_Beta"       :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^isub>\<beta> _" [80,80] 80)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   195
  "_Beta_star"  :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^isub>\<beta>\<^sup>* _" [80,80] 80)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   196
translations 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   197
  "t1 \<longrightarrow>\<^isub>\<beta> t2" \<rightleftharpoons> "(t1,t2) \<in> Beta"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   198
  "t1 \<longrightarrow>\<^isub>\<beta>\<^sup>* t2" \<rightleftharpoons> "(t1,t2) \<in> Beta\<^sup>*"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   199
inductive Beta
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   200
  intros
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   201
  b1[intro!]: "s1\<longrightarrow>\<^isub>\<beta>s2 \<Longrightarrow> (App s1 t)\<longrightarrow>\<^isub>\<beta>(App s2 t)"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   202
  b2[intro!]: "s1\<longrightarrow>\<^isub>\<beta>s2 \<Longrightarrow> (App t s1)\<longrightarrow>\<^isub>\<beta>(App t s2)"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   203
  b3[intro!]: "s1\<longrightarrow>\<^isub>\<beta>s2 \<Longrightarrow> (Lam [a].s1)\<longrightarrow>\<^isub>\<beta> (Lam [(a::name)].s2)"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   204
  b4[intro!]: "(App (Lam [(a::name)].s1) s2)\<longrightarrow>\<^isub>\<beta>(s1[a::=s2])"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   205
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   206
lemma eqvt_beta: 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   207
  fixes pi :: "name prm"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   208
  and   t  :: "lam"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   209
  and   s  :: "lam"
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   210
  assumes a: "t\<longrightarrow>\<^isub>\<beta>s"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   211
  shows "(pi\<bullet>t)\<longrightarrow>\<^isub>\<beta>(pi\<bullet>s)"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   212
  using a by (induct, auto)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   213
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   214
lemma beta_induct[consumes 1, case_names b1 b2 b3 b4]:
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   215
  fixes  P :: "'a::fs_name\<Rightarrow>lam \<Rightarrow> lam \<Rightarrow>bool"
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   216
  and    t :: "lam"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   217
  and    s :: "lam"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   218
  and    x :: "'a::fs_name"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   219
  assumes a: "t\<longrightarrow>\<^isub>\<beta>s"
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   220
  and a1:    "\<And>t s1 s2 x. s1\<longrightarrow>\<^isub>\<beta>s2 \<Longrightarrow> (\<And>z. P z s1 s2) \<Longrightarrow> P x (App s1 t) (App s2 t)"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   221
  and a2:    "\<And>t s1 s2 x. s1\<longrightarrow>\<^isub>\<beta>s2 \<Longrightarrow> (\<And>z. P z s1 s2) \<Longrightarrow> P x (App t s1) (App t s2)"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   222
  and a3:    "\<And>a s1 s2 x. a\<sharp>x \<Longrightarrow> s1\<longrightarrow>\<^isub>\<beta>s2 \<Longrightarrow> (\<And>z. P z s1 s2) \<Longrightarrow> P x (Lam [a].s1) (Lam [a].s2)"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   223
  and a4:    "\<And>a t1 s1 x. a\<sharp>(s1,x) \<Longrightarrow> P x (App (Lam [a].t1) s1) (t1[a::=s1])"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   224
  shows "P x t s"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   225
proof -
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   226
  from a have "\<And>(pi::name prm) x. P x (pi\<bullet>t) (pi\<bullet>s)"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   227
  proof (induct)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   228
    case b1 thus ?case using a1 by (simp, blast intro: eqvt_beta)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   229
  next
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   230
    case b2 thus ?case using a2 by (simp, blast intro: eqvt_beta)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   231
  next
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   232
    case (b3 a s1 s2)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   233
    have j1: "s1 \<longrightarrow>\<^isub>\<beta> s2" by fact
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   234
    have j2: "\<And>x (pi::name prm). P x (pi\<bullet>s1) (pi\<bullet>s2)" by fact
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   235
    show ?case 
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   236
    proof (simp)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   237
      have f: "\<exists>c::name. c\<sharp>(pi\<bullet>a,pi\<bullet>s1,pi\<bullet>s2,x)"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   238
	by (rule at_exists_fresh[OF at_name_inst], simp add: fs_name1)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   239
      then obtain c::"name" 
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   240
	where f1: "c\<noteq>(pi\<bullet>a)" and f2: "c\<sharp>x" and f3: "c\<sharp>(pi\<bullet>s1)" and f4: "c\<sharp>(pi\<bullet>s2)"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   241
	by (force simp add: fresh_prod fresh_atm)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   242
      have x: "P x (Lam [c].(([(c,pi\<bullet>a)]@pi)\<bullet>s1)) (Lam [c].(([(c,pi\<bullet>a)]@pi)\<bullet>s2))"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   243
	using a3 f2 j1 j2 by (simp, blast intro: eqvt_beta)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   244
      have alpha1: "(Lam [c].([(c,pi\<bullet>a)]\<bullet>(pi\<bullet>s1))) = (Lam [(pi\<bullet>a)].(pi\<bullet>s1))" using f1 f3
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   245
	by (simp add: lam.inject alpha)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   246
      have alpha2: "(Lam [c].([(c,pi\<bullet>a)]\<bullet>(pi\<bullet>s2))) = (Lam [(pi\<bullet>a)].(pi\<bullet>s2))" using f1 f3
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   247
	by (simp add: lam.inject alpha)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   248
      show " P x (Lam [(pi\<bullet>a)].(pi\<bullet>s1)) (Lam [(pi\<bullet>a)].(pi\<bullet>s2))"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   249
	using x alpha1 alpha2 by (simp only: pt_name2)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   250
    qed
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   251
  next
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   252
    case (b4 a s1 s2)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   253
    show ?case
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   254
    proof (simp add: subst_eqvt)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   255
      have f: "\<exists>c::name. c\<sharp>(pi\<bullet>a,pi\<bullet>s1,pi\<bullet>s2,x)"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   256
	by (rule at_exists_fresh[OF at_name_inst], simp add: fs_name1)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   257
      then obtain c::"name" 
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   258
	where f1: "c\<noteq>(pi\<bullet>a)" and f2: "c\<sharp>(pi\<bullet>s2,x)" and f3: "c\<sharp>(pi\<bullet>s1)" and f4: "c\<sharp>(pi\<bullet>s2)"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   259
	by (force simp add: fresh_prod fresh_atm)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   260
      have x: "P x (App (Lam [c].(([(c,pi\<bullet>a)]@pi)\<bullet>s1)) (pi\<bullet>s2)) ((([(c,pi\<bullet>a)]@pi)\<bullet>s1)[c::=(pi\<bullet>s2)])"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   261
	using a4 f2 by (blast intro!: eqvt_beta)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   262
      have alpha1: "(Lam [c].([(c,pi\<bullet>a)]\<bullet>(pi\<bullet>s1))) = (Lam [(pi\<bullet>a)].(pi\<bullet>s1))" using f1 f3
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   263
	by (simp add: lam.inject alpha)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   264
      have alpha2: "(([(c,pi\<bullet>a)]@pi)\<bullet>s1)[c::=(pi\<bullet>s2)] = (pi\<bullet>s1)[(pi\<bullet>a)::=(pi\<bullet>s2)]"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   265
	using f3 by (simp only: subst_rename[symmetric] pt_name2)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   266
      show "P x (App (Lam [(pi\<bullet>a)].(pi\<bullet>s1)) (pi\<bullet>s2)) ((pi\<bullet>s1)[(pi\<bullet>a)::=(pi\<bullet>s2)])"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   267
	using x alpha1 alpha2 by (simp only: pt_name2)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   268
    qed
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   269
  qed
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   270
  hence "P x (([]::name prm)\<bullet>t) (([]::name prm)\<bullet>s)" by blast 
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   271
  thus ?thesis by simp
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   272
qed
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   273
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   274
section {* One-Reduction *}
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   275
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   276
consts
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   277
  One :: "(lam\<times>lam) set"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   278
syntax 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   279
  "_One"       :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^isub>1 _" [80,80] 80)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   280
  "_One_star"  :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^isub>1\<^sup>* _" [80,80] 80)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   281
translations 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   282
  "t1 \<longrightarrow>\<^isub>1 t2" \<rightleftharpoons> "(t1,t2) \<in> One"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   283
  "t1 \<longrightarrow>\<^isub>1\<^sup>* t2" \<rightleftharpoons> "(t1,t2) \<in> One\<^sup>*"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   284
inductive One
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   285
  intros
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   286
  o1[intro!]:      "M\<longrightarrow>\<^isub>1M"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   287
  o2[simp,intro!]: "\<lbrakk>t1\<longrightarrow>\<^isub>1t2;s1\<longrightarrow>\<^isub>1s2\<rbrakk> \<Longrightarrow> (App t1 s1)\<longrightarrow>\<^isub>1(App t2 s2)"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   288
  o3[simp,intro!]: "s1\<longrightarrow>\<^isub>1s2 \<Longrightarrow> (Lam [(a::name)].s1)\<longrightarrow>\<^isub>1(Lam [a].s2)"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   289
  o4[simp,intro!]: "\<lbrakk>s1\<longrightarrow>\<^isub>1s2;t1\<longrightarrow>\<^isub>1t2\<rbrakk> \<Longrightarrow> (App (Lam [(a::name)].t1) s1)\<longrightarrow>\<^isub>1(t2[a::=s2])"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   290
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   291
lemma eqvt_one: 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   292
  fixes pi :: "name prm"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   293
  and   t  :: "lam"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   294
  and   s  :: "lam"
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   295
  assumes a: "t\<longrightarrow>\<^isub>1s"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   296
  shows "(pi\<bullet>t)\<longrightarrow>\<^isub>1(pi\<bullet>s)"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   297
  using a by (induct, auto)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   298
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   299
lemma one_induct[consumes 1, case_names o1 o2 o3 o4]:
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   300
  fixes  P :: "'a::fs_name\<Rightarrow>lam \<Rightarrow> lam \<Rightarrow>bool"
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   301
  and    t :: "lam"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   302
  and    s :: "lam"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   303
  and    x :: "'a::fs_name"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   304
  assumes a: "t\<longrightarrow>\<^isub>1s"
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   305
  and a1:    "\<And>t x. P x t t"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   306
  and a2:    "\<And>t1 t2 s1 s2 x. t1\<longrightarrow>\<^isub>1t2 \<Longrightarrow> (\<And>z. P z t1 t2) \<Longrightarrow> s1\<longrightarrow>\<^isub>1s2 \<Longrightarrow> (\<And>z. P z s1 s2) \<Longrightarrow> 
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   307
              P x (App t1 s1) (App t2 s2)"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   308
  and a3:    "\<And>a s1 s2 x. a\<sharp>x \<Longrightarrow> s1\<longrightarrow>\<^isub>1s2 \<Longrightarrow> (\<And>z. P z s1 s2) \<Longrightarrow> P x (Lam [a].s1) (Lam [a].s2)"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   309
  and a4:    "\<And>a t1 t2 s1 s2 x. 
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   310
              a\<sharp>(s1,s2,x) \<Longrightarrow> t1\<longrightarrow>\<^isub>1t2 \<Longrightarrow> (\<And>z. P z t1 t2) \<Longrightarrow> s1\<longrightarrow>\<^isub>1s2 \<Longrightarrow> (\<And>z. P z s1 s2) 
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   311
              \<Longrightarrow> P x (App (Lam [a].t1) s1) (t2[a::=s2])"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   312
  shows "P x t s"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   313
proof -
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   314
  from a have "\<And>(pi::name prm) x. P x (pi\<bullet>t) (pi\<bullet>s)"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   315
  proof (induct)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   316
    case o1 show ?case using a1 by force
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   317
  next
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   318
    case (o2 s1 s2 t1 t2) 
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   319
    thus ?case using a2 by (simp, blast intro: eqvt_one)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   320
  next
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   321
    case (o3 a t1 t2)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   322
    have j1: "t1 \<longrightarrow>\<^isub>1 t2" by fact
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   323
    have j2: "\<And>(pi::name prm) x. P x (pi\<bullet>t1) (pi\<bullet>t2)" by fact
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   324
    show ?case 
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   325
    proof (simp)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   326
      have f: "\<exists>c::name. c\<sharp>(pi\<bullet>a,pi\<bullet>t1,pi\<bullet>t2,x)"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   327
	by (rule at_exists_fresh[OF at_name_inst], simp add: fs_name1)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   328
      then obtain c::"name" 
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   329
	where f1: "c\<noteq>(pi\<bullet>a)" and f2: "c\<sharp>x" and f3: "c\<sharp>(pi\<bullet>t1)" and f4: "c\<sharp>(pi\<bullet>t2)"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   330
	by (force simp add: fresh_prod fresh_atm)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   331
      have x: "P x (Lam [c].(([(c,pi\<bullet>a)]@pi)\<bullet>t1)) (Lam [c].(([(c,pi\<bullet>a)]@pi)\<bullet>t2))"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   332
	using a3 f2 j1 j2 by (simp, blast intro: eqvt_one)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   333
      have alpha1: "(Lam [c].([(c,pi\<bullet>a)]\<bullet>(pi\<bullet>t1))) = (Lam [(pi\<bullet>a)].(pi\<bullet>t1))" using f1 f3
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   334
	by (simp add: lam.inject alpha)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   335
      have alpha2: "(Lam [c].([(c,pi\<bullet>a)]\<bullet>(pi\<bullet>t2))) = (Lam [(pi\<bullet>a)].(pi\<bullet>t2))" using f1 f3
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   336
	by (simp add: lam.inject alpha)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   337
      show " P x (Lam [(pi\<bullet>a)].(pi\<bullet>t1)) (Lam [(pi\<bullet>a)].(pi\<bullet>t2))"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   338
	using x alpha1 alpha2 by (simp only: pt_name2)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   339
    qed
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   340
  next
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   341
    case (o4 a s1 s2 t1 t2)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   342
    have j0: "t1 \<longrightarrow>\<^isub>1 t2" by fact
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   343
    have j1: "s1 \<longrightarrow>\<^isub>1 s2" by fact
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   344
    have j2: "\<And>(pi::name prm) x. P x (pi\<bullet>t1) (pi\<bullet>t2)" by fact
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   345
    have j3: "\<And>(pi::name prm) x. P x (pi\<bullet>s1) (pi\<bullet>s2)" by fact
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   346
    show ?case
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   347
    proof (simp)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   348
      have f: "\<exists>c::name. c\<sharp>(pi\<bullet>a,pi\<bullet>t1,pi\<bullet>t2,pi\<bullet>s1,pi\<bullet>s2,x)"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   349
	by (rule at_exists_fresh[OF at_name_inst], simp add: fs_name1)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   350
      then obtain c::"name" 
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   351
	where f1: "c\<noteq>(pi\<bullet>a)" and f2: "c\<sharp>(pi\<bullet>s1,pi\<bullet>s2,x)" and f3: "c\<sharp>(pi\<bullet>t1)" and f4: "c\<sharp>(pi\<bullet>t2)"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   352
	by (force simp add: fresh_prod at_fresh[OF at_name_inst])
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   353
      have x: "P x (App (Lam [c].(([(c,pi\<bullet>a)]@pi)\<bullet>t1)) (pi\<bullet>s1)) ((([(c,pi\<bullet>a)]@pi)\<bullet>t2)[c::=(pi\<bullet>s2)])"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   354
	using a4 f2 j0 j1 j2 j3 by (simp, blast intro!: eqvt_one)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   355
      have alpha1: "(Lam [c].([(c,pi\<bullet>a)]\<bullet>(pi\<bullet>t1))) = (Lam [(pi\<bullet>a)].(pi\<bullet>t1))" using f1 f3
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   356
	by (simp add: lam.inject alpha)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   357
      have alpha2: "(([(c,pi\<bullet>a)]@pi)\<bullet>t2)[c::=(pi\<bullet>s2)] = (pi\<bullet>t2)[(pi\<bullet>a)::=(pi\<bullet>s2)]"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   358
	using f4 by (simp only: subst_rename[symmetric] pt_name2)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   359
      show "P x (App (Lam [(pi\<bullet>a)].(pi\<bullet>t1)) (pi\<bullet>s1)) ((pi\<bullet>t2)[(pi\<bullet>a)::=(pi\<bullet>s2)])"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   360
	using x alpha1 alpha2 by (simp only: pt_name2)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   361
    qed
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   362
  qed
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   363
  hence "P x (([]::name prm)\<bullet>t) (([]::name prm)\<bullet>s)" by blast
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   364
  thus ?thesis by simp
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   365
qed
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   366
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   367
lemma fresh_fact':  
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   368
  assumes a: "a\<sharp>t2" 
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   369
  shows "a\<sharp>(t1[a::=t2])"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   370
using a 
18312
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   371
proof (nominal_induct t1 avoiding: a t2 rule: lam_induct)
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   372
  case (Var b) 
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   373
  thus ?case by (simp add: fresh_atm)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   374
next
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   375
  case App thus ?case by simp
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   376
next
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   377
  case (Lam c t)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   378
  have "a\<sharp>t2" "c\<sharp>a" "c\<sharp>t2" by fact
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   379
  moreover
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   380
  have ih: "\<And>a t2. a\<sharp>t2 \<Longrightarrow> a\<sharp>(t[a::=t2])" by fact
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   381
  ultimately show ?case by (simp add: abs_fresh)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   382
qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   383
18312
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   384
lemma one_fresh_preserv:
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   385
  fixes    t :: "lam"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   386
  and      s :: "lam"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   387
  and      a :: "name"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   388
  assumes a: "t\<longrightarrow>\<^isub>1s"
18312
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   389
  and     b: "a\<sharp>t"
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   390
  shows "a\<sharp>s"
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   391
using a b
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   392
proof (induct)
18312
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   393
  case o1 thus ?case by simp
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   394
next
18312
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   395
  case o2 thus ?case by simp
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   396
next
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   397
  case (o3 c s1 s2)
18312
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   398
  have ih: "a\<sharp>s1 \<Longrightarrow>  a\<sharp>s2" by fact
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   399
  have c: "a\<sharp>Lam [c].s1" by fact
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   400
  show ?case
18312
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   401
  proof (cases "a=c")
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   402
    assume "a=c" thus "a\<sharp>Lam [c].s2" by (simp add: abs_fresh)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   403
  next
18312
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   404
    assume d: "a\<noteq>c" 
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   405
    with c have "a\<sharp>s1" by (simp add: abs_fresh)
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   406
    hence "a\<sharp>s2" using ih by simp
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   407
    thus "a\<sharp>Lam [c].s2" using d by (simp add: abs_fresh) 
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   408
  qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   409
next 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   410
  case (o4 c t1 t2 s1 s2)
18312
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   411
  have i1: "a\<sharp>t1 \<Longrightarrow> a\<sharp>t2" by fact
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   412
  have i2: "a\<sharp>s1 \<Longrightarrow> a\<sharp>s2" by fact
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   413
  have as: "a\<sharp>App (Lam [c].s1) t1" by fact
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   414
  hence c1: "a\<sharp>Lam [c].s1" and c2: "a\<sharp>t1" by (simp add: fresh_prod)+
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   415
  from c2 i1 have c3: "a\<sharp>t2" by simp
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   416
  show "a\<sharp>s2[c::=t2]"
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   417
  proof (cases "a=c")
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   418
    assume "a=c"
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   419
    thus "a\<sharp>s2[c::=t2]" using c3 by (simp add: fresh_fact')
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   420
  next
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   421
    assume d1: "a\<noteq>c"
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   422
    from c1 d1 have "a\<sharp>s1" by (simp add: abs_fresh)
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   423
    hence "a\<sharp>s2" using i2 by simp
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   424
    thus "a\<sharp>s2[c::=t2]" using c3 by (simp add: fresh_fact)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   425
  qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   426
qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   427
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   428
lemma one_abs: 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   429
  fixes    t :: "lam"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   430
  and      t':: "lam"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   431
  and      a :: "name"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   432
  shows "(Lam [a].t)\<longrightarrow>\<^isub>1t'\<Longrightarrow>\<exists>t''. t'=Lam [a].t'' \<and> t\<longrightarrow>\<^isub>1t''"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   433
  apply(ind_cases "(Lam [a].t)\<longrightarrow>\<^isub>1t'")
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   434
  apply(auto simp add: lam.distinct lam.inject alpha)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   435
  apply(rule_tac x="[(a,aa)]\<bullet>s2" in exI)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   436
  apply(rule conjI)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   437
  apply(rule pt_bij2[OF pt_name_inst, OF at_name_inst, symmetric])
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   438
  apply(simp)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   439
  apply(rule pt_name3)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   440
  apply(rule at_ds5[OF at_name_inst])
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   441
  apply(frule_tac a="a" in one_fresh_preserv)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   442
  apply(assumption)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   443
  apply(rule conjI)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   444
  apply(simp add: pt_fresh_left[OF pt_name_inst, OF at_name_inst])
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   445
  apply(simp add: calc_atm)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   446
  apply(force intro!: eqvt_one)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   447
  done
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   448
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   449
lemma one_app: 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   450
  "App t1 t2 \<longrightarrow>\<^isub>1 t' \<Longrightarrow> 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   451
  (\<exists>s1 s2. t' = App s1 s2 \<and> t1 \<longrightarrow>\<^isub>1 s1 \<and> t2 \<longrightarrow>\<^isub>1 s2) \<or> 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   452
  (\<exists>a s s1 s2. t1 = Lam [a].s \<and> t' = s1[a::=s2] \<and> s \<longrightarrow>\<^isub>1 s1 \<and> t2 \<longrightarrow>\<^isub>1 s2)" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   453
  apply(ind_cases "App t1 s1 \<longrightarrow>\<^isub>1 t'")
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   454
  apply(auto simp add: lam.distinct lam.inject)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   455
  done
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   456
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   457
lemma one_red: 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   458
  "App (Lam [a].t1) t2 \<longrightarrow>\<^isub>1 M \<Longrightarrow>
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   459
  (\<exists>s1 s2. M = App (Lam [a].s1) s2 \<and> t1 \<longrightarrow>\<^isub>1 s1 \<and> t2 \<longrightarrow>\<^isub>1 s2) \<or> 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   460
  (\<exists>s1 s2. M = s1[a::=s2] \<and> t1 \<longrightarrow>\<^isub>1 s1 \<and> t2 \<longrightarrow>\<^isub>1 s2)" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   461
  apply(ind_cases "App (Lam [a].t1) s1 \<longrightarrow>\<^isub>1 M")
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   462
  apply(simp_all add: lam.inject)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   463
  apply(force)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   464
  apply(erule conjE)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   465
  apply(drule sym[of "Lam [a].t1"])
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   466
  apply(simp)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   467
  apply(drule one_abs)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   468
  apply(erule exE)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   469
  apply(simp)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   470
  apply(force simp add: alpha)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   471
  apply(erule conjE)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   472
  apply(simp add: lam.inject alpha)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   473
  apply(erule disjE)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   474
  apply(simp)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   475
  apply(force)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   476
  apply(simp)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   477
  apply(rule disjI2)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   478
  apply(rule_tac x="[(a,aa)]\<bullet>t2a" in exI)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   479
  apply(rule_tac x="s2" in exI)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   480
  apply(auto)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   481
  apply(subgoal_tac "a\<sharp>t2a")(*A*)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   482
  apply(simp add: subst_rename)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   483
  (*A*)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   484
  apply(force intro: one_fresh_preserv)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   485
  apply(force intro: eqvt_one)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   486
  done
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   487
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   488
text {* first case in Lemma 3.2.4*}
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   489
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   490
lemma one_subst_aux:
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   491
  fixes    M :: "lam"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   492
  and      N :: "lam"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   493
  and      N':: "lam"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   494
  and      x :: "name"
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   495
  assumes a: "N\<longrightarrow>\<^isub>1N'"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   496
  shows "M[x::=N] \<longrightarrow>\<^isub>1 M[x::=N']"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   497
using a
18312
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   498
proof (nominal_induct M avoiding: x N N' rule: lam_induct)
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   499
  case (Var y) 
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   500
  show "Var y[x::=N] \<longrightarrow>\<^isub>1 Var y[x::=N']" by (cases "x=y", auto)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   501
next
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   502
  case (App P Q) (* application case - third line *)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   503
  thus "(App P Q)[x::=N] \<longrightarrow>\<^isub>1  (App P Q)[x::=N']" using o2 by simp
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   504
next 
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   505
  case (Lam y P) (* abstraction case - fourth line *)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   506
  thus "(Lam [y].P)[x::=N] \<longrightarrow>\<^isub>1 (Lam [y].P)[x::=N']" using o3 by simp
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   507
qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   508
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   509
lemma one_subst_aux:
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   510
  fixes    M :: "lam"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   511
  and      N :: "lam"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   512
  and      N':: "lam"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   513
  and      x :: "name"
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   514
  assumes a: "N\<longrightarrow>\<^isub>1N'"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   515
  shows "M[x::=N] \<longrightarrow>\<^isub>1 M[x::=N']"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   516
using a
18312
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   517
apply(nominal_induct M avoiding: x N N' rule: lam_induct)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   518
apply(auto simp add: fresh_prod fresh_atm)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   519
done
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   520
18312
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   521
lemma one_subst: 
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   522
  fixes    M :: "lam"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   523
  and      M':: "lam"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   524
  and      N :: "lam"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   525
  and      N':: "lam"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   526
  and      x :: "name"
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   527
  assumes a: "M\<longrightarrow>\<^isub>1M'"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   528
  and     b: "N\<longrightarrow>\<^isub>1N'"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   529
  shows "M[x::=N]\<longrightarrow>\<^isub>1M'[x::=N']" 
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   530
using prems
18312
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   531
proof (nominal_induct M M' avoiding: N N' x rule: one_induct)
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   532
  case (o1 M)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   533
  thus ?case by (simp add: one_subst_aux)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   534
next
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   535
  case (o2 M1 M2 N1 N2)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   536
  thus ?case by simp
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   537
next
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   538
  case (o3 a M1 M2)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   539
  thus ?case by simp
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   540
next
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   541
  case (o4 a M1 M2 N1 N2)
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   542
  have e3: "a\<sharp>N1" "a\<sharp>N2" "a\<sharp>N" "a\<sharp>N'" "a\<sharp>x" by fact
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   543
  show ?case
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   544
  proof -
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   545
    have "(App (Lam [a].M1) N1)[x::=N] = App (Lam [a].(M1[x::=N])) (N1[x::=N])" using e3 by simp
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   546
    also have "App (Lam [a].(M1[x::=N])) (N1[x::=N]) \<longrightarrow>\<^isub>1 M2[x::=N'][a::=N2[x::=N']]" 
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   547
      using  o4 b by force
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   548
    also have "M2[x::=N'][a::=N2[x::=N']] = M2[a::=N2][x::=N']" 
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   549
      using e3 by (simp add: subs_lemma fresh_atm)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   550
    ultimately show "(App (Lam [a].M1) N1)[x::=N] \<longrightarrow>\<^isub>1 M2[a::=N2][x::=N']" by simp
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   551
  qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   552
qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   553
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   554
lemma one_subst[rule_format]: 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   555
  assumes a: "M\<longrightarrow>\<^isub>1M'" 
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   556
  and     b: "N\<longrightarrow>\<^isub>1N'"
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   557
  shows "M[x::=N]\<longrightarrow>\<^isub>1M'[x::=N']" 
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   558
using a b
18312
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   559
apply(nominal_induct M M' avoiding: N N' x rule: one_induct)
18303
b18fabea0fd0 modified almost everything for the new nominal_induct
urbanc
parents: 18269
diff changeset
   560
apply(auto simp add: one_subst_aux subs_lemma fresh_atm)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   561
done
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   562
18312
c68296902ddb cleaned up further the proofs (diamond still needs work);
urbanc
parents: 18303
diff changeset
   563
(* FIXME: change to make use of the new induction facilities *)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   564
lemma diamond[rule_format]:
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   565
  fixes    M :: "lam"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   566
  and      M1:: "lam"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   567
  assumes a: "M\<longrightarrow>\<^isub>1M1" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   568
  shows "\<forall>M2. (M\<longrightarrow>\<^isub>1M2) \<longrightarrow> (\<exists>M3. M1\<longrightarrow>\<^isub>1M3 \<and> M2\<longrightarrow>\<^isub>1M3)"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   569
  using a
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   570
proof (induct)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   571
  case (o1 M) (* case 1 --- M1 = M *)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   572
  show "\<forall>M2. M\<longrightarrow>\<^isub>1M2 \<longrightarrow> (\<exists>M3. M\<longrightarrow>\<^isub>1M3 \<and>  M2\<longrightarrow>\<^isub>1M3)" by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   573
next
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   574
  case (o4 x Q Q' P P') (* case 2 --- a beta-reduction occurs*)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   575
  assume i1: "\<forall>M2. Q \<longrightarrow>\<^isub>1M2 \<longrightarrow> (\<exists>M3. Q'\<longrightarrow>\<^isub>1M3 \<and> M2\<longrightarrow>\<^isub>1M3)"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   576
  assume i2: "\<forall>M2. P \<longrightarrow>\<^isub>1M2 \<longrightarrow> (\<exists>M3. P'\<longrightarrow>\<^isub>1M3 \<and> M2\<longrightarrow>\<^isub>1M3)"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   577
  show "\<forall>M2. App (Lam [x].P) Q\<longrightarrow>\<^isub>1M2 \<longrightarrow> (\<exists>M3. P'[x::=Q']\<longrightarrow>\<^isub>1M3 \<and> M2\<longrightarrow>\<^isub>1M3)"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   578
  proof (intro strip)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   579
    fix M2
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   580
    assume "App (Lam [x].P) Q \<longrightarrow>\<^isub>1 M2"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   581
    hence "(\<exists>P' Q'. M2 = App (Lam [x].P') Q' \<and> P\<longrightarrow>\<^isub>1P' \<and> Q\<longrightarrow>\<^isub>1Q') \<or> 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   582
           (\<exists>P' Q'. M2 = P'[x::=Q'] \<and> P\<longrightarrow>\<^isub>1P' \<and> Q\<longrightarrow>\<^isub>1Q')" by (simp add: one_red)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   583
    moreover (* subcase 2.1 *)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   584
    { assume "\<exists>P' Q'. M2 = App (Lam [x].P') Q' \<and> P\<longrightarrow>\<^isub>1P' \<and> Q\<longrightarrow>\<^isub>1Q'"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   585
      then obtain P'' and Q'' where 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   586
	b1: "M2=App (Lam [x].P'') Q''" and b2: "P\<longrightarrow>\<^isub>1P''" and b3: "Q\<longrightarrow>\<^isub>1Q''" by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   587
      from b2 i2 have "(\<exists>M3. P'\<longrightarrow>\<^isub>1M3 \<and> P''\<longrightarrow>\<^isub>1M3)" by simp
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   588
      then obtain P''' where
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   589
	c1: "P'\<longrightarrow>\<^isub>1P'''" and c2: "P''\<longrightarrow>\<^isub>1P'''" by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   590
      from b3 i1 have "(\<exists>M3. Q'\<longrightarrow>\<^isub>1M3 \<and> Q''\<longrightarrow>\<^isub>1M3)" by simp
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   591
      then obtain Q''' where
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   592
	d1: "Q'\<longrightarrow>\<^isub>1Q'''" and d2: "Q''\<longrightarrow>\<^isub>1Q'''" by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   593
      from c1 c2 d1 d2 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   594
      have "P'[x::=Q']\<longrightarrow>\<^isub>1P'''[x::=Q'''] \<and> App (Lam [x].P'') Q'' \<longrightarrow>\<^isub>1 P'''[x::=Q''']" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   595
	by (force simp add: one_subst)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   596
      hence "\<exists>M3. P'[x::=Q']\<longrightarrow>\<^isub>1M3 \<and> M2\<longrightarrow>\<^isub>1M3" using b1 by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   597
    }
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   598
    moreover (* subcase 2.2 *)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   599
    { assume "\<exists>P' Q'. M2 = P'[x::=Q'] \<and> P\<longrightarrow>\<^isub>1P' \<and> Q\<longrightarrow>\<^isub>1Q'"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   600
      then obtain P'' Q'' where
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   601
	b1: "M2=P''[x::=Q'']" and b2: "P\<longrightarrow>\<^isub>1P''" and  b3: "Q\<longrightarrow>\<^isub>1Q''" by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   602
      from b2 i2 have "(\<exists>M3. P'\<longrightarrow>\<^isub>1M3 \<and> P''\<longrightarrow>\<^isub>1M3)" by simp
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   603
      then obtain P''' where
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   604
	c1: "P'\<longrightarrow>\<^isub>1P'''" and c2: "P''\<longrightarrow>\<^isub>1P'''" by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   605
      from b3 i1 have "(\<exists>M3. Q'\<longrightarrow>\<^isub>1M3 \<and> Q''\<longrightarrow>\<^isub>1M3)" by simp
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   606
      then obtain Q''' where
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   607
	d1: "Q'\<longrightarrow>\<^isub>1Q'''" and d2: "Q''\<longrightarrow>\<^isub>1Q'''" by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   608
      from c1 c2 d1 d2 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   609
      have "P'[x::=Q']\<longrightarrow>\<^isub>1P'''[x::=Q'''] \<and> P''[x::=Q'']\<longrightarrow>\<^isub>1P'''[x::=Q''']" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   610
	by (force simp add: one_subst)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   611
      hence "\<exists>M3. P'[x::=Q']\<longrightarrow>\<^isub>1M3 \<and> M2\<longrightarrow>\<^isub>1M3" using b1 by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   612
    }
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   613
    ultimately show "\<exists>M3. P'[x::=Q']\<longrightarrow>\<^isub>1M3 \<and> M2\<longrightarrow>\<^isub>1M3" by blast
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   614
  qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   615
next
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   616
  case (o2 Q Q' P P') (* case 3 *)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   617
  assume i0: "P\<longrightarrow>\<^isub>1P'"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   618
  assume i1: "\<forall>M2. Q \<longrightarrow>\<^isub>1M2 \<longrightarrow> (\<exists>M3. Q'\<longrightarrow>\<^isub>1M3 \<and> M2\<longrightarrow>\<^isub>1M3)"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   619
  assume i2: "\<forall>M2. P \<longrightarrow>\<^isub>1M2 \<longrightarrow> (\<exists>M3. P'\<longrightarrow>\<^isub>1M3 \<and> M2\<longrightarrow>\<^isub>1M3)"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   620
  show "\<forall>M2. App P Q\<longrightarrow>\<^isub>1M2 \<longrightarrow> (\<exists>M3. App P' Q'\<longrightarrow>\<^isub>1M3 \<and> M2\<longrightarrow>\<^isub>1M3)"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   621
  proof (intro strip)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   622
    fix M2
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   623
    assume "App P Q \<longrightarrow>\<^isub>1 M2"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   624
    hence "(\<exists>P'' Q''. M2 = App P'' Q'' \<and> P\<longrightarrow>\<^isub>1P'' \<and> Q\<longrightarrow>\<^isub>1Q'') \<or> 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   625
           (\<exists>x P' P'' Q'. P = Lam [x].P' \<and> M2 = P''[x::=Q'] \<and> P'\<longrightarrow>\<^isub>1 P'' \<and> Q\<longrightarrow>\<^isub>1Q')" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   626
      by (simp add: one_app[simplified])
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   627
    moreover (* subcase 3.1 *)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   628
    { assume "\<exists>P'' Q''. M2 = App P'' Q'' \<and> P\<longrightarrow>\<^isub>1P'' \<and> Q\<longrightarrow>\<^isub>1Q''"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   629
      then obtain P'' and Q'' where 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   630
	b1: "M2=App P'' Q''" and b2: "P\<longrightarrow>\<^isub>1P''" and b3: "Q\<longrightarrow>\<^isub>1Q''" by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   631
      from b2 i2 have "(\<exists>M3. P'\<longrightarrow>\<^isub>1M3 \<and> P''\<longrightarrow>\<^isub>1M3)" by simp
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   632
      then obtain P''' where
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   633
	c1: "P'\<longrightarrow>\<^isub>1P'''" and c2: "P''\<longrightarrow>\<^isub>1P'''" by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   634
      from b3 i1 have "(\<exists>M3. Q'\<longrightarrow>\<^isub>1M3 \<and> Q''\<longrightarrow>\<^isub>1M3)" by simp
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   635
      then obtain Q''' where
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   636
	d1: "Q'\<longrightarrow>\<^isub>1Q'''" and d2: "Q''\<longrightarrow>\<^isub>1Q'''" by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   637
      from c1 c2 d1 d2 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   638
      have "App P' Q'\<longrightarrow>\<^isub>1App P''' Q''' \<and> App P'' Q'' \<longrightarrow>\<^isub>1 App P''' Q'''" by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   639
      hence "\<exists>M3. App P' Q'\<longrightarrow>\<^isub>1M3 \<and> M2\<longrightarrow>\<^isub>1M3" using b1 by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   640
    }
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   641
    moreover (* subcase 3.2 *)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   642
    { assume "\<exists>x P1 P'' Q''. P = Lam [x].P1 \<and> M2 = P''[x::=Q''] \<and> P1\<longrightarrow>\<^isub>1 P'' \<and> Q\<longrightarrow>\<^isub>1Q''"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   643
      then obtain x P1 P1'' Q'' where
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   644
	b0: "P=Lam [x].P1" and b1: "M2=P1''[x::=Q'']" and 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   645
        b2: "P1\<longrightarrow>\<^isub>1P1''" and  b3: "Q\<longrightarrow>\<^isub>1Q''" by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   646
      from b0 i0 have "\<exists>P1'. P'=Lam [x].P1' \<and> P1\<longrightarrow>\<^isub>1P1'" by (simp add: one_abs)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   647
      then obtain P1' where g1: "P'=Lam [x].P1'" and g2: "P1\<longrightarrow>\<^isub>1P1'" by force 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   648
      from g1 b0 b2 i2 have "(\<exists>M3. (Lam [x].P1')\<longrightarrow>\<^isub>1M3 \<and> (Lam [x].P1'')\<longrightarrow>\<^isub>1M3)" by simp
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   649
      then obtain P1''' where
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   650
	c1: "(Lam [x].P1')\<longrightarrow>\<^isub>1P1'''" and c2: "(Lam [x].P1'')\<longrightarrow>\<^isub>1P1'''" by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   651
      from c1 have "\<exists>R1. P1'''=Lam [x].R1 \<and> P1'\<longrightarrow>\<^isub>1R1" by (simp add: one_abs)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   652
      then obtain R1 where r1: "P1'''=Lam [x].R1" and r2: "P1'\<longrightarrow>\<^isub>1R1" by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   653
      from c2 have "\<exists>R2. P1'''=Lam [x].R2 \<and> P1''\<longrightarrow>\<^isub>1R2" by (simp add: one_abs)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   654
      then obtain R2 where r3: "P1'''=Lam [x].R2" and r4: "P1''\<longrightarrow>\<^isub>1R2" by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   655
      from r1 r3 have r5: "R1=R2" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   656
	by (simp add: lam.inject alpha)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   657
      from b3 i1 have "(\<exists>M3. Q'\<longrightarrow>\<^isub>1M3 \<and> Q''\<longrightarrow>\<^isub>1M3)" by simp
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   658
      then obtain Q''' where
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   659
	d1: "Q'\<longrightarrow>\<^isub>1Q'''" and d2: "Q''\<longrightarrow>\<^isub>1Q'''" by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   660
      from g1 r2 d1 r4 r5 d2 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   661
      have "App P' Q'\<longrightarrow>\<^isub>1R1[x::=Q'''] \<and> P1''[x::=Q'']\<longrightarrow>\<^isub>1R1[x::=Q''']" by (simp add: one_subst)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   662
      hence "\<exists>M3. App P' Q'\<longrightarrow>\<^isub>1M3 \<and> M2\<longrightarrow>\<^isub>1M3" using b1 by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   663
    }
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   664
    ultimately show "\<exists>M3. App P' Q'\<longrightarrow>\<^isub>1M3 \<and> M2\<longrightarrow>\<^isub>1M3" by blast
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   665
  qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   666
next
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   667
  case (o3 x P P') (* case 4 *)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   668
  assume i1: "P\<longrightarrow>\<^isub>1P'"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   669
  assume i2: "\<forall>M2. P \<longrightarrow>\<^isub>1M2 \<longrightarrow> (\<exists>M3. P'\<longrightarrow>\<^isub>1M3 \<and> M2\<longrightarrow>\<^isub>1M3)"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   670
  show "\<forall>M2. (Lam [x].P)\<longrightarrow>\<^isub>1M2 \<longrightarrow> (\<exists>M3. (Lam [x].P')\<longrightarrow>\<^isub>1M3 \<and> M2\<longrightarrow>\<^isub>1M3)"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   671
  proof (intro strip)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   672
    fix M2
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   673
    assume "(Lam [x].P)\<longrightarrow>\<^isub>1 M2"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   674
    hence "\<exists>P''. M2=Lam [x].P'' \<and> P\<longrightarrow>\<^isub>1P''" by (simp add: one_abs)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   675
    then obtain P'' where b1: "M2=Lam [x].P''" and b2: "P\<longrightarrow>\<^isub>1P''" by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   676
    from i2 b1 b2 have "\<exists>M3. (Lam [x].P')\<longrightarrow>\<^isub>1M3 \<and> (Lam [x].P'')\<longrightarrow>\<^isub>1M3" by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   677
    then obtain M3 where c1: "(Lam [x].P')\<longrightarrow>\<^isub>1M3" and c2: "(Lam [x].P'')\<longrightarrow>\<^isub>1M3" by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   678
    from c1 have "\<exists>R1. M3=Lam [x].R1 \<and> P'\<longrightarrow>\<^isub>1R1" by (simp add: one_abs)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   679
    then obtain R1 where r1: "M3=Lam [x].R1" and r2: "P'\<longrightarrow>\<^isub>1R1" by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   680
    from c2 have "\<exists>R2. M3=Lam [x].R2 \<and> P''\<longrightarrow>\<^isub>1R2" by (simp add: one_abs)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   681
    then obtain R2 where r3: "M3=Lam [x].R2" and r4: "P''\<longrightarrow>\<^isub>1R2" by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   682
    from r1 r3 have r5: "R1=R2" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   683
      by (simp add: lam.inject alpha)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   684
    from r2 r4 have "(Lam [x].P')\<longrightarrow>\<^isub>1(Lam [x].R1) \<and> (Lam [x].P'')\<longrightarrow>\<^isub>1(Lam [x].R2)" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   685
      by (simp add: one_subst)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   686
    thus "\<exists>M3. (Lam [x].P')\<longrightarrow>\<^isub>1M3 \<and> M2\<longrightarrow>\<^isub>1M3" using b1 r5 by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   687
  qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   688
qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   689
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   690
lemma one_abs_cong: 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   691
  fixes a :: "name"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   692
  assumes a: "t1\<longrightarrow>\<^isub>\<beta>\<^sup>*t2" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   693
  shows "(Lam [a].t1)\<longrightarrow>\<^isub>\<beta>\<^sup>*(Lam [a].t2)"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   694
  using a
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   695
proof induct
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   696
  case 1 thus ?case by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   697
next
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   698
  case (2 y z) 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   699
  thus ?case by (force dest: b3 intro: rtrancl_trans)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   700
qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   701
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   702
lemma one_pr_congL: 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   703
  assumes a: "t1\<longrightarrow>\<^isub>\<beta>\<^sup>*t2" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   704
  shows "App t1 s\<longrightarrow>\<^isub>\<beta>\<^sup>* App t2 s"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   705
  using a
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   706
proof induct
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   707
  case 1 thus ?case by (force)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   708
next
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   709
  case 2 thus ?case by (force dest: b1 intro: rtrancl_trans)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   710
qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   711
  
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   712
lemma one_pr_congR: 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   713
  assumes a: "t1\<longrightarrow>\<^isub>\<beta>\<^sup>*t2" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   714
  shows "App s t1 \<longrightarrow>\<^isub>\<beta>\<^sup>* App s t2"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   715
using a
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   716
proof induct
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   717
  case 1 thus ?case by (force)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   718
next 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   719
  case 2 thus ?case by (force dest: b2 intro: rtrancl_trans)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   720
qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   721
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   722
lemma one_pr_cong: 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   723
  assumes a1: "t1\<longrightarrow>\<^isub>\<beta>\<^sup>*t2" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   724
  and a2: "s1\<longrightarrow>\<^isub>\<beta>\<^sup>*s2" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   725
  shows "App t1 s1\<longrightarrow>\<^isub>\<beta>\<^sup>* App t2 s2"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   726
proof -
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   727
  have b1: "App t1 s1 \<longrightarrow>\<^isub>\<beta>\<^sup>* App t2 s1" using a1 by (rule one_pr_congL)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   728
  have b2: "App t2 s1 \<longrightarrow>\<^isub>\<beta>\<^sup>* App t2 s2" using a2 by (rule one_pr_congR)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   729
  show ?thesis using b1 and b2 by (force intro: rtrancl_trans)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   730
qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   731
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   732
lemma one_beta_star: 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   733
  assumes a: "(t1\<longrightarrow>\<^isub>1t2)" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   734
  shows "(t1\<longrightarrow>\<^isub>\<beta>\<^sup>*t2)"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   735
  using a
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   736
proof induct
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   737
  case o1 thus ?case by (force)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   738
next
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   739
  case o2 thus ?case by (force intro!: one_pr_cong)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   740
next
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   741
  case o3 thus ?case by (force intro!: one_abs_cong)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   742
next 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   743
  case (o4 a s1 s2 t1 t2)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   744
  have a1: "t1\<longrightarrow>\<^isub>\<beta>\<^sup>*t2" and a2: "s1\<longrightarrow>\<^isub>\<beta>\<^sup>*s2" .
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   745
  have c1: "(App (Lam [a].t2) s2) \<longrightarrow>\<^isub>\<beta> (t2 [a::= s2])" by (rule b4)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   746
  from a1 a2 have c2: "App (Lam [a].t1 ) s1 \<longrightarrow>\<^isub>\<beta>\<^sup>* App (Lam [a].t2 ) s2" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   747
    by (force intro!: one_pr_cong one_abs_cong)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   748
  show ?case using c1 c2 by (force intro: rtrancl_trans)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   749
qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   750
 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   751
lemma one_star_abs_cong: 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   752
  fixes a :: "name"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   753
  assumes a: "t1\<longrightarrow>\<^isub>1\<^sup>*t2" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   754
  shows "(Lam  [a].t1)\<longrightarrow>\<^isub>1\<^sup>* (Lam [a].t2)"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   755
  using a
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   756
proof induct
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   757
  case 1 thus ?case by (force)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   758
next
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   759
  case 2 thus ?case by (force intro: rtrancl_trans)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   760
qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   761
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   762
lemma one_star_pr_congL: 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   763
  assumes a: "t1\<longrightarrow>\<^isub>1\<^sup>*t2" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   764
  shows "App t1 s\<longrightarrow>\<^isub>1\<^sup>* App t2 s"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   765
  using a
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   766
proof induct
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   767
  case 1 thus ?case by (force)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   768
next
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   769
  case 2 thus ?case by (force intro: rtrancl_trans)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   770
qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   771
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   772
lemma one_star_pr_congR: 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   773
  assumes a: "t1\<longrightarrow>\<^isub>1\<^sup>*t2" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   774
  shows "App s t1 \<longrightarrow>\<^isub>1\<^sup>* App s t2"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   775
  using a
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   776
proof induct
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   777
  case 1 thus ?case by (force)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   778
next
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   779
  case 2 thus ?case by (force intro: rtrancl_trans)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   780
qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   781
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   782
lemma beta_one_star: 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   783
  assumes a: "t1\<longrightarrow>\<^isub>\<beta>t2" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   784
  shows "t1\<longrightarrow>\<^isub>1\<^sup>*t2"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   785
  using a
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   786
proof induct
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   787
  case b1 thus ?case by (force intro!: one_star_pr_congL)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   788
next
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   789
  case b2 thus ?case by (force intro!: one_star_pr_congR)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   790
next
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   791
  case b3 thus ?case by (force intro!: one_star_abs_cong)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   792
next
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   793
  case b4 thus ?case by (force)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   794
qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   795
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   796
lemma trans_closure: 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   797
  shows "(t1\<longrightarrow>\<^isub>1\<^sup>*t2) = (t1\<longrightarrow>\<^isub>\<beta>\<^sup>*t2)"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   798
proof
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   799
  assume "t1 \<longrightarrow>\<^isub>1\<^sup>* t2"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   800
  thus "t1\<longrightarrow>\<^isub>\<beta>\<^sup>*t2"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   801
  proof induct
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   802
    case 1 thus ?case by (force)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   803
  next
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   804
    case 2 thus ?case by (force intro: rtrancl_trans simp add: one_beta_star)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   805
  qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   806
next
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   807
  assume "t1 \<longrightarrow>\<^isub>\<beta>\<^sup>* t2" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   808
  thus "t1\<longrightarrow>\<^isub>1\<^sup>*t2"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   809
  proof induct
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   810
    case 1 thus ?case by (force)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   811
  next
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   812
    case 2 thus ?case by (force intro: rtrancl_trans simp add: beta_one_star)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   813
  qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   814
qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   815
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   816
lemma cr_one:
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   817
  assumes a: "t\<longrightarrow>\<^isub>1\<^sup>*t1" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   818
  and b: "t\<longrightarrow>\<^isub>1t2"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   819
  shows "\<exists>t3. t1\<longrightarrow>\<^isub>1t3 \<and> t2\<longrightarrow>\<^isub>1\<^sup>*t3"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   820
proof -
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   821
  have stronger: "\<forall>t2. t\<longrightarrow>\<^isub>1t2\<longrightarrow>(\<exists>t3. t1\<longrightarrow>\<^isub>1t3\<and>t2\<longrightarrow>\<^isub>1\<^sup>*t3)"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   822
    using a 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   823
  proof induct
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   824
    case 1 show ?case by (force)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   825
  next
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   826
    case (2 s1 s2)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   827
    assume b: "s1 \<longrightarrow>\<^isub>1 s2"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   828
    assume h: "\<forall>t2. t \<longrightarrow>\<^isub>1 t2 \<longrightarrow> (\<exists>t3. s1 \<longrightarrow>\<^isub>1 t3 \<and> t2 \<longrightarrow>\<^isub>1\<^sup>* t3)"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   829
    show ?case
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   830
    proof (rule allI, rule impI)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   831
      fix t2
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   832
      assume  c: "t \<longrightarrow>\<^isub>1 t2"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   833
      show "(\<exists>t3.  s2 \<longrightarrow>\<^isub>1 t3 \<and>  t2 \<longrightarrow>\<^isub>1\<^sup>* t3)" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   834
      proof -
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   835
	from c h have "(\<exists>t3. s1 \<longrightarrow>\<^isub>1 t3 \<and> t2 \<longrightarrow>\<^isub>1\<^sup>* t3)" by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   836
	then obtain t3 where c1: "s1 \<longrightarrow>\<^isub>1 t3"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   837
                         and c2: "t2 \<longrightarrow>\<^isub>1\<^sup>* t3" by (force)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   838
	have "(\<exists>t4. s2 \<longrightarrow>\<^isub>1 t4 \<and> t3 \<longrightarrow>\<^isub>1 t4)" using b c1 by (force intro: diamond)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   839
	thus ?thesis using c2 by (force intro: rtrancl_trans)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   840
      qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   841
    qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   842
  qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   843
  from a b stronger show ?thesis by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   844
qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   845
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   846
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   847
lemma cr_one_star: 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   848
  assumes a: "t\<longrightarrow>\<^isub>1\<^sup>*t2"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   849
      and b: "t\<longrightarrow>\<^isub>1\<^sup>*t1"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   850
    shows "(\<exists>t3. t1\<longrightarrow>\<^isub>1\<^sup>*t3\<and>t2\<longrightarrow>\<^isub>1\<^sup>*t3)"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   851
using a
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   852
proof induct
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   853
  case 1
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   854
  show ?case using b by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   855
next 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   856
  case (2 s1 s2)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   857
  assume d: "s1 \<longrightarrow>\<^isub>1 s2"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   858
  assume "\<exists>t3.  t1 \<longrightarrow>\<^isub>1\<^sup>* t3 \<and>  s1 \<longrightarrow>\<^isub>1\<^sup>* t3"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   859
  then obtain t3 where f1: "t1 \<longrightarrow>\<^isub>1\<^sup>* t3"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   860
                   and f2: "s1 \<longrightarrow>\<^isub>1\<^sup>* t3" by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   861
  from cr_one d f2 have "\<exists>t4. t3\<longrightarrow>\<^isub>1t4 \<and> s2\<longrightarrow>\<^isub>1\<^sup>*t4" by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   862
  then obtain t4 where g1: "t3\<longrightarrow>\<^isub>1t4"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   863
                   and g2: "s2\<longrightarrow>\<^isub>1\<^sup>*t4" by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   864
  have "t1\<longrightarrow>\<^isub>1\<^sup>*t4" using f1 g1 by (force intro: rtrancl_trans)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   865
  thus ?case using g2 by (force)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   866
qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   867
  
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   868
lemma cr_beta_star: 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   869
  assumes a1: "t\<longrightarrow>\<^isub>\<beta>\<^sup>*t1" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   870
  and a2: "t\<longrightarrow>\<^isub>\<beta>\<^sup>*t2" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   871
  shows "(\<exists>t3. t1\<longrightarrow>\<^isub>\<beta>\<^sup>*t3\<and>t2\<longrightarrow>\<^isub>\<beta>\<^sup>*t3)"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   872
proof -
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   873
  from a1 have b1: "t\<longrightarrow>\<^isub>1\<^sup>*t1" by (simp add: trans_closure[symmetric])
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   874
  from a2 have b2: "t\<longrightarrow>\<^isub>1\<^sup>*t2" by (simp add: trans_closure[symmetric])
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   875
  from b1 and b2 have c: "\<exists>t3. (t1\<longrightarrow>\<^isub>1\<^sup>*t3 \<and> t2\<longrightarrow>\<^isub>1\<^sup>*t3)" by (force intro!: cr_one_star) 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   876
  from c obtain t3 where d1: "t1\<longrightarrow>\<^isub>1\<^sup>*t3" and d2: "t2\<longrightarrow>\<^isub>1\<^sup>*t3" by force
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   877
  from d1 have e1: "t1\<longrightarrow>\<^isub>\<beta>\<^sup>*t3" by (simp add: trans_closure)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   878
  from d2 have e2: "t2\<longrightarrow>\<^isub>\<beta>\<^sup>*t3" by (simp add: trans_closure)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   879
  show ?thesis using e1 and e2 by (force)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   880
qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   881
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   882
end
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   883