src/HOL/Nominal/Examples/SN.thy
author urbanc
Wed, 28 Mar 2007 17:27:44 +0200
changeset 22540 e4817fa0f6a1
parent 22538 3ccb92dfb5e9
child 22541 c33b542394f3
permissions -rw-r--r--
adapted to new nominal_inductive
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"
e61d2424863d initial cleanup to use the new induction method
urbanc
parents: 18269
diff changeset
    27
  assumes a: "a\<sharp>t1"
e61d2424863d initial cleanup to use the new induction method
urbanc
parents: 18269
diff changeset
    28
  and     b: "a\<sharp>t2"
22540
e4817fa0f6a1 adapted to new nominal_inductive
urbanc
parents: 22538
diff changeset
    29
  shows "a\<sharp>t1[b::=t2]"
18313
e61d2424863d initial cleanup to use the new induction method
urbanc
parents: 18269
diff changeset
    30
using a b
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
    31
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
    32
   (auto simp add: abs_fresh fresh_atm)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    33
22540
e4817fa0f6a1 adapted to new nominal_inductive
urbanc
parents: 22538
diff changeset
    34
lemma fresh_fact': 
e4817fa0f6a1 adapted to new nominal_inductive
urbanc
parents: 22538
diff changeset
    35
  fixes a::"name"
e4817fa0f6a1 adapted to new nominal_inductive
urbanc
parents: 22538
diff changeset
    36
  assumes a: "a\<sharp>t2"
e4817fa0f6a1 adapted to new nominal_inductive
urbanc
parents: 22538
diff changeset
    37
  shows "a\<sharp>t1[a::=t2]"
e4817fa0f6a1 adapted to new nominal_inductive
urbanc
parents: 22538
diff changeset
    38
using a 
e4817fa0f6a1 adapted to new nominal_inductive
urbanc
parents: 22538
diff changeset
    39
by (nominal_induct t1 avoiding: a t2 rule: lam.induct)
e4817fa0f6a1 adapted to new nominal_inductive
urbanc
parents: 22538
diff changeset
    40
   (auto simp add: abs_fresh fresh_atm)
e4817fa0f6a1 adapted to new nominal_inductive
urbanc
parents: 22538
diff changeset
    41
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
    42
lemma subst_lemma:  
18313
e61d2424863d initial cleanup to use the new induction method
urbanc
parents: 18269
diff changeset
    43
  assumes a: "x\<noteq>y"
e61d2424863d initial cleanup to use the new induction method
urbanc
parents: 18269
diff changeset
    44
  and     b: "x\<sharp>L"
e61d2424863d initial cleanup to use the new induction method
urbanc
parents: 18269
diff changeset
    45
  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
    46
using a b
18659
2ff0ae57431d changes to make use of the new induction principle proved by
urbanc
parents: 18654
diff changeset
    47
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
    48
   (auto simp add: fresh_fact forget)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    49
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
    50
lemma id_subs: 
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
    51
  shows "t[x::=Var x] = t"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
    52
  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
    53
     (simp_all add: fresh_atm)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
    54
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
    55
lemma subst_eqvt[eqvt]:
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
    56
  fixes pi::"name prm" 
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
    57
  and   t::"lam"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
    58
  shows "pi\<bullet>(t[x::=t']) = (pi\<bullet>t)[(pi\<bullet>x)::=(pi\<bullet>t')]"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
    59
by (nominal_induct t avoiding: x t' rule: lam.induct)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
    60
   (perm_simp add: fresh_bij)+
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    61
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    62
inductive2 Beta :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^isub>\<beta> _" [80,80] 80)
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    63
where
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    64
  b1[intro!]: "s1\<longrightarrow>\<^isub>\<beta>s2 \<Longrightarrow> (App s1 t)\<longrightarrow>\<^isub>\<beta>(App s2 t)"
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    65
| b2[intro!]: "s1\<longrightarrow>\<^isub>\<beta>s2 \<Longrightarrow> (App t s1)\<longrightarrow>\<^isub>\<beta>(App t s2)"
22538
3ccb92dfb5e9 tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents: 22531
diff changeset
    66
| b3[intro!]: "s1\<longrightarrow>\<^isub>\<beta>s2 \<Longrightarrow> (Lam [a].s1)\<longrightarrow>\<^isub>\<beta> (Lam [a].s2)"
22540
e4817fa0f6a1 adapted to new nominal_inductive
urbanc
parents: 22538
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
3ccb92dfb5e9 tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents: 22531
diff changeset
    69
nominal_inductive Beta
22540
e4817fa0f6a1 adapted to new nominal_inductive
urbanc
parents: 22538
diff changeset
    70
  by (simp_all add: abs_fresh fresh_fact')
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    71
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    72
abbreviation "Beta_star"  :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^isub>\<beta>\<^sup>* _" [80,80] 80) where
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    73
  "t1 \<longrightarrow>\<^isub>\<beta>\<^sup>* t2 \<equiv> Beta\<^sup>*\<^sup>* t1 t2"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
    74
18378
urbanc
parents: 18348
diff changeset
    75
lemma supp_beta: 
urbanc
parents: 18348
diff changeset
    76
  assumes a: "t\<longrightarrow>\<^isub>\<beta> s"
urbanc
parents: 18348
diff changeset
    77
  shows "(supp s)\<subseteq>((supp t)::name set)"
urbanc
parents: 18348
diff changeset
    78
using a
urbanc
parents: 18348
diff changeset
    79
by (induct)
urbanc
parents: 18348
diff changeset
    80
   (auto intro!: simp add: abs_supp lam.supp subst_supp)
urbanc
parents: 18348
diff changeset
    81
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    82
lemma beta_abs: "Lam [a].t\<longrightarrow>\<^isub>\<beta> t'\<Longrightarrow>\<exists>t''. t'=Lam [a].t'' \<and> t\<longrightarrow>\<^isub>\<beta> t''"
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    83
apply(ind_cases2 "Lam [a].t  \<longrightarrow>\<^isub>\<beta> t'")
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    84
apply(auto simp add: lam.distinct lam.inject)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    85
apply(auto simp add: alpha)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    86
apply(rule_tac x="[(a,aa)]\<bullet>s2" in exI)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    87
apply(rule conjI)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    88
apply(rule sym)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    89
apply(rule pt_bij2[OF pt_name_inst, OF at_name_inst])
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    90
apply(simp)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    91
apply(rule pt_name3)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    92
apply(simp add: at_ds5[OF at_name_inst])
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    93
apply(rule conjI)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    94
apply(simp add: pt_fresh_left[OF pt_name_inst, OF at_name_inst] calc_atm)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    95
apply(force dest!: supp_beta simp add: fresh_def)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
    96
apply(force intro!: eqvt)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    97
done
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
    98
18313
e61d2424863d initial cleanup to use the new induction method
urbanc
parents: 18269
diff changeset
    99
lemma beta_subst: 
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   100
  assumes a: "M \<longrightarrow>\<^isub>\<beta> M'"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   101
  shows "M[x::=N]\<longrightarrow>\<^isub>\<beta> M'[x::=N]" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   102
using a
22538
3ccb92dfb5e9 tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents: 22531
diff changeset
   103
apply(nominal_induct M M' avoiding: x N rule: Beta.strong_induct)
3ccb92dfb5e9 tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents: 22531
diff changeset
   104
apply(auto simp add: fresh_atm subst_lemma fresh_fact)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   105
done
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   106
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   107
section {* types *}
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   108
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
   109
nominal_datatype ty =
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
   110
    TVar "nat"
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   111
  | TArr "ty" "ty" (infix "\<rightarrow>" 200)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   112
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   113
lemma perm_ty[simp]:
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   114
  fixes pi ::"name prm"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   115
  and   \<tau>  ::"ty"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   116
  shows "pi\<bullet>\<tau> = \<tau>"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
   117
by (nominal_induct \<tau> rule: ty.induct) 
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
   118
   (simp_all add: perm_nat_def)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   119
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   120
lemma fresh_ty:
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   121
  fixes a ::"name"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   122
  and   \<tau>  ::"ty"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   123
  shows "a\<sharp>\<tau>"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   124
  by (simp add: fresh_def supp_def)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   125
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   126
(* valid contexts *)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   127
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   128
consts
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   129
  "dom_ty" :: "(name\<times>ty) list \<Rightarrow> (name list)"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   130
primrec
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   131
  "dom_ty []    = []"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   132
  "dom_ty (x#\<Gamma>) = (fst x)#(dom_ty \<Gamma>)" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   133
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   134
inductive2 valid :: "(name\<times>ty) list \<Rightarrow> bool"
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   135
where
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   136
  v1[intro]: "valid []"
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   137
| v2[intro]: "\<lbrakk>valid \<Gamma>;a\<sharp>\<Gamma>\<rbrakk>\<Longrightarrow> valid ((a,\<sigma>)#\<Gamma>)"
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   138
22538
3ccb92dfb5e9 tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents: 22531
diff changeset
   139
equivariance valid 
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   140
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   141
(* typing judgements *)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   142
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
   143
lemma fresh_context: 
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   144
  fixes  \<Gamma> :: "(name\<times>ty)list"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   145
  and    a :: "name"
18378
urbanc
parents: 18348
diff changeset
   146
  assumes a: "a\<sharp>\<Gamma>"
urbanc
parents: 18348
diff changeset
   147
  shows "\<not>(\<exists>\<tau>::ty. (a,\<tau>)\<in>set \<Gamma>)"
urbanc
parents: 18348
diff changeset
   148
using a
urbanc
parents: 18348
diff changeset
   149
apply(induct \<Gamma>)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   150
apply(auto simp add: fresh_prod fresh_list_cons fresh_atm)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   151
done
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   152
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
   153
inductive_cases2 valid_elim[elim]: "valid ((a,\<tau>)#\<Gamma>)"
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   154
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
   155
lemma valid_unicity: 
18378
urbanc
parents: 18348
diff changeset
   156
  assumes a: "valid \<Gamma>"
urbanc
parents: 18348
diff changeset
   157
  and     b: "(c,\<sigma>)\<in>set \<Gamma>"
urbanc
parents: 18348
diff changeset
   158
  and     c: "(c,\<tau>)\<in>set \<Gamma>"
urbanc
parents: 18348
diff changeset
   159
  shows "\<sigma>=\<tau>" 
urbanc
parents: 18348
diff changeset
   160
using a b c
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
   161
by (induct \<Gamma>) (auto dest!: fresh_context)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   162
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   163
inductive2 typing :: "(name\<times>ty) list\<Rightarrow>lam\<Rightarrow>ty\<Rightarrow>bool" (" _ \<turnstile> _ : _ " [80,80,80] 80)
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   164
where
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   165
  t1[intro]: "\<lbrakk>valid \<Gamma>; (a,\<tau>)\<in>set \<Gamma>\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> Var a : \<tau>"
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   166
| t2[intro]: "\<lbrakk>\<Gamma> \<turnstile> t1 : \<tau>\<rightarrow>\<sigma>; \<Gamma> \<turnstile> t2 : \<tau>\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> App t1 t2 : \<sigma>"
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   167
| 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
   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
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21377
diff changeset
   173
  "sub" :: "(name\<times>ty) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" (" _ \<lless> _ " [80,80] 80) where
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   174
  "\<Gamma>1 \<lless> \<Gamma>2 \<equiv> \<forall>a \<sigma>. (a,\<sigma>)\<in>set \<Gamma>1 \<longrightarrow>  (a,\<sigma>)\<in>set \<Gamma>2"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   175
18313
e61d2424863d initial cleanup to use the new induction method
urbanc
parents: 18269
diff changeset
   176
lemma weakening: 
e61d2424863d initial cleanup to use the new induction method
urbanc
parents: 18269
diff changeset
   177
  assumes a: "\<Gamma>1 \<turnstile> t : \<sigma>" 
e61d2424863d initial cleanup to use the new induction method
urbanc
parents: 18269
diff changeset
   178
  and     b: "valid \<Gamma>2" 
e61d2424863d initial cleanup to use the new induction method
urbanc
parents: 18269
diff changeset
   179
  and     c: "\<Gamma>1 \<lless> \<Gamma>2"
e61d2424863d initial cleanup to use the new induction method
urbanc
parents: 18269
diff changeset
   180
  shows "\<Gamma>2 \<turnstile> t:\<sigma>"
e61d2424863d initial cleanup to use the new induction method
urbanc
parents: 18269
diff changeset
   181
using a b c
22538
3ccb92dfb5e9 tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents: 22531
diff changeset
   182
apply(nominal_induct \<Gamma>1 t \<sigma> avoiding: \<Gamma>2 rule: typing.strong_induct)
21107
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 19972
diff changeset
   183
apply(auto | atomize)+
18313
e61d2424863d initial cleanup to use the new induction method
urbanc
parents: 18269
diff changeset
   184
(* FIXME: before using meta-connectives and the new induction *)
e61d2424863d initial cleanup to use the new induction method
urbanc
parents: 18269
diff changeset
   185
(* method, this was completely automatic *)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   186
done
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   187
18378
urbanc
parents: 18348
diff changeset
   188
lemma in_ctxt: 
urbanc
parents: 18348
diff changeset
   189
  assumes a: "(a,\<tau>)\<in>set \<Gamma>"
urbanc
parents: 18348
diff changeset
   190
  shows "a\<in>set(dom_ty \<Gamma>)"
urbanc
parents: 18348
diff changeset
   191
using a
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
   192
by (induct \<Gamma>) (auto)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   193
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   194
lemma free_vars: 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   195
  assumes a: "\<Gamma> \<turnstile> t : \<tau>"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   196
  shows " (supp t)\<subseteq>set(dom_ty \<Gamma>)"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   197
using a
22538
3ccb92dfb5e9 tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents: 22531
diff changeset
   198
by (nominal_induct \<Gamma> t \<tau> rule: typing.strong_induct)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
   199
   (auto simp add: lam.supp abs_supp supp_atm in_ctxt)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   200
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   201
lemma t1_elim: "\<Gamma> \<turnstile> Var a : \<tau> \<Longrightarrow> valid \<Gamma> \<and> (a,\<tau>) \<in> set \<Gamma>"
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   202
apply(ind_cases2 "\<Gamma> \<turnstile> Var a : \<tau>")
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   203
apply(auto simp add: lam.inject lam.distinct)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   204
done
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   205
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   206
lemma t2_elim: "\<Gamma> \<turnstile> App t1 t2 : \<sigma> \<Longrightarrow> \<exists>\<tau>. (\<Gamma> \<turnstile> t1 : \<tau>\<rightarrow>\<sigma> \<and> \<Gamma> \<turnstile> t2 : \<tau>)"
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   207
apply(ind_cases2 "\<Gamma> \<turnstile> App t1 t2 : \<sigma>")
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   208
apply(auto simp add: lam.inject lam.distinct)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   209
done
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   210
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   211
lemma t3_elim: "\<lbrakk>\<Gamma> \<turnstile> Lam [a].t : \<sigma>;a\<sharp>\<Gamma>\<rbrakk>\<Longrightarrow> \<exists>\<tau> \<tau>'. \<sigma>=\<tau>\<rightarrow>\<tau>' \<and> ((a,\<tau>)#\<Gamma>) \<turnstile> t : \<tau>'"
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   212
apply(ind_cases2 "\<Gamma> \<turnstile> Lam [a].t : \<sigma>")
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   213
apply(auto simp add: lam.distinct lam.inject alpha) 
22538
3ccb92dfb5e9 tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents: 22531
diff changeset
   214
apply(drule_tac pi="[(a,aa)]::name prm" in typing_eqvt)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   215
apply(simp)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   216
apply(subgoal_tac "([(a,aa)]::name prm)\<bullet>\<Gamma> = \<Gamma>")(*A*)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   217
apply(force simp add: calc_atm)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   218
(*A*)
22538
3ccb92dfb5e9 tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents: 22531
diff changeset
   219
apply(force intro!: perm_fresh_fresh)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   220
done
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   221
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   222
lemma typing_valid: 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   223
  assumes a: "\<Gamma> \<turnstile> t : \<tau>" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   224
  shows "valid \<Gamma>"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
   225
using a by (induct, auto)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   226
18378
urbanc
parents: 18348
diff changeset
   227
lemma ty_subs:
urbanc
parents: 18348
diff changeset
   228
  assumes a: "((c,\<sigma>)#\<Gamma>) \<turnstile> t1:\<tau>"
urbanc
parents: 18348
diff changeset
   229
  and     b: "\<Gamma>\<turnstile> t2:\<sigma>"
urbanc
parents: 18348
diff changeset
   230
  shows  "\<Gamma> \<turnstile> t1[c::=t2]:\<tau>"
urbanc
parents: 18348
diff changeset
   231
using a b
18659
2ff0ae57431d changes to make use of the new induction principle proved by
urbanc
parents: 18654
diff changeset
   232
proof(nominal_induct t1 avoiding: \<Gamma> \<sigma> \<tau> c t2 rule: lam.induct)
18313
e61d2424863d initial cleanup to use the new induction method
urbanc
parents: 18269
diff changeset
   233
  case (Var a) 
18378
urbanc
parents: 18348
diff changeset
   234
  have a1: "\<Gamma> \<turnstile>t2:\<sigma>" by fact
urbanc
parents: 18348
diff changeset
   235
  have a2: "((c,\<sigma>)#\<Gamma>) \<turnstile> Var a:\<tau>" by fact
urbanc
parents: 18348
diff changeset
   236
  hence a21: "(a,\<tau>)\<in>set((c,\<sigma>)#\<Gamma>)" and a22: "valid((c,\<sigma>)#\<Gamma>)" by (auto dest: t1_elim)
urbanc
parents: 18348
diff changeset
   237
  from a22 have a23: "valid \<Gamma>" and a24: "c\<sharp>\<Gamma>" by (auto dest: valid_elim) 
urbanc
parents: 18348
diff changeset
   238
  from a24 have a25: "\<not>(\<exists>\<tau>. (c,\<tau>)\<in>set \<Gamma>)" by (rule fresh_context)
urbanc
parents: 18348
diff changeset
   239
  show "\<Gamma>\<turnstile>(Var a)[c::=t2] : \<tau>"
urbanc
parents: 18348
diff changeset
   240
  proof (cases "a=c", simp_all)
urbanc
parents: 18348
diff changeset
   241
    assume case1: "a=c"
urbanc
parents: 18348
diff changeset
   242
    show "\<Gamma> \<turnstile> t2:\<tau>" using a1
urbanc
parents: 18348
diff changeset
   243
    proof (cases "\<sigma>=\<tau>")
urbanc
parents: 18348
diff changeset
   244
      assume "\<sigma>=\<tau>" thus ?thesis using a1 by simp 
urbanc
parents: 18348
diff changeset
   245
    next
urbanc
parents: 18348
diff changeset
   246
      assume a3: "\<sigma>\<noteq>\<tau>"
urbanc
parents: 18348
diff changeset
   247
      show ?thesis
urbanc
parents: 18348
diff changeset
   248
      proof (rule ccontr)
urbanc
parents: 18348
diff changeset
   249
	from a3 a21 have "(a,\<tau>)\<in>set \<Gamma>" by force
urbanc
parents: 18348
diff changeset
   250
	with case1 a25 show False by force 
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   251
      qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   252
    qed
18378
urbanc
parents: 18348
diff changeset
   253
  next
urbanc
parents: 18348
diff changeset
   254
    assume case2: "a\<noteq>c"
urbanc
parents: 18348
diff changeset
   255
    with a21 have a26: "(a,\<tau>)\<in>set \<Gamma>" by force 
urbanc
parents: 18348
diff changeset
   256
    from a23 a26 show "\<Gamma> \<turnstile> Var a:\<tau>" by force
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   257
  qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   258
next
18313
e61d2424863d initial cleanup to use the new induction method
urbanc
parents: 18269
diff changeset
   259
  case (App s1 s2)
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   260
  have ih_s1: "\<And>c \<sigma> \<tau> t2 \<Gamma>. ((c,\<sigma>)#\<Gamma>) \<turnstile> s1:\<tau> \<Longrightarrow> \<Gamma>\<turnstile> t2: \<sigma> \<Longrightarrow> \<Gamma> \<turnstile> s1[c::=t2]:\<tau>" by fact
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   261
  have ih_s2: "\<And>c \<sigma> \<tau> t2 \<Gamma>. ((c,\<sigma>)#\<Gamma>) \<turnstile> s2:\<tau> \<Longrightarrow> \<Gamma>\<turnstile> t2: \<sigma> \<Longrightarrow> \<Gamma> \<turnstile> s2[c::=t2]:\<tau>" by fact
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   262
  have "((c,\<sigma>)#\<Gamma>)\<turnstile>App s1 s2 : \<tau>" by fact
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   263
  hence "\<exists>\<tau>'. ((c,\<sigma>)#\<Gamma>)\<turnstile>s1:\<tau>'\<rightarrow>\<tau> \<and> ((c,\<sigma>)#\<Gamma>)\<turnstile>s2:\<tau>'" by (rule t2_elim) 
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   264
  then obtain \<tau>' where "((c,\<sigma>)#\<Gamma>)\<turnstile>s1:\<tau>'\<rightarrow>\<tau>" and "((c,\<sigma>)#\<Gamma>)\<turnstile>s2:\<tau>'" by blast
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   265
  moreover
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   266
  have "\<Gamma> \<turnstile>t2:\<sigma>" by fact
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   267
  ultimately show "\<Gamma> \<turnstile>  (App s1 s2)[c::=t2] : \<tau>" using ih_s1 ih_s2 by (simp, blast)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   268
next
18313
e61d2424863d initial cleanup to use the new induction method
urbanc
parents: 18269
diff changeset
   269
  case (Lam a s)
e61d2424863d initial cleanup to use the new induction method
urbanc
parents: 18269
diff changeset
   270
  have "a\<sharp>\<Gamma>" "a\<sharp>\<sigma>" "a\<sharp>\<tau>" "a\<sharp>c" "a\<sharp>t2" by fact 
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   271
  hence f1: "a\<sharp>\<Gamma>" and f2: "a\<noteq>c" and f2': "c\<sharp>a" and f3: "a\<sharp>t2" and f4: "a\<sharp>((c,\<sigma>)#\<Gamma>)"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   272
    by (auto simp add: fresh_atm fresh_prod fresh_list_cons)
18378
urbanc
parents: 18348
diff changeset
   273
  have c1: "((c,\<sigma>)#\<Gamma>)\<turnstile>Lam [a].s : \<tau>" by fact
urbanc
parents: 18348
diff changeset
   274
  hence "\<exists>\<tau>1 \<tau>2. \<tau>=\<tau>1\<rightarrow>\<tau>2 \<and> ((a,\<tau>1)#(c,\<sigma>)#\<Gamma>) \<turnstile> s : \<tau>2" using f4 by (auto dest: t3_elim) 
urbanc
parents: 18348
diff changeset
   275
  then obtain \<tau>1 \<tau>2 where c11: "\<tau>=\<tau>1\<rightarrow>\<tau>2" and c12: "((a,\<tau>1)#(c,\<sigma>)#\<Gamma>) \<turnstile> s : \<tau>2" by force
urbanc
parents: 18348
diff changeset
   276
  from c12 have "valid ((a,\<tau>1)#(c,\<sigma>)#\<Gamma>)" by (rule typing_valid)
urbanc
parents: 18348
diff changeset
   277
  hence ca: "valid \<Gamma>" and cb: "a\<sharp>\<Gamma>" and cc: "c\<sharp>\<Gamma>" 
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
   278
    by (auto simp add: fresh_list_cons) 
18378
urbanc
parents: 18348
diff changeset
   279
  from c12 have c14: "((c,\<sigma>)#(a,\<tau>1)#\<Gamma>) \<turnstile> s : \<tau>2"
urbanc
parents: 18348
diff changeset
   280
  proof -
21107
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 19972
diff changeset
   281
    have c2: "((a,\<tau>1)#(c,\<sigma>)#\<Gamma>) \<lless> ((c,\<sigma>)#(a,\<tau>1)#\<Gamma>)" by force
18378
urbanc
parents: 18348
diff changeset
   282
    have c3: "valid ((c,\<sigma>)#(a,\<tau>1)#\<Gamma>)"
urbanc
parents: 18348
diff changeset
   283
      by (rule v2, rule v2, auto simp add: fresh_list_cons fresh_prod ca cb cc f2' fresh_ty)
urbanc
parents: 18348
diff changeset
   284
    from c12 c2 c3 show ?thesis by (force intro: weakening)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   285
  qed
18378
urbanc
parents: 18348
diff changeset
   286
  assume c8: "\<Gamma> \<turnstile> t2 : \<sigma>"
urbanc
parents: 18348
diff changeset
   287
  have c81: "((a,\<tau>1)#\<Gamma>)\<turnstile>t2 :\<sigma>"
urbanc
parents: 18348
diff changeset
   288
  proof -
21107
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents: 19972
diff changeset
   289
    have c82: "\<Gamma> \<lless> ((a,\<tau>1)#\<Gamma>)" by force
18378
urbanc
parents: 18348
diff changeset
   290
    have c83: "valid ((a,\<tau>1)#\<Gamma>)" using f1 ca by force
urbanc
parents: 18348
diff changeset
   291
    with c8 c82 c83 show ?thesis by (force intro: weakening)
urbanc
parents: 18348
diff changeset
   292
  qed
urbanc
parents: 18348
diff changeset
   293
  show "\<Gamma> \<turnstile> (Lam [a].s)[c::=t2] : \<tau>"
urbanc
parents: 18348
diff changeset
   294
    using c11 prems c14 c81 f1 by force
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   295
qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   296
18378
urbanc
parents: 18348
diff changeset
   297
lemma subject: 
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   298
  assumes a: "t1\<longrightarrow>\<^isub>\<beta>t2"
18378
urbanc
parents: 18348
diff changeset
   299
  and     b: "\<Gamma> \<turnstile> t1:\<tau>"
urbanc
parents: 18348
diff changeset
   300
  shows "\<Gamma> \<turnstile> t2:\<tau>"
urbanc
parents: 18348
diff changeset
   301
using a b
22538
3ccb92dfb5e9 tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents: 22531
diff changeset
   302
proof (nominal_induct t1 t2 avoiding: \<Gamma> \<tau> rule: Beta.strong_induct)
3ccb92dfb5e9 tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents: 22531
diff changeset
   303
  case (b1 s1 s2 t) --"App-case left"
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   304
  have ih: "\<And>\<Gamma> \<tau>. \<Gamma> \<turnstile> s1:\<tau> \<Longrightarrow> \<Gamma> \<turnstile> s2 : \<tau>" by fact
18378
urbanc
parents: 18348
diff changeset
   305
  have "\<Gamma> \<turnstile> App s1 t : \<tau>" by fact 
urbanc
parents: 18348
diff changeset
   306
  hence "\<exists>\<sigma>. \<Gamma> \<turnstile> s1 : \<sigma>\<rightarrow>\<tau> \<and> \<Gamma> \<turnstile> t : \<sigma>" by (rule t2_elim)
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   307
  then obtain \<sigma> where "\<Gamma> \<turnstile> s1 : \<sigma>\<rightarrow>\<tau>" and "\<Gamma> \<turnstile> t : \<sigma>" by blast
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   308
  with ih show "\<Gamma> \<turnstile> App s2 t : \<tau>" by blast
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   309
next
22538
3ccb92dfb5e9 tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents: 22531
diff changeset
   310
  case (b2 s1 s2 t) --"App-case right"
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   311
  have ih: "\<And>\<Gamma> \<tau>. \<Gamma> \<turnstile> s1 : \<tau> \<Longrightarrow> \<Gamma> \<turnstile> s2 : \<tau>" by fact 
18378
urbanc
parents: 18348
diff changeset
   312
  have "\<Gamma> \<turnstile> App t s1 : \<tau>" by fact
urbanc
parents: 18348
diff changeset
   313
  hence "\<exists>\<sigma>. \<Gamma> \<turnstile> t : \<sigma>\<rightarrow>\<tau> \<and> \<Gamma> \<turnstile> s1 : \<sigma>" by (rule t2_elim)
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   314
  then obtain \<sigma> where "\<Gamma> \<turnstile> t : \<sigma>\<rightarrow>\<tau>" and "\<Gamma> \<turnstile> s1 : \<sigma>" by blast
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   315
  with ih show "\<Gamma> \<turnstile> App t s2 : \<tau>" by blast
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   316
next
22538
3ccb92dfb5e9 tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents: 22531
diff changeset
   317
  case (b3 s1 s2 a) --"Lam-case"
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   318
  have fr: "a\<sharp>\<Gamma>" "a\<sharp>\<tau>" by fact
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   319
  have ih: "\<And>\<Gamma> \<tau>. \<Gamma> \<turnstile> s1 : \<tau> \<Longrightarrow> \<Gamma> \<turnstile> s2 : \<tau>" by fact
18378
urbanc
parents: 18348
diff changeset
   320
  have "\<Gamma> \<turnstile> Lam [a].s1 : \<tau>" by fact
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   321
  with fr have "\<exists>\<tau>1 \<tau>2. \<tau>=\<tau>1\<rightarrow>\<tau>2 \<and> ((a,\<tau>1)#\<Gamma>) \<turnstile> s1 : \<tau>2" by (simp add: t3_elim)
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   322
  then obtain \<tau>1 \<tau>2 where "\<tau>=\<tau>1\<rightarrow>\<tau>2" and "((a,\<tau>1)#\<Gamma>) \<turnstile> s1 : \<tau>2" by blast
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   323
  with ih show "\<Gamma> \<turnstile> Lam [a].s2 : \<tau>" using fr by blast
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   324
next
22540
e4817fa0f6a1 adapted to new nominal_inductive
urbanc
parents: 22538
diff changeset
   325
  case (b4 a s2 s1) --"Beta-redex"
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   326
  have fr: "a\<sharp>\<Gamma>" by fact
18378
urbanc
parents: 18348
diff changeset
   327
  have "\<Gamma> \<turnstile> App (Lam [a].s1) s2 : \<tau>" by fact
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   328
  hence "\<exists>\<sigma>. (\<Gamma> \<turnstile> (Lam [a].s1) : \<sigma>\<rightarrow>\<tau> \<and> \<Gamma> \<turnstile> s2 : \<sigma>)" by (simp add: t2_elim)
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   329
  then obtain \<sigma> where a1: "\<Gamma> \<turnstile> (Lam [a].s1) : \<sigma>\<rightarrow>\<tau>" and a2: "\<Gamma> \<turnstile> s2 : \<sigma>" by blast
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
   330
  from a1 have "((a,\<sigma>)#\<Gamma>) \<turnstile> s1 : \<tau>" using fr by (auto dest!: t3_elim simp add: ty.inject) 
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   331
  with a2 show "\<Gamma> \<turnstile> s1[a::=s2] : \<tau>" by (simp add: ty_subs)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   332
qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   333
18378
urbanc
parents: 18348
diff changeset
   334
lemma subject_automatic: 
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   335
  assumes a: "t1\<longrightarrow>\<^isub>\<beta>t2"
18378
urbanc
parents: 18348
diff changeset
   336
  and     b: "\<Gamma> \<turnstile> t1:\<tau>"
urbanc
parents: 18348
diff changeset
   337
  shows "\<Gamma> \<turnstile> t2:\<tau>"
urbanc
parents: 18348
diff changeset
   338
using a b
22538
3ccb92dfb5e9 tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents: 22531
diff changeset
   339
apply(nominal_induct t1 t2 avoiding: \<Gamma> \<tau> rule: Beta.strong_induct)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
   340
apply(auto dest!: t2_elim t3_elim intro: ty_subs simp add: ty.inject)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   341
done
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   342
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   343
subsection {* some facts about beta *}
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   344
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   345
constdefs
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   346
  "NORMAL" :: "lam \<Rightarrow> bool"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   347
  "NORMAL t \<equiv> \<not>(\<exists>t'. t\<longrightarrow>\<^isub>\<beta> t')"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   348
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   349
lemma NORMAL_Var:
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   350
  shows "NORMAL (Var a)"
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   351
proof -
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   352
  { assume "\<exists>t'. (Var a) \<longrightarrow>\<^isub>\<beta> t'"
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   353
    then obtain t' where "(Var a) \<longrightarrow>\<^isub>\<beta> t'" by blast
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   354
    hence False by (cases, auto) 
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   355
  }
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   356
  thus "NORMAL (Var a)" by (force simp add: NORMAL_def)
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   357
qed
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   358
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   359
constdefs
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   360
  "SN" :: "lam \<Rightarrow> bool"
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   361
  "SN t \<equiv> termi Beta t"
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   362
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   363
lemma SN_preserved: "\<lbrakk>SN(t1);t1\<longrightarrow>\<^isub>\<beta> t2\<rbrakk>\<Longrightarrow>SN(t2)"
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   364
apply(simp add: SN_def)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   365
apply(drule_tac a="t2" in acc_downward)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   366
apply(auto)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   367
done
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   368
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   369
lemma SN_intro: "(\<forall>t2. t1\<longrightarrow>\<^isub>\<beta>t2 \<longrightarrow> SN(t2))\<Longrightarrow>SN(t1)"
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   370
apply(simp add: SN_def)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   371
apply(rule accI)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   372
apply(auto)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   373
done
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   374
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   375
section {* Candidates *}
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   376
22440
7e4f4f19002f deleted function for defining candidates and used nominal_primrec instead
urbanc
parents: 22420
diff changeset
   377
consts
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   378
  RED :: "ty \<Rightarrow> lam set"
22440
7e4f4f19002f deleted function for defining candidates and used nominal_primrec instead
urbanc
parents: 22420
diff changeset
   379
7e4f4f19002f deleted function for defining candidates and used nominal_primrec instead
urbanc
parents: 22420
diff changeset
   380
nominal_primrec
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
   381
  "RED (TVar X) = {t. SN(t)}"
22440
7e4f4f19002f deleted function for defining candidates and used nominal_primrec instead
urbanc
parents: 22420
diff changeset
   382
  "RED (\<tau>\<rightarrow>\<sigma>) =   {t. \<forall>u. (u\<in>RED \<tau> \<longrightarrow> (App t u)\<in>RED \<sigma>)}"
7e4f4f19002f deleted function for defining candidates and used nominal_primrec instead
urbanc
parents: 22420
diff changeset
   383
apply(rule TrueI)+
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
   384
done
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
   385
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   386
constdefs
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   387
  NEUT :: "lam \<Rightarrow> bool"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   388
  "NEUT t \<equiv> (\<exists>a. t=Var a)\<or>(\<exists>t1 t2. t=App t1 t2)" 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   389
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   390
(* a slight hack to get the first element of applications *)
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   391
inductive2 FST :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<guillemotright> _" [80,80] 80)
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   392
where
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
   393
  fst[intro!]:  "(App t s) \<guillemotright> t"
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   394
18378
urbanc
parents: 18348
diff changeset
   395
lemma fst_elim[elim!]: 
urbanc
parents: 18348
diff changeset
   396
  shows "(App t s) \<guillemotright> t' \<Longrightarrow> t=t'"
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   397
apply(ind_cases2 "App t s \<guillemotright> t'")
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   398
apply(simp add: lam.inject)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   399
done
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   400
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   401
lemma qq3: "SN(App t s)\<Longrightarrow>SN(t)"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   402
apply(simp add: SN_def)
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   403
apply(subgoal_tac "\<forall>z. (App t s \<guillemotright> z) \<longrightarrow> termi Beta z")(*A*)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   404
apply(force)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   405
(*A*)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   406
apply(erule acc_induct)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   407
apply(clarify)
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   408
apply(ind_cases2 "x \<guillemotright> z" for x z)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   409
apply(clarify)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   410
apply(rule accI)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   411
apply(auto intro: b1)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   412
done
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   413
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   414
section {* Candidates *}
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   415
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   416
constdefs
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   417
  "CR1" :: "ty \<Rightarrow> bool"
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   418
  "CR1 \<tau> \<equiv> \<forall> t. (t\<in>RED \<tau> \<longrightarrow> SN(t))"
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   419
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   420
  "CR2" :: "ty \<Rightarrow> bool"
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   421
  "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
   422
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   423
  "CR3_RED" :: "lam \<Rightarrow> ty \<Rightarrow> bool"
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   424
  "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
   425
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   426
  "CR3" :: "ty \<Rightarrow> bool"
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   427
  "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
   428
   
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   429
  "CR4" :: "ty \<Rightarrow> bool"
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   430
  "CR4 \<tau> \<equiv> \<forall>t. (NEUT t \<and> NORMAL t) \<longrightarrow>t\<in>RED \<tau>"
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   431
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   432
lemma CR3_CR4: "CR3 \<tau> \<Longrightarrow> CR4 \<tau>"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   433
apply(simp (no_asm_use) add: CR3_def CR3_RED_def CR4_def NORMAL_def)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   434
apply(blast)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   435
done
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   436
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   437
lemma sub_ind: 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   438
  "SN(u)\<Longrightarrow>(u\<in>RED \<tau>\<longrightarrow>(\<forall>t. (NEUT t\<and>CR2 \<tau>\<and>CR3 \<sigma>\<and>CR3_RED t (\<tau>\<rightarrow>\<sigma>))\<longrightarrow>(App t u)\<in>RED \<sigma>))"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   439
apply(simp add: SN_def)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   440
apply(erule acc_induct)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   441
apply(auto)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   442
apply(simp add: CR3_def)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   443
apply(rotate_tac 5)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   444
apply(drule_tac x="App t x" in spec)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   445
apply(drule mp)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   446
apply(rule conjI)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   447
apply(force simp only: NEUT_def)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   448
apply(simp (no_asm) add: CR3_RED_def)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   449
apply(clarify)
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   450
apply(ind_cases2 "App t x \<longrightarrow>\<^isub>\<beta> t'" for x t t')
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   451
apply(simp_all add: lam.inject)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   452
apply(simp only:  CR3_RED_def)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   453
apply(drule_tac x="s2" in spec)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   454
apply(simp)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   455
apply(drule_tac x="s2" in spec)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   456
apply(simp)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   457
apply(drule mp)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   458
apply(simp (no_asm_use) add: CR2_def)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   459
apply(blast)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   460
apply(drule_tac x="ta" in spec)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   461
apply(force)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   462
apply(auto simp only: NEUT_def lam.inject lam.distinct)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   463
done
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   464
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   465
lemma RED_props: 
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   466
  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
   467
proof (nominal_induct \<tau> rule: ty.induct)
18611
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   468
  case (TVar a)
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   469
  { case 1 show "CR1 (TVar a)" by (simp add: CR1_def)
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   470
  next
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   471
    case 2 show "CR2 (TVar a)" by (force intro: SN_preserved simp add: CR2_def)
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   472
  next
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   473
    case 3 show "CR3 (TVar a)" by (force intro: SN_intro simp add: CR3_def CR3_RED_def)
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   474
  }
18599
e01112713fdc changed PRO_RED proof to conform with the new induction rules
urbanc
parents: 18383
diff changeset
   475
next
18611
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   476
  case (TArr \<tau>1 \<tau>2)
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   477
  { case 1
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   478
    have ih_CR3_\<tau>1: "CR3 \<tau>1" by fact
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   479
    have ih_CR1_\<tau>2: "CR1 \<tau>2" by fact
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   480
    show "CR1 (\<tau>1 \<rightarrow> \<tau>2)"
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   481
    proof (simp add: CR1_def, intro strip)
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   482
      fix t
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   483
      assume a: "\<forall>u. u \<in> RED \<tau>1 \<longrightarrow> App t u \<in> RED \<tau>2"
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   484
      from ih_CR3_\<tau>1 have "CR4 \<tau>1" by (simp add: CR3_CR4) 
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   485
      moreover
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   486
      have "NEUT (Var a)" by (force simp add: NEUT_def)
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   487
      moreover
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   488
      have "NORMAL (Var a)" by (rule NORMAL_Var)
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   489
      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
   490
      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
   491
      hence "SN (App t (Var a))" using ih_CR1_\<tau>2 by (simp add: CR1_def)
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   492
      thus "SN(t)" by (rule qq3)
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   493
    qed
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   494
  next
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   495
    case 2
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   496
    have ih_CR1_\<tau>1: "CR1 \<tau>1" by fact
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   497
    have ih_CR2_\<tau>2: "CR2 \<tau>2" by fact
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   498
    show "CR2 (\<tau>1 \<rightarrow> \<tau>2)"
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   499
    proof (simp add: CR2_def, intro strip)
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   500
      fix t1 t2 u
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   501
      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
   502
	and  "u \<in> RED \<tau>1"
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   503
      hence "t1 \<longrightarrow>\<^isub>\<beta> t2" and "App t1 u \<in> RED \<tau>2" by simp_all
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   504
      thus "App t2 u \<in> RED \<tau>2" using ih_CR2_\<tau>2 by (force simp add: CR2_def)
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   505
    qed
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   506
  next
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   507
    case 3
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   508
    have ih_CR1_\<tau>1: "CR1 \<tau>1" by fact
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   509
    have ih_CR2_\<tau>1: "CR2 \<tau>1" by fact
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   510
    have ih_CR3_\<tau>2: "CR3 \<tau>2" by fact
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   511
    show "CR3 (\<tau>1 \<rightarrow> \<tau>2)"
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   512
    proof (simp add: CR3_def, intro strip)
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   513
      fix t u
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   514
      assume a1: "u \<in> RED \<tau>1"
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   515
      assume a2: "NEUT t \<and> CR3_RED t (\<tau>1 \<rightarrow> \<tau>2)"
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   516
      from a1 have "SN(u)" using ih_CR1_\<tau>1 by (simp add: CR1_def)
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   517
      hence "u\<in>RED \<tau>1\<longrightarrow>(\<forall>t. (NEUT t\<and>CR2 \<tau>1\<and>CR3 \<tau>2\<and>CR3_RED t (\<tau>1\<rightarrow>\<tau>2))\<longrightarrow>(App t u)\<in>RED \<tau>2)" 
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   518
	by (rule sub_ind)
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   519
      with a1 a2 show "(App t u)\<in>RED \<tau>2" using ih_CR2_\<tau>1 ih_CR3_\<tau>2 by simp
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   520
    qed
687c9bffbca1 another change for the new induct-method
urbanc
parents: 18599
diff changeset
   521
  }
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   522
qed
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   523
    
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   524
lemma double_acc_aux:
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   525
  assumes a_acc: "acc r a"
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   526
  and b_acc: "acc r b"
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   527
  and hyp: "\<And>x z.
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   528
    (\<And>y. r y x \<Longrightarrow> acc r y) \<Longrightarrow>
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   529
    (\<And>y. r y x \<Longrightarrow> P y z) \<Longrightarrow>
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   530
    (\<And>u. r u z \<Longrightarrow> acc r u) \<Longrightarrow>
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   531
    (\<And>u. r u z \<Longrightarrow> P x u) \<Longrightarrow> P x z"
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   532
  shows "P a b"
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   533
proof -
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   534
  from a_acc
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   535
  have r: "\<And>b. acc r b \<Longrightarrow> P a b"
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   536
  proof (induct a rule: acc.induct)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   537
    case (accI x)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   538
    note accI' = accI
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   539
    have "acc r b" .
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   540
    thus ?case
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   541
    proof (induct b rule: acc.induct)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   542
      case (accI y)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   543
      show ?case
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   544
	apply (rule hyp)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   545
	apply (erule accI')
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   546
	apply (erule accI')
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   547
	apply (rule acc.accI)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   548
	apply (erule accI)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   549
	apply (erule accI)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   550
	apply (erule accI)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   551
	done
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   552
    qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   553
  qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   554
  from b_acc show ?thesis by (rule r)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   555
qed
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   556
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   557
lemma double_acc:
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   558
  "\<lbrakk>acc r a; acc r b; \<forall>x z. ((\<forall>y. r y x \<longrightarrow> P y z) \<and> (\<forall>u. r u z \<longrightarrow> P x u)) \<longrightarrow> P x z\<rbrakk> \<Longrightarrow> P a b"
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   559
apply(rule_tac r="r" in double_acc_aux)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   560
apply(assumption)+
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   561
apply(blast)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   562
done
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   563
18263
7f75925498da cleaned up all examples so that they work with the
urbanc
parents: 18106
diff changeset
   564
lemma abs_RED: "(\<forall>s\<in>RED \<tau>. t[x::=s]\<in>RED \<sigma>)\<longrightarrow>Lam [x].t\<in>RED (\<tau>\<rightarrow>\<sigma>)"
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   565
apply(simp)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   566
apply(clarify)
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   567
apply(subgoal_tac "termi Beta t")(*1*)
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   568
apply(erule rev_mp)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   569
apply(subgoal_tac "u \<in> RED \<tau>")(*A*)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   570
apply(erule rev_mp)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   571
apply(rule_tac a="t" and b="u" in double_acc)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   572
apply(assumption)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   573
apply(subgoal_tac "CR1 \<tau>")(*A*)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   574
apply(simp add: CR1_def SN_def)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   575
(*A*)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   576
apply(force simp add: RED_props)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   577
apply(simp)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   578
apply(clarify)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   579
apply(subgoal_tac "CR3 \<sigma>")(*B*)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   580
apply(simp add: CR3_def)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   581
apply(rotate_tac 6)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   582
apply(drule_tac x="App(Lam[x].xa ) z" in spec)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   583
apply(drule mp)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   584
apply(rule conjI)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   585
apply(force simp add: NEUT_def)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   586
apply(simp add: CR3_RED_def)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   587
apply(clarify)
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   588
apply(ind_cases2 "App(Lam[x].xa) z \<longrightarrow>\<^isub>\<beta> t'" for xa z t')
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   589
apply(auto simp add: lam.inject lam.distinct)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   590
apply(drule beta_abs)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   591
apply(auto)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   592
apply(drule_tac x="t''" in spec)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   593
apply(simp)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   594
apply(drule mp)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   595
apply(clarify)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   596
apply(drule_tac x="s" in bspec)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   597
apply(assumption)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   598
apply(subgoal_tac "xa [ x ::= s ] \<longrightarrow>\<^isub>\<beta>  t'' [ x ::= s ]")(*B*)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   599
apply(subgoal_tac "CR2 \<sigma>")(*C*)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   600
apply(simp (no_asm_use) add: CR2_def)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   601
apply(blast)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   602
(*C*)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   603
apply(force simp add: RED_props)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   604
(*B*)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   605
apply(force intro!: beta_subst)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   606
apply(assumption)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   607
apply(rotate_tac 3)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   608
apply(drule_tac x="s2" in spec)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   609
apply(subgoal_tac "s2\<in>RED \<tau>")(*D*)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   610
apply(simp)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   611
(*D*)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   612
apply(subgoal_tac "CR2 \<tau>")(*E*)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   613
apply(simp (no_asm_use) add: CR2_def)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   614
apply(blast)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   615
(*E*)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   616
apply(force simp add: RED_props)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   617
apply(simp add: alpha)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   618
apply(erule disjE)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   619
apply(force)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   620
apply(auto)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   621
apply(simp add: subst_rename)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   622
apply(drule_tac x="z" in bspec)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   623
apply(assumption)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   624
(*B*)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   625
apply(force simp add: RED_props)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   626
(*1*)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   627
apply(drule_tac x="Var x" in bspec)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   628
apply(subgoal_tac "CR3 \<tau>")(*2*) 
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   629
apply(drule CR3_CR4)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   630
apply(simp add: CR4_def)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   631
apply(drule_tac x="Var x" in spec)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   632
apply(drule mp)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   633
apply(rule conjI)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   634
apply(force simp add: NEUT_def)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   635
apply(simp add: NORMAL_def)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   636
apply(clarify)
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   637
apply(ind_cases2 "Var x \<longrightarrow>\<^isub>\<beta> t'" for t')
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   638
apply(auto simp add: lam.inject lam.distinct)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   639
apply(force simp add: RED_props)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   640
apply(simp add: id_subs)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   641
apply(subgoal_tac "CR1 \<sigma>")(*3*)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   642
apply(simp add: CR1_def SN_def)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   643
(*3*)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   644
apply(force simp add: RED_props)
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   645
done
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   646
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
   647
lemma psubst_subst:
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
   648
  assumes h:"c\<sharp>\<theta>"
22420
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   649
  shows "(\<theta><t>)[c::=s] = ((c,s)#\<theta>)<t>"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
   650
  using h
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
   651
by (nominal_induct t avoiding: \<theta> c s rule: lam.induct)
22420
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   652
   (auto simp add: fresh_list_cons fresh_atm forget lookup_fresh lookup_fresh')
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   653
 
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22271
diff changeset
   654
22420
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   655
abbreviation 
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   656
 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
   657
where
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   658
 "\<theta> maps x to e\<equiv> (lookup \<theta> x) = e"
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   659
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   660
abbreviation 
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   661
  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
   662
where
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   663
  "\<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
   664
18106
836135c8acb2 Initial commit.
urbanc
parents:
diff changeset
   665
lemma all_RED: 
22420
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   666
  assumes a: "\<Gamma> \<turnstile> t : \<tau>"
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   667
  and     b: "\<theta> closes \<Gamma>"
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   668
  shows "\<theta><t> \<in> RED \<tau>"
18345
d37fb71754fe added an Isar-proof for the abs_ALL lemma
urbanc
parents: 18313
diff changeset
   669
using a b
18659
2ff0ae57431d changes to make use of the new induction principle proved by
urbanc
parents: 18654
diff changeset
   670
proof(nominal_induct t avoiding: \<Gamma> \<tau> \<theta> rule: lam.induct)
18345
d37fb71754fe added an Isar-proof for the abs_ALL lemma
urbanc
parents: 18313
diff changeset
   671
  case (Lam a t) --"lambda case"
22420
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   672
  have ih: "\<And>\<Gamma> \<tau> \<theta>. \<Gamma> \<turnstile> t : \<tau> \<Longrightarrow> \<theta> closes \<Gamma> \<Longrightarrow> \<theta><t> \<in> RED \<tau>" 
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   673
  and  \<theta>_cond: "\<theta> closes \<Gamma>" 
18345
d37fb71754fe added an Isar-proof for the abs_ALL lemma
urbanc
parents: 18313
diff changeset
   674
  and fresh: "a\<sharp>\<Gamma>" "a\<sharp>\<theta>" 
d37fb71754fe added an Isar-proof for the abs_ALL lemma
urbanc
parents: 18313
diff changeset
   675
  and "\<Gamma> \<turnstile> Lam [a].t:\<tau>" by fact
d37fb71754fe added an Isar-proof for the abs_ALL lemma
urbanc
parents: 18313
diff changeset
   676
  hence "\<exists>\<tau>1 \<tau>2. \<tau>=\<tau>1\<rightarrow>\<tau>2 \<and> ((a,\<tau>1)#\<Gamma>)\<turnstile>t:\<tau>2" using t3_elim fresh by simp
d37fb71754fe added an Isar-proof for the abs_ALL lemma
urbanc
parents: 18313
diff changeset
   677
  then obtain \<tau>1 \<tau>2 where \<tau>_inst: "\<tau>=\<tau>1\<rightarrow>\<tau>2" and typing: "((a,\<tau>1)#\<Gamma>)\<turnstile>t:\<tau>2" by blast
22420
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   678
  from ih have "\<forall>s\<in>RED \<tau>1. \<theta><t>[a::=s] \<in> RED \<tau>2" using fresh typing \<theta>_cond
18345
d37fb71754fe added an Isar-proof for the abs_ALL lemma
urbanc
parents: 18313
diff changeset
   679
    by (force dest: fresh_context simp add: psubst_subst)
22420
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   680
  hence "(Lam [a].(\<theta><t>)) \<in> RED (\<tau>1 \<rightarrow> \<tau>2)" by (simp only: abs_RED)
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   681
  thus "\<theta><(Lam [a].t)> \<in> RED \<tau>" using fresh \<tau>_inst by simp
18345
d37fb71754fe added an Isar-proof for the abs_ALL lemma
urbanc
parents: 18313
diff changeset
   682
qed (force dest!: t1_elim t2_elim)+
d37fb71754fe added an Isar-proof for the abs_ALL lemma
urbanc
parents: 18313
diff changeset
   683
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   684
(* identity substitution generated from a context \<Gamma> *)
18382
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   685
consts
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   686
  "id" :: "(name\<times>ty) list \<Rightarrow> (name\<times>lam) list"
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   687
primrec
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   688
  "id []    = []"
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   689
  "id (x#\<Gamma>) = ((fst x),Var (fst x))#(id \<Gamma>)"
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   690
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   691
lemma id_var:
22420
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   692
  shows "(id \<Gamma>) maps a to Var a"
18382
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   693
apply(induct \<Gamma>, auto)
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   694
done
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   695
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   696
lemma id_fresh:
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   697
  fixes a::"name"
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   698
  assumes a: "a\<sharp>\<Gamma>"
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   699
  shows "a\<sharp>(id \<Gamma>)"
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   700
using a
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   701
apply(induct \<Gamma>)
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   702
apply(auto simp add: fresh_list_nil fresh_list_cons fresh_prod)
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   703
done
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   704
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   705
lemma id_apply:  
22420
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   706
  shows "(id \<Gamma>)<t> = t"
18659
2ff0ae57431d changes to make use of the new induction principle proved by
urbanc
parents: 18654
diff changeset
   707
apply(nominal_induct t avoiding: \<Gamma> rule: lam.induct)
18382
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   708
apply(auto)
22420
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   709
apply(rule id_var)
18382
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   710
apply(drule id_fresh)+
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   711
apply(simp)
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   712
done
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   713
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   714
lemma id_mem:
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   715
  assumes a: "(a,\<tau>)\<in>set \<Gamma>"
22420
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   716
  shows "lookup (id \<Gamma>) a = Var a"
18382
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   717
using a
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   718
apply(induct \<Gamma>, auto)
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   719
done
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   720
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   721
lemma id_prop:
22420
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   722
  shows "(id \<Gamma>) closes \<Gamma>"
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   723
apply(auto)
18382
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   724
apply(simp add: id_mem)
22420
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   725
apply(subgoal_tac "CR3 T") --"A"
18382
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   726
apply(drule CR3_CR4)
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   727
apply(simp add: CR4_def)
22420
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   728
apply(drule_tac x="Var x" in spec)
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   729
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
   730
--"A"
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   731
apply(rule RED_props)
18382
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   732
done
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   733
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   734
lemma typing_implies_RED:  
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   735
  assumes a: "\<Gamma>\<turnstile>t:\<tau>"
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   736
  shows "t \<in> RED \<tau>"
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   737
proof -
22420
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   738
  have "(id \<Gamma>)<t>\<in>RED \<tau>" 
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   739
  proof -
22420
4ccc8c1b08a3 updated this file to the new infrastructure
urbanc
parents: 22418
diff changeset
   740
    have "(id \<Gamma>) closes \<Gamma>" by (rule id_prop)
18383
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   741
    with a show ?thesis by (rule all_RED)
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   742
  qed
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   743
  thus"t \<in> RED \<tau>" by (simp add: id_apply)
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   744
qed
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   745
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   746
lemma typing_implies_SN: 
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   747
  assumes a: "\<Gamma>\<turnstile>t:\<tau>"
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   748
  shows "SN(t)"
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   749
proof -
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   750
  from a have "t \<in> RED \<tau>" by (rule typing_implies_RED)
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   751
  moreover
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   752
  have "CR1 \<tau>" by (rule RED_props)
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   753
  ultimately show "SN(t)" by (simp add: CR1_def)
5f40a59a798b ISAR-fied some proofs
urbanc
parents: 18382
diff changeset
   754
qed
18382
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   755
44578c918349 completed the sn proof and changed the manual
urbanc
parents: 18378
diff changeset
   756
end