src/HOL/Nominal/Examples/SN.thy
author wenzelm
Thu, 03 Jan 2008 22:25:15 +0100
changeset 25819 e6feb08b7f4b
parent 24899 08865bb87098
child 25831 7711d60a5293
permissions -rw-r--r--
replaced thread_properties by simplified version in position.ML;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
18263
7f75925498da cleaned up all examples so that they work with the
urbanc
parents: 18106
diff changeset
     1
(* $Id$ *)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
     2
19496
79dbe35c6cba Capitalized theory names.
berghofe
parents: 19218
diff changeset
     3
theory SN
21107
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 19972
diff changeset
     4
imports Lam_Funs
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
     5
begin
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
     6
18269
3f36e2165e51 some small tuning
urbanc
parents: 18263
diff changeset
     7
text {* Strong Normalisation proof from the Proofs and Types book *}
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
     8
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
     9
section {* Beta Reduction *}
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    10
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
    11
lemma subst_rename: 
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
    12
  assumes a: "c\<sharp>t1"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
    13
  shows "t1[a::=t2] = ([(c,a)]\<bullet>t1)[c::=t2]"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
    14
using a
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
    15
by (nominal_induct t1 avoiding: a c t2 rule: lam.induct)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
    16
   (auto simp add: calc_atm fresh_atm abs_fresh)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    17
18313
e61d2424863d initial cleanup to use the new induction method
urbanc
parents: 18269
diff changeset
    18
lemma forget: 
e61d2424863d initial cleanup to use the new induction method
urbanc
parents: 18269
diff changeset
    19
  assumes a: "a\<sharp>t1"
e61d2424863d initial cleanup to use the new induction method
urbanc
parents: 18269
diff changeset
    20
  shows "t1[a::=t2] = t1"
e61d2424863d initial cleanup to use the new induction method
urbanc
parents: 18269
diff changeset
    21
  using a
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
    22
by (nominal_induct t1 avoiding: a t2 rule: lam.induct)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
    23
   (auto simp add: abs_fresh fresh_atm)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    24
18313
e61d2424863d initial cleanup to use the new induction method
urbanc
parents: 18269
diff changeset
    25
lemma fresh_fact: 
e61d2424863d initial cleanup to use the new induction method
urbanc
parents: 18269
diff changeset
    26
  fixes a::"name"
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
    27
  assumes a: "a\<sharp>t1" "a\<sharp>t2"
22540
e4817fa0f6a1 adapted to new nominal_inductive
urbanc
parents: 22538
diff changeset
    28
  shows "a\<sharp>t1[b::=t2]"
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
    29
using a
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
    30
by (nominal_induct t1 avoiding: a b t2 rule: lam.induct)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
    31
   (auto simp add: abs_fresh fresh_atm)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    32
22540
e4817fa0f6a1 adapted to new nominal_inductive
urbanc
parents: 22538
diff changeset
    33
lemma fresh_fact': 
e4817fa0f6a1 adapted to new nominal_inductive
urbanc
parents: 22538
diff changeset
    34
  fixes a::"name"
e4817fa0f6a1 adapted to new nominal_inductive
urbanc
parents: 22538
diff changeset
    35
  assumes a: "a\<sharp>t2"
e4817fa0f6a1 adapted to new nominal_inductive
urbanc
parents: 22538
diff changeset
    36
  shows "a\<sharp>t1[a::=t2]"
e4817fa0f6a1 adapted to new nominal_inductive
urbanc
parents: 22538
diff changeset
    37
using a 
e4817fa0f6a1 adapted to new nominal_inductive
urbanc
parents: 22538
diff changeset
    38
by (nominal_induct t1 avoiding: a t2 rule: lam.induct)
e4817fa0f6a1 adapted to new nominal_inductive
urbanc
parents: 22538
diff changeset
    39
   (auto simp add: abs_fresh fresh_atm)
e4817fa0f6a1 adapted to new nominal_inductive
urbanc
parents: 22538
diff changeset
    40
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
    41
lemma subst_lemma:  
18313
e61d2424863d initial cleanup to use the new induction method
urbanc
parents: 18269
diff changeset
    42
  assumes a: "x\<noteq>y"
e61d2424863d initial cleanup to use the new induction method
urbanc
parents: 18269
diff changeset
    43
  and     b: "x\<sharp>L"
e61d2424863d initial cleanup to use the new induction method
urbanc
parents: 18269
diff changeset
    44
  shows "M[x::=N][y::=L] = M[y::=L][x::=N[y::=L]]"
e61d2424863d initial cleanup to use the new induction method
urbanc
parents: 18269
diff changeset
    45
using a b
18659
2ff0ae57431d changes to make use of the new induction principle proved by
urbanc
parents: 18654
diff changeset
    46
by (nominal_induct M avoiding: x y N L rule: lam.induct)
18313
e61d2424863d initial cleanup to use the new induction method
urbanc
parents: 18269
diff changeset
    47
   (auto simp add: fresh_fact forget)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    48
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
    49
lemma id_subs: 
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
    50
  shows "t[x::=Var x] = t"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
    51
  by (nominal_induct t avoiding: x rule: lam.induct)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
    52
     (simp_all add: fresh_atm)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
    53
23142
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
    54
lemma psubst_subst:
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
    55
  assumes h:"c\<sharp>\<theta>"
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
    56
  shows "(\<theta><t>)[c::=s] = ((c,s)#\<theta>)<t>"
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
    57
  using h
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
    58
by (nominal_induct t avoiding: \<theta> c s rule: lam.induct)
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
    59
   (auto simp add: fresh_list_cons fresh_atm forget lookup_fresh lookup_fresh')
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
    60
 
23760
aca2c7f80e2f Renamed inductive2 to inductive.
berghofe
parents: 23393
diff changeset
    61
inductive 
23142
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
    62
  Beta :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^isub>\<beta> _" [80,80] 80)
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    63
where
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
    64
  b1[intro!]: "s1 \<longrightarrow>\<^isub>\<beta> s2 \<Longrightarrow> App s1 t \<longrightarrow>\<^isub>\<beta> App s2 t"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
    65
| b2[intro!]: "s1\<longrightarrow>\<^isub>\<beta>s2 \<Longrightarrow> App t s1 \<longrightarrow>\<^isub>\<beta> App t s2"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
    66
| b3[intro!]: "s1\<longrightarrow>\<^isub>\<beta>s2 \<Longrightarrow> Lam [a].s1 \<longrightarrow>\<^isub>\<beta> Lam [a].s2"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
    67
| b4[intro!]: "a\<sharp>s2 \<Longrightarrow> App (Lam [a].s1) s2\<longrightarrow>\<^isub>\<beta> (s1[a::=s2])"
22538
3ccb92dfb5e9 tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents: 22531
diff changeset
    68
22730
8bcc8809ed3b nominal_inductive no longer proves equivariance.
berghofe
parents: 22650
diff changeset
    69
equivariance Beta
8bcc8809ed3b nominal_inductive no longer proves equivariance.
berghofe
parents: 22650
diff changeset
    70
22538
3ccb92dfb5e9 tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents: 22531
diff changeset
    71
nominal_inductive Beta
22540
e4817fa0f6a1 adapted to new nominal_inductive
urbanc
parents: 22538
diff changeset
    72
  by (simp_all add: abs_fresh fresh_fact')
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    73
18378
urbanc
parents: 18348
diff changeset
    74
lemma supp_beta: 
urbanc
parents: 18348
diff changeset
    75
  assumes a: "t\<longrightarrow>\<^isub>\<beta> s"
urbanc
parents: 18348
diff changeset
    76
  shows "(supp s)\<subseteq>((supp t)::name set)"
urbanc
parents: 18348
diff changeset
    77
using a
urbanc
parents: 18348
diff changeset
    78
by (induct)
urbanc
parents: 18348
diff changeset
    79
   (auto intro!: simp add: abs_supp lam.supp subst_supp)
urbanc
parents: 18348
diff changeset
    80
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
    81
lemma beta_abs: 
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
    82
  assumes a: "Lam [a].t\<longrightarrow>\<^isub>\<beta> t'"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
    83
  shows "\<exists>t''. t'=Lam [a].t'' \<and> t\<longrightarrow>\<^isub>\<beta> t''"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
    84
using a
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
    85
apply -
23760
aca2c7f80e2f Renamed inductive2 to inductive.
berghofe
parents: 23393
diff changeset
    86
apply(ind_cases "Lam [a].t  \<longrightarrow>\<^isub>\<beta> t'")
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    87
apply(auto simp add: lam.distinct lam.inject)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    88
apply(auto simp add: alpha)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    89
apply(rule_tac x="[(a,aa)]\<bullet>s2" in exI)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    90
apply(rule conjI)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    91
apply(rule sym)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    92
apply(rule pt_bij2[OF pt_name_inst, OF at_name_inst])
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    93
apply(simp)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    94
apply(rule pt_name3)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    95
apply(simp add: at_ds5[OF at_name_inst])
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    96
apply(rule conjI)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    97
apply(simp add: pt_fresh_left[OF pt_name_inst, OF at_name_inst] calc_atm)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    98
apply(force dest!: supp_beta simp add: fresh_def)
22541
c33b542394f3 the name for the collection of equivariance lemmas is now eqvts (changed from eqvt) in order to avoid clashes with eqvt-lemmas generated in nominal_inductive
urbanc
parents: 22540
diff changeset
    99
apply(force intro!: eqvts)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   100
done
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   101
18313
e61d2424863d initial cleanup to use the new induction method
urbanc
parents: 18269
diff changeset
   102
lemma beta_subst: 
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   103
  assumes a: "M \<longrightarrow>\<^isub>\<beta> M'"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   104
  shows "M[x::=N]\<longrightarrow>\<^isub>\<beta> M'[x::=N]" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   105
using a
23142
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   106
by (nominal_induct M M' avoiding: x N rule: Beta.strong_induct)
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   107
   (auto simp add: fresh_atm subst_lemma fresh_fact)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   108
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   109
section {* types *}
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   110
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
   111
nominal_datatype ty =
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
   112
    TVar "nat"
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   113
  | TArr "ty" "ty" (infix "\<rightarrow>" 200)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   114
23142
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   115
lemma perm_ty:
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   116
  fixes pi ::"name prm"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   117
  and   \<tau>  ::"ty"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   118
  shows "pi\<bullet>\<tau> = \<tau>"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
   119
by (nominal_induct \<tau> rule: ty.induct) 
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
   120
   (simp_all add: perm_nat_def)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   121
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   122
lemma fresh_ty:
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   123
  fixes a ::"name"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   124
  and   \<tau>  ::"ty"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   125
  shows "a\<sharp>\<tau>"
23142
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   126
  by (simp add: fresh_def perm_ty supp_def)
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   127
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   128
(* domain of a typing context *)
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   129
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   130
fun
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   131
  "dom_ty" :: "(name\<times>ty) list \<Rightarrow> (name list)"
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   132
where
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   133
  "dom_ty []    = []"
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   134
| "dom_ty ((x,\<tau>)#\<Gamma>) = (x)#(dom_ty \<Gamma>)" 
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   135
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   136
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   137
(* valid contexts *)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   138
23760
aca2c7f80e2f Renamed inductive2 to inductive.
berghofe
parents: 23393
diff changeset
   139
inductive 
23142
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   140
  valid :: "(name\<times>ty) list \<Rightarrow> bool"
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   141
where
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   142
  v1[intro]: "valid []"
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   143
| v2[intro]: "\<lbrakk>valid \<Gamma>;a\<sharp>\<Gamma>\<rbrakk>\<Longrightarrow> valid ((a,\<sigma>)#\<Gamma>)"
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   144
22538
3ccb92dfb5e9 tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents: 22531
diff changeset
   145
equivariance valid 
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   146
23760
aca2c7f80e2f Renamed inductive2 to inductive.
berghofe
parents: 23393
diff changeset
   147
inductive_cases valid_elim[elim]: "valid ((a,\<tau>)#\<Gamma>)"
23142
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   148
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   149
(* typing judgements *)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   150
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
   151
lemma fresh_context: 
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   152
  fixes  \<Gamma> :: "(name\<times>ty)list"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   153
  and    a :: "name"
18378
urbanc
parents: 18348
diff changeset
   154
  assumes a: "a\<sharp>\<Gamma>"
urbanc
parents: 18348
diff changeset
   155
  shows "\<not>(\<exists>\<tau>::ty. (a,\<tau>)\<in>set \<Gamma>)"
urbanc
parents: 18348
diff changeset
   156
using a
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   157
by (induct \<Gamma>)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   158
   (auto simp add: fresh_prod fresh_list_cons fresh_atm)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   159
23760
aca2c7f80e2f Renamed inductive2 to inductive.
berghofe
parents: 23393
diff changeset
   160
inductive 
23142
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   161
  typing :: "(name\<times>ty) list\<Rightarrow>lam\<Rightarrow>ty\<Rightarrow>bool" ("_ \<turnstile> _ : _" [60,60,60] 60)
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   162
where
22650
0c5b22076fb3 tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
urbanc
parents: 22542
diff changeset
   163
  t1[intro]: "\<lbrakk>valid \<Gamma>; (a,\<tau>)\<in>set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var a : \<tau>"
0c5b22076fb3 tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
urbanc
parents: 22542
diff changeset
   164
| t2[intro]: "\<lbrakk>\<Gamma> \<turnstile> t1 : \<tau>\<rightarrow>\<sigma>; \<Gamma> \<turnstile> t2 : \<tau>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App t1 t2 : \<sigma>"
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   165
| t3[intro]: "\<lbrakk>a\<sharp>\<Gamma>;((a,\<tau>)#\<Gamma>) \<turnstile> t : \<sigma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [a].t : \<tau>\<rightarrow>\<sigma>"
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   166
22730
8bcc8809ed3b nominal_inductive no longer proves equivariance.
berghofe
parents: 22650
diff changeset
   167
equivariance typing
8bcc8809ed3b nominal_inductive no longer proves equivariance.
berghofe
parents: 22650
diff changeset
   168
22538
3ccb92dfb5e9 tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents: 22531
diff changeset
   169
nominal_inductive typing
3ccb92dfb5e9 tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents: 22531
diff changeset
   170
  by (simp_all add: abs_fresh fresh_ty)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   171
21107
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 19972
diff changeset
   172
abbreviation
22650
0c5b22076fb3 tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
urbanc
parents: 22542
diff changeset
   173
  "sub" :: "(name\<times>ty) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" ("_ \<subseteq> _" [60,60] 60) 
0c5b22076fb3 tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
urbanc
parents: 22542
diff changeset
   174
where
0c5b22076fb3 tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
urbanc
parents: 22542
diff changeset
   175
  "\<Gamma>1 \<subseteq> \<Gamma>2 \<equiv> \<forall>a \<sigma>. (a,\<sigma>)\<in>set \<Gamma>1 \<longrightarrow>  (a,\<sigma>)\<in>set \<Gamma>2"
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   176
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   177
subsection {* some facts about beta *}
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   178
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   179
constdefs
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   180
  "NORMAL" :: "lam \<Rightarrow> bool"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   181
  "NORMAL t \<equiv> \<not>(\<exists>t'. t\<longrightarrow>\<^isub>\<beta> t')"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   182
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   183
lemma NORMAL_Var:
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   184
  shows "NORMAL (Var a)"
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   185
proof -
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   186
  { assume "\<exists>t'. (Var a) \<longrightarrow>\<^isub>\<beta> t'"
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   187
    then obtain t' where "(Var a) \<longrightarrow>\<^isub>\<beta> t'" by blast
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   188
    hence False by (cases, auto) 
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   189
  }
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   190
  thus "NORMAL (Var a)" by (force simp add: NORMAL_def)
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   191
qed
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   192
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   193
inductive 
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   194
  SN :: "lam \<Rightarrow> bool"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   195
where
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   196
  SN_intro: "(\<And>t'. t \<longrightarrow>\<^isub>\<beta> t' \<Longrightarrow> SN t') \<Longrightarrow> SN t"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   197
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   198
lemma SN_elim:
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   199
  assumes a: "SN M"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   200
  shows "(\<forall>M. (\<forall>N. M \<longrightarrow>\<^isub>\<beta> N \<longrightarrow> P N)\<longrightarrow> P M) \<longrightarrow> P M"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   201
using a
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   202
by (induct rule: SN.SN.induct) (blast)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   203
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   204
lemma SN_preserved: 
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   205
  assumes a: "SN t1" "t1\<longrightarrow>\<^isub>\<beta> t2"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   206
  shows "SN t2"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   207
using a 
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   208
by (cases) (auto)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   209
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   210
lemma double_SN_aux:
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   211
  assumes a: "SN a"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   212
  and b: "SN b"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   213
  and hyp: "\<And>x z.
24899
08865bb87098 Isar-fied many proofs
urbanc
parents: 23970
diff changeset
   214
    \<lbrakk>\<And>y. x \<longrightarrow>\<^isub>\<beta> y \<Longrightarrow> SN y; \<And>y. x \<longrightarrow>\<^isub>\<beta> y \<Longrightarrow> P y z;
08865bb87098 Isar-fied many proofs
urbanc
parents: 23970
diff changeset
   215
     \<And>u. z \<longrightarrow>\<^isub>\<beta> u \<Longrightarrow> SN u; \<And>u. z \<longrightarrow>\<^isub>\<beta> u \<Longrightarrow> P x u\<rbrakk> \<Longrightarrow> P x z"
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   216
  shows "P a b"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   217
proof -
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   218
  from a
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   219
  have r: "\<And>b. SN b \<Longrightarrow> P a b"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   220
  proof (induct a rule: SN.SN.induct)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   221
    case (SN_intro x)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   222
    note SNI' = SN_intro
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   223
    have "SN b" by fact
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   224
    thus ?case
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   225
    proof (induct b rule: SN.SN.induct)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   226
      case (SN_intro y)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   227
      show ?case
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   228
	apply (rule hyp)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   229
	apply (erule SNI')
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   230
	apply (erule SNI')
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   231
	apply (rule SN.SN_intro)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   232
	apply (erule SN_intro)+
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   233
	done
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   234
    qed
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   235
  qed
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   236
  from b show ?thesis by (rule r)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   237
qed
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   238
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   239
lemma double_SN[consumes 2]:
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   240
  assumes a: "SN a"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   241
  and     b: "SN b" 
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   242
  and     c: "\<And>x z. \<lbrakk>\<And>y. x \<longrightarrow>\<^isub>\<beta> y \<Longrightarrow> P y z; \<And>u. z \<longrightarrow>\<^isub>\<beta> u \<Longrightarrow> P x u\<rbrakk> \<Longrightarrow> P x z"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   243
  shows "P a b"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   244
using a b c
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   245
apply(rule_tac double_SN_aux)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   246
apply(assumption)+
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   247
apply(blast)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   248
done
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   249
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   250
section {* Candidates *}
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   251
22440
7e4f4f19002f deleted function for defining candidates and used nominal_primrec instead
urbanc
parents: 22420
diff changeset
   252
consts
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   253
  RED :: "ty \<Rightarrow> lam set"
22440
7e4f4f19002f deleted function for defining candidates and used nominal_primrec instead
urbanc
parents: 22420
diff changeset
   254
7e4f4f19002f deleted function for defining candidates and used nominal_primrec instead
urbanc
parents: 22420
diff changeset
   255
nominal_primrec
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
   256
  "RED (TVar X) = {t. SN(t)}"
22440
7e4f4f19002f deleted function for defining candidates and used nominal_primrec instead
urbanc
parents: 22420
diff changeset
   257
  "RED (\<tau>\<rightarrow>\<sigma>) =   {t. \<forall>u. (u\<in>RED \<tau> \<longrightarrow> (App t u)\<in>RED \<sigma>)}"
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   258
by (rule TrueI)+
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
   259
24899
08865bb87098 Isar-fied many proofs
urbanc
parents: 23970
diff changeset
   260
(* neutral terms *)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   261
constdefs
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   262
  NEUT :: "lam \<Rightarrow> bool"
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   263
  "NEUT t \<equiv> (\<exists>a. t = Var a) \<or> (\<exists>t1 t2. t = App t1 t2)" 
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   264
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   265
(* a slight hack to get the first element of applications *)
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   266
(* this is needed to get (SN t) from SN (App t s)         *) 
23760
aca2c7f80e2f Renamed inductive2 to inductive.
berghofe
parents: 23393
diff changeset
   267
inductive 
23142
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   268
  FST :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<guillemotright> _" [80,80] 80)
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   269
where
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
   270
  fst[intro!]:  "(App t s) \<guillemotright> t"
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   271
24899
08865bb87098 Isar-fied many proofs
urbanc
parents: 23970
diff changeset
   272
consts
08865bb87098 Isar-fied many proofs
urbanc
parents: 23970
diff changeset
   273
  fst_app_aux::"lam\<Rightarrow>lam option"
08865bb87098 Isar-fied many proofs
urbanc
parents: 23970
diff changeset
   274
08865bb87098 Isar-fied many proofs
urbanc
parents: 23970
diff changeset
   275
nominal_primrec
08865bb87098 Isar-fied many proofs
urbanc
parents: 23970
diff changeset
   276
  "fst_app_aux (Var a)     = None"
08865bb87098 Isar-fied many proofs
urbanc
parents: 23970
diff changeset
   277
  "fst_app_aux (App t1 t2) = Some t1"
08865bb87098 Isar-fied many proofs
urbanc
parents: 23970
diff changeset
   278
  "fst_app_aux (Lam [x].t) = None"
08865bb87098 Isar-fied many proofs
urbanc
parents: 23970
diff changeset
   279
apply(finite_guess)+
08865bb87098 Isar-fied many proofs
urbanc
parents: 23970
diff changeset
   280
apply(rule TrueI)+
08865bb87098 Isar-fied many proofs
urbanc
parents: 23970
diff changeset
   281
apply(simp add: fresh_none)
08865bb87098 Isar-fied many proofs
urbanc
parents: 23970
diff changeset
   282
apply(fresh_guess)+
08865bb87098 Isar-fied many proofs
urbanc
parents: 23970
diff changeset
   283
done
08865bb87098 Isar-fied many proofs
urbanc
parents: 23970
diff changeset
   284
08865bb87098 Isar-fied many proofs
urbanc
parents: 23970
diff changeset
   285
definition
08865bb87098 Isar-fied many proofs
urbanc
parents: 23970
diff changeset
   286
  fst_app_def[simp]: "fst_app t = the (fst_app_aux t)"
08865bb87098 Isar-fied many proofs
urbanc
parents: 23970
diff changeset
   287
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   288
lemma SN_of_FST_of_App: 
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   289
  assumes a: "SN (App t s)"
24899
08865bb87098 Isar-fied many proofs
urbanc
parents: 23970
diff changeset
   290
  shows "SN (fst_app (App t s))"
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   291
using a
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   292
proof - 
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   293
  from a have "\<forall>z. (App t s \<guillemotright> z) \<longrightarrow> SN z"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   294
    by (induct rule: SN.SN.induct)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   295
       (blast elim: FST.cases intro: SN_intro)
24899
08865bb87098 Isar-fied many proofs
urbanc
parents: 23970
diff changeset
   296
  then have "SN t" by blast
08865bb87098 Isar-fied many proofs
urbanc
parents: 23970
diff changeset
   297
  then show "SN (fst_app (App t s))" by simp
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   298
qed
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   299
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   300
section {* Candidates *}
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   301
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   302
constdefs
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   303
  "CR1" :: "ty \<Rightarrow> bool"
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   304
  "CR1 \<tau> \<equiv> \<forall>t. (t\<in>RED \<tau> \<longrightarrow> SN t)"
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   305
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   306
  "CR2" :: "ty \<Rightarrow> bool"
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   307
  "CR2 \<tau> \<equiv> \<forall>t t'. (t\<in>RED \<tau> \<and> t \<longrightarrow>\<^isub>\<beta> t') \<longrightarrow> t'\<in>RED \<tau>"
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   308
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   309
  "CR3_RED" :: "lam \<Rightarrow> ty \<Rightarrow> bool"
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   310
  "CR3_RED t \<tau> \<equiv> \<forall>t'. t\<longrightarrow>\<^isub>\<beta> t' \<longrightarrow>  t'\<in>RED \<tau>" 
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   311
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   312
  "CR3" :: "ty \<Rightarrow> bool"
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   313
  "CR3 \<tau> \<equiv> \<forall>t. (NEUT t \<and> CR3_RED t \<tau>) \<longrightarrow> t\<in>RED \<tau>"
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   314
   
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   315
  "CR4" :: "ty \<Rightarrow> bool"
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   316
  "CR4 \<tau> \<equiv> \<forall>t. (NEUT t \<and> NORMAL t) \<longrightarrow>t\<in>RED \<tau>"
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   317
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   318
lemma CR3_implies_CR4: 
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   319
  assumes a: "CR3 \<tau>" 
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   320
  shows "CR4 \<tau>"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   321
using a by (auto simp add: CR3_def CR3_RED_def CR4_def NORMAL_def)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   322
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   323
(* sub_induction in the arrow-type case for the next proof *) 
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   324
lemma sub_induction: 
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   325
  assumes a: "SN(u)"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   326
  and     b: "u\<in>RED \<tau>"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   327
  and     c1: "NEUT t"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   328
  and     c2: "CR2 \<tau>"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   329
  and     c3: "CR3 \<sigma>"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   330
  and     c4: "CR3_RED t (\<tau>\<rightarrow>\<sigma>)"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   331
  shows "(App t u)\<in>RED \<sigma>"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   332
using a b
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   333
proof (induct)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   334
  fix u
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   335
  assume as: "u\<in>RED \<tau>"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   336
  assume ih: " \<And>u'. \<lbrakk>u \<longrightarrow>\<^isub>\<beta> u'; u' \<in> RED \<tau>\<rbrakk> \<Longrightarrow> App t u' \<in> RED \<sigma>"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   337
  have "NEUT (App t u)" using c1 by (auto simp add: NEUT_def)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   338
  moreover
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   339
  have "CR3_RED (App t u) \<sigma>" unfolding CR3_RED_def
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   340
  proof (intro strip)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   341
    fix r
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   342
    assume red: "App t u \<longrightarrow>\<^isub>\<beta> r"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   343
    moreover
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   344
    { assume "\<exists>t'. t \<longrightarrow>\<^isub>\<beta> t' \<and> r = App t' u"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   345
      then obtain t' where a1: "t \<longrightarrow>\<^isub>\<beta> t'" and a2: "r = App t' u" by blast
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   346
      have "t'\<in>RED (\<tau>\<rightarrow>\<sigma>)" using c4 a1 by (simp add: CR3_RED_def)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   347
      then have "App t' u\<in>RED \<sigma>" using as by simp
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   348
      then have "r\<in>RED \<sigma>" using a2 by simp
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   349
    }
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   350
    moreover
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   351
    { assume "\<exists>u'. u \<longrightarrow>\<^isub>\<beta> u' \<and> r = App t u'"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   352
      then obtain u' where b1: "u \<longrightarrow>\<^isub>\<beta> u'" and b2: "r = App t u'" by blast
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   353
      have "u'\<in>RED \<tau>" using as b1 c2 by (auto simp add: CR2_def)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   354
      with ih have "App t u' \<in> RED \<sigma>" using b1 by simp
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   355
      then have "r\<in>RED \<sigma>" using b2 by simp
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   356
    }
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   357
    moreover
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   358
    { assume "\<exists>x t'. t = Lam [x].t'"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   359
      then obtain x t' where "t = Lam [x].t'" by blast
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   360
      then have "NEUT (Lam [x].t')" using c1 by simp
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   361
      then have "False" by (simp add: NEUT_def)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   362
      then have "r\<in>RED \<sigma>" by simp
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   363
    }
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   364
    ultimately show "r \<in> RED \<sigma>" by (cases) (auto simp add: lam.inject)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   365
  qed
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   366
  ultimately show "App t u \<in> RED \<sigma>" using c3 by (simp add: CR3_def)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   367
qed 
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   368
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   369
(* properties of the candiadates *)
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   370
lemma RED_props: 
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   371
  shows "CR1 \<tau>" and "CR2 \<tau>" and "CR3 \<tau>"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
   372
proof (nominal_induct \<tau> rule: ty.induct)
18611
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   373
  case (TVar a)
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   374
  { case 1 show "CR1 (TVar a)" by (simp add: CR1_def)
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   375
  next
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   376
    case 2 show "CR2 (TVar a)" by (auto intro: SN_preserved simp add: CR2_def)
18611
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   377
  next
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   378
    case 3 show "CR3 (TVar a)" by (auto intro: SN_intro simp add: CR3_def CR3_RED_def)
18611
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   379
  }
18599
e01112713fdc changed PRO_RED proof to conform with the new induction rules
urbanc
parents: 18383
diff changeset
   380
next
18611
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   381
  case (TArr \<tau>1 \<tau>2)
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   382
  { case 1
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   383
    have ih_CR3_\<tau>1: "CR3 \<tau>1" by fact
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   384
    have ih_CR1_\<tau>2: "CR1 \<tau>2" by fact
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   385
    show "CR1 (\<tau>1 \<rightarrow> \<tau>2)" unfolding CR1_def
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   386
    proof (simp, intro strip)
18611
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   387
      fix t
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   388
      assume a: "\<forall>u. u \<in> RED \<tau>1 \<longrightarrow> App t u \<in> RED \<tau>2"
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   389
      from ih_CR3_\<tau>1 have "CR4 \<tau>1" by (simp add: CR3_implies_CR4) 
18611
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   390
      moreover
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   391
      have "NEUT (Var a)" by (force simp add: NEUT_def)
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   392
      moreover
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   393
      have "NORMAL (Var a)" by (rule NORMAL_Var)
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   394
      ultimately have "(Var a)\<in> RED \<tau>1" by (simp add: CR4_def)
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   395
      with a have "App t (Var a) \<in> RED \<tau>2" by simp
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   396
      hence "SN (App t (Var a))" using ih_CR1_\<tau>2 by (simp add: CR1_def)
24899
08865bb87098 Isar-fied many proofs
urbanc
parents: 23970
diff changeset
   397
      thus "SN(t)" by (auto dest: SN_of_FST_of_App)
18611
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   398
    qed
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   399
  next
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   400
    case 2
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   401
    have ih_CR1_\<tau>1: "CR1 \<tau>1" by fact
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   402
    have ih_CR2_\<tau>2: "CR2 \<tau>2" by fact
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   403
    show "CR2 (\<tau>1 \<rightarrow> \<tau>2)" unfolding CR2_def
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   404
    proof (simp, intro strip)
18611
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   405
      fix t1 t2 u
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   406
      assume "(\<forall>u. u \<in> RED \<tau>1 \<longrightarrow> App t1 u \<in> RED \<tau>2) \<and>  t1 \<longrightarrow>\<^isub>\<beta> t2" 
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   407
	and  "u \<in> RED \<tau>1"
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   408
      hence "t1 \<longrightarrow>\<^isub>\<beta> t2" and "App t1 u \<in> RED \<tau>2" by simp_all
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   409
      thus "App t2 u \<in> RED \<tau>2" using ih_CR2_\<tau>2 by (auto simp add: CR2_def)
18611
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   410
    qed
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   411
  next
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   412
    case 3
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   413
    have ih_CR1_\<tau>1: "CR1 \<tau>1" by fact
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   414
    have ih_CR2_\<tau>1: "CR2 \<tau>1" by fact
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   415
    have ih_CR3_\<tau>2: "CR3 \<tau>2" by fact
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   416
    show "CR3 (\<tau>1 \<rightarrow> \<tau>2)" unfolding CR3_def
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   417
    proof (simp, intro strip)
18611
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   418
      fix t u
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   419
      assume a1: "u \<in> RED \<tau>1"
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   420
      assume a2: "NEUT t \<and> CR3_RED t (\<tau>1 \<rightarrow> \<tau>2)"
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   421
      have "SN(u)" using a1 ih_CR1_\<tau>1 by (simp add: CR1_def)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   422
      then show "(App t u)\<in>RED \<tau>2" using ih_CR2_\<tau>1 ih_CR3_\<tau>2 a1 a2 by (blast intro: sub_induction)
18611
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   423
    qed
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   424
  }
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   425
qed
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   426
   
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   427
(* not as simple as on paper, because of the stronger double_SN induction *) 
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   428
lemma abs_RED: 
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   429
  assumes asm: "\<forall>s\<in>RED \<tau>. t[x::=s]\<in>RED \<sigma>"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   430
  shows "Lam [x].t\<in>RED (\<tau>\<rightarrow>\<sigma>)"
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   431
proof -
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   432
  have b1: "SN t" 
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   433
  proof -
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   434
    have "Var x\<in>RED \<tau>"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   435
    proof -
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   436
      have "CR4 \<tau>" by (simp add: RED_props CR3_implies_CR4)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   437
      moreover
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   438
      have "NEUT (Var x)" by (auto simp add: NEUT_def)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   439
      moreover
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   440
      have "NORMAL (Var x)" by (auto elim: Beta.cases simp add: NORMAL_def)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   441
      ultimately show "Var x\<in>RED \<tau>" by (simp add: CR4_def)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   442
    qed
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   443
    then have "t[x::=Var x]\<in>RED \<sigma>" using asm by simp
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   444
    then have "t\<in>RED \<sigma>" by (simp add: id_subs)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   445
    moreover 
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   446
    have "CR1 \<sigma>" by (simp add: RED_props)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   447
    ultimately show "SN t" by (simp add: CR1_def)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   448
  qed
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   449
  show "Lam [x].t\<in>RED (\<tau>\<rightarrow>\<sigma>)"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   450
  proof (simp, intro strip)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   451
    fix u
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   452
    assume b2: "u\<in>RED \<tau>"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   453
    then have b3: "SN u" using RED_props by (auto simp add: CR1_def)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   454
    show "App (Lam [x].t) u \<in> RED \<sigma>" using b1 b3 b2 asm
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   455
    proof(induct t u rule: double_SN)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   456
      fix t u
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   457
      assume ih1: "\<And>t'.  \<lbrakk>t \<longrightarrow>\<^isub>\<beta> t'; u\<in>RED \<tau>; \<forall>s\<in>RED \<tau>. t'[x::=s]\<in>RED \<sigma>\<rbrakk> \<Longrightarrow> App (Lam [x].t') u \<in> RED \<sigma>" 
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   458
      assume ih2: "\<And>u'.  \<lbrakk>u \<longrightarrow>\<^isub>\<beta> u'; u'\<in>RED \<tau>; \<forall>s\<in>RED \<tau>. t[x::=s]\<in>RED \<sigma>\<rbrakk> \<Longrightarrow> App (Lam [x].t) u' \<in> RED \<sigma>"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   459
      assume as1: "u \<in> RED \<tau>"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   460
      assume as2: "\<forall>s\<in>RED \<tau>. t[x::=s]\<in>RED \<sigma>"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   461
      have "CR3_RED (App (Lam [x].t) u) \<sigma>" unfolding CR3_RED_def
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   462
      proof(intro strip)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   463
	fix r
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   464
	assume red: "App (Lam [x].t) u \<longrightarrow>\<^isub>\<beta> r"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   465
	moreover
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   466
	{ assume "\<exists>t'. t \<longrightarrow>\<^isub>\<beta> t' \<and> r = App (Lam [x].t') u"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   467
	  then obtain t' where a1: "t \<longrightarrow>\<^isub>\<beta> t'" and a2: "r = App (Lam [x].t') u" by blast
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   468
	  have "App (Lam [x].t') u\<in>RED \<sigma>" using ih1 a1 as1 as2
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   469
	    apply(auto)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   470
	    apply(drule_tac x="t'" in meta_spec)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   471
	    apply(simp)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   472
	    apply(drule meta_mp)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   473
	    apply(auto)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   474
	    apply(drule_tac x="s" in bspec)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   475
	    apply(simp)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   476
	    apply(subgoal_tac "CR2 \<sigma>")
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   477
	    apply(unfold CR2_def)[1]
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   478
	    apply(drule_tac x="t[x::=s]" in spec)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   479
	    apply(drule_tac x="t'[x::=s]" in spec)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   480
	    apply(simp add: beta_subst)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   481
	    apply(simp add: RED_props)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   482
	    done
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   483
	  then have "r\<in>RED \<sigma>" using a2 by simp
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   484
	}
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   485
	moreover
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   486
	{ assume "\<exists>u'. u \<longrightarrow>\<^isub>\<beta> u' \<and> r = App (Lam [x].t) u'"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   487
	  then obtain u' where b1: "u \<longrightarrow>\<^isub>\<beta> u'" and b2: "r = App (Lam [x].t) u'" by blast
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   488
	  have "App (Lam [x].t) u'\<in>RED \<sigma>" using ih2 b1 as1 as2
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   489
	    apply(auto)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   490
	    apply(drule_tac x="u'" in meta_spec)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   491
	    apply(simp)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   492
	    apply(drule meta_mp)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   493
	    apply(subgoal_tac "CR2 \<tau>")
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   494
	    apply(unfold CR2_def)[1]
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   495
	    apply(drule_tac x="u" in spec)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   496
	    apply(drule_tac x="u'" in spec)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   497
	    apply(simp)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   498
	    apply(simp add: RED_props)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   499
	    apply(simp)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   500
	    done
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   501
	  then have "r\<in>RED \<sigma>" using b2 by simp
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   502
	}
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   503
	moreover
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   504
	{ assume "r = t[x::=u]"
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   505
	  then have "r\<in>RED \<sigma>" using as1 as2 by auto
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   506
	}
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   507
	ultimately show "r \<in> RED \<sigma>" 
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   508
	apply(cases) 
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   509
	apply(auto simp add: lam.inject)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   510
	apply(drule beta_abs)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   511
	apply(auto simp add: alpha subst_rename)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   512
	done
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   513
    qed
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   514
    moreover 
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   515
    have "NEUT (App (Lam [x].t) u)" unfolding NEUT_def by (auto)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   516
    ultimately show "App (Lam [x].t) u \<in> RED \<sigma>"  using RED_props by (simp add: CR3_def)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   517
  qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   518
qed
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   519
qed
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   520
22420
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   521
abbreviation 
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   522
 mapsto :: "(name\<times>lam) list \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> bool" ("_ maps _ to _" [55,55,55] 55) 
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   523
where
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   524
 "\<theta> maps x to e\<equiv> (lookup \<theta> x) = e"
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   525
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   526
abbreviation 
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   527
  closes :: "(name\<times>lam) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" ("_ closes _" [55,55] 55) 
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   528
where
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   529
  "\<theta> closes \<Gamma> \<equiv> \<forall>x T. ((x,T) \<in> set \<Gamma> \<longrightarrow> (\<exists>t. \<theta> maps x to t \<and> t \<in> RED T))"
21107
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 19972
diff changeset
   530
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   531
lemma all_RED: 
22420
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   532
  assumes a: "\<Gamma> \<turnstile> t : \<tau>"
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   533
  and     b: "\<theta> closes \<Gamma>"
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   534
  shows "\<theta><t> \<in> RED \<tau>"
18345
d37fb71754fe added an Isar-proof for the abs_ALL lemma
urbanc
parents: 18313
diff changeset
   535
using a b
23142
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   536
proof(nominal_induct  avoiding: \<theta> rule: typing.strong_induct)
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   537
  case (t3 a \<Gamma> \<sigma> t \<tau> \<theta>) --"lambda case"
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   538
  have ih: "\<And>\<theta>. \<theta> closes ((a,\<sigma>)#\<Gamma>) \<Longrightarrow> \<theta><t> \<in> RED \<tau>" by fact
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   539
  have \<theta>_cond: "\<theta> closes \<Gamma>" by fact
23393
31781b2de73d tuned proofs: avoid implicit prems;
wenzelm
parents: 23142
diff changeset
   540
  have fresh: "a\<sharp>\<Gamma>" "a\<sharp>\<theta>" by fact+
24899
08865bb87098 Isar-fied many proofs
urbanc
parents: 23970
diff changeset
   541
  from ih have "\<forall>s\<in>RED \<sigma>. ((a,s)#\<theta>)<t> \<in> RED \<tau>" using fresh \<theta>_cond fresh_context by simp
08865bb87098 Isar-fied many proofs
urbanc
parents: 23970
diff changeset
   542
  then have "\<forall>s\<in>RED \<sigma>. \<theta><t>[a::=s] \<in> RED \<tau>" using fresh by (simp add: psubst_subst)
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   543
  then have "Lam [a].(\<theta><t>) \<in> RED (\<sigma> \<rightarrow> \<tau>)" by (simp only: abs_RED)
23142
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   544
  then show "\<theta><(Lam [a].t)> \<in> RED (\<sigma> \<rightarrow> \<tau>)" using fresh by simp
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   545
qed auto
18345
d37fb71754fe added an Isar-proof for the abs_ALL lemma
urbanc
parents: 18313
diff changeset
   546
23142
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   547
section {* identity substitution generated from a context \<Gamma> *}
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   548
fun
18382
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   549
  "id" :: "(name\<times>ty) list \<Rightarrow> (name\<times>lam) list"
23142
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   550
where
18382
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   551
  "id []    = []"
23142
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   552
| "id ((x,\<tau>)#\<Gamma>) = (x,Var x)#(id \<Gamma>)"
18382
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   553
23142
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   554
lemma id_maps:
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   555
  shows "(id \<Gamma>) maps a to (Var a)"
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   556
by (induct \<Gamma>) (auto)
18382
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   557
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   558
lemma id_fresh:
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   559
  fixes a::"name"
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   560
  assumes a: "a\<sharp>\<Gamma>"
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   561
  shows "a\<sharp>(id \<Gamma>)"
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   562
using a
23142
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   563
by (induct \<Gamma>)
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   564
   (auto simp add: fresh_list_nil fresh_list_cons)
18382
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   565
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   566
lemma id_apply:  
22420
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   567
  shows "(id \<Gamma>)<t> = t"
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   568
by (nominal_induct t avoiding: \<Gamma> rule: lam.induct)
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   569
   (auto simp add: id_maps id_fresh)
18382
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   570
23142
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   571
lemma id_closes:
22420
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   572
  shows "(id \<Gamma>) closes \<Gamma>"
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   573
apply(auto)
23142
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   574
apply(simp add: id_maps)
22420
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   575
apply(subgoal_tac "CR3 T") --"A"
23970
a252a7da82b9 cleaned up the proofs a bit
urbanc
parents: 23760
diff changeset
   576
apply(drule CR3_implies_CR4)
18382
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   577
apply(simp add: CR4_def)
22420
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   578
apply(drule_tac x="Var x" in spec)
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   579
apply(force simp add: NEUT_def NORMAL_Var)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
   580
--"A"
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   581
apply(rule RED_props)
18382
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   582
done
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   583
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   584
lemma typing_implies_RED:  
23142
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   585
  assumes a: "\<Gamma> \<turnstile> t : \<tau>"
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   586
  shows "t \<in> RED \<tau>"
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   587
proof -
22420
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   588
  have "(id \<Gamma>)<t>\<in>RED \<tau>" 
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   589
  proof -
23142
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   590
    have "(id \<Gamma>) closes \<Gamma>" by (rule id_closes)
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   591
    with a show ?thesis by (rule all_RED)
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   592
  qed
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   593
  thus"t \<in> RED \<tau>" by (simp add: id_apply)
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   594
qed
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   595
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   596
lemma typing_implies_SN: 
23142
cb1dbe64a4d5 tuned the proof
urbanc
parents: 22730
diff changeset
   597
  assumes a: "\<Gamma> \<turnstile> t : \<tau>"
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   598
  shows "SN(t)"
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   599
proof -
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   600
  from a have "t \<in> RED \<tau>" by (rule typing_implies_RED)
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   601
  moreover
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   602
  have "CR1 \<tau>" by (rule RED_props)
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   603
  ultimately show "SN(t)" by (simp add: CR1_def)
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   604
qed
18382
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   605
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   606
end