src/HOL/Nominal/Examples/Fsub.thy
author wenzelm
Wed, 28 Nov 2012 15:59:18 +0100
changeset 50252 4aa34bd43228
parent 49171 3d7a695385f1
child 53015 a1119cf551e8
permissions -rw-r--r--
eliminated slightly odd identifiers;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
18577
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
     1
(*<*)
19501
9afa7183dfc2 Capitalized theory names.
berghofe
parents: 18747
diff changeset
     2
theory Fsub
9afa7183dfc2 Capitalized theory names.
berghofe
parents: 18747
diff changeset
     3
imports "../Nominal" 
18269
3f36e2165e51 some small tuning
urbanc
parents: 18266
diff changeset
     4
begin
18577
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
     5
(*>*)
18269
3f36e2165e51 some small tuning
urbanc
parents: 18266
diff changeset
     6
18577
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
     7
text{* Authors: Christian Urban,
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
     8
                Benjamin Pierce,
18650
urbanc
parents: 18628
diff changeset
     9
                Dimitrios Vytiniotis
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    10
                Stephanie Weirich
18650
urbanc
parents: 18628
diff changeset
    11
                Steve Zdancewic
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    12
                Julien Narboux
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    13
                Stefan Berghofer
18266
55c201fe4c95 added an authors section (please let me know if somebody is left out or unhappy)
urbanc
parents: 18263
diff changeset
    14
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    15
       with great help from Markus Wenzel. *}
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    16
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
    17
section {* Types for Names, Nominal Datatype Declaration for Types and Terms *}
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
    18
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    19
no_syntax
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    20
  "_Map" :: "maplets => 'a ~=> 'b"  ("(1[_])")
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    21
18577
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
    22
text {* The main point of this solution is to use names everywhere (be they bound, 
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
    23
  binding or free). In System \FSUB{} there are two kinds of names corresponding to 
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
    24
  type-variables and to term-variables. These two kinds of names are represented in 
18577
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
    25
  the nominal datatype package as atom-types @{text "tyvrs"} and @{text "vrs"}: *}
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
    26
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    27
atom_decl tyvrs vrs
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    28
18577
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
    29
text{* There are numerous facts that come with this declaration: for example that 
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
    30
  there are infinitely many elements in @{text "tyvrs"} and @{text "vrs"}. *}
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
    31
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
    32
text{* The constructors for types and terms in System \FSUB{} contain abstractions 
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
    33
  over type-variables and term-variables. The nominal datatype-package uses 
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
    34
  @{text "\<guillemotleft>\<dots>\<guillemotright>\<dots>"} to indicate where abstractions occur. *}
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
    35
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
    36
nominal_datatype ty = 
18577
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
    37
    Tvar   "tyvrs"
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
    38
  | Top
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    39
  | Arrow  "ty" "ty"         (infixr "\<rightarrow>" 200)
18577
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
    40
  | Forall "\<guillemotleft>tyvrs\<guillemotright>ty" "ty" 
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    41
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
    42
nominal_datatype trm = 
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
    43
    Var   "vrs"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    44
  | Abs   "\<guillemotleft>vrs\<guillemotright>trm" "ty" 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    45
  | TAbs  "\<guillemotleft>tyvrs\<guillemotright>trm" "ty"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    46
  | App   "trm" "trm" (infixl "\<cdot>" 200)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    47
  | TApp  "trm" "ty"  (infixl "\<cdot>\<^sub>\<tau>" 200)
18577
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
    48
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
    49
text {* To be polite to the eye, some more familiar notation is introduced. 
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
    50
  Because of the change in the order of arguments, one needs to use 
18577
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
    51
  translation rules, instead of syntax annotations at the term-constructors
18650
urbanc
parents: 18628
diff changeset
    52
  as given above for @{term "Arrow"}. *}
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    53
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    54
abbreviation
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    55
  Forall_syn :: "tyvrs \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> ty"  ("(3\<forall>_<:_./ _)" [0, 0, 10] 10) 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    56
where
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    57
  "\<forall>X<:T\<^isub>1. T\<^isub>2 \<equiv> ty.Forall X T\<^isub>2 T\<^isub>1"
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
    58
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    59
abbreviation
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    60
  Abs_syn    :: "vrs \<Rightarrow> ty \<Rightarrow> trm \<Rightarrow> trm"  ("(3\<lambda>_:_./ _)" [0, 0, 10] 10) 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    61
where
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    62
  "\<lambda>x:T. t \<equiv> trm.Abs x t T"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    63
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    64
abbreviation
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    65
  TAbs_syn   :: "tyvrs \<Rightarrow> ty \<Rightarrow> trm \<Rightarrow> trm" ("(3\<lambda>_<:_./ _)" [0, 0, 10] 10) 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    66
where
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    67
  "\<lambda>X<:T. t \<equiv> trm.TAbs X t T"
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    68
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
    69
text {* Again there are numerous facts that are proved automatically for @{typ "ty"} 
18650
urbanc
parents: 18628
diff changeset
    70
  and @{typ "trm"}: for example that the set of free variables, i.e.~the @{text "support"}, 
urbanc
parents: 18628
diff changeset
    71
  is finite. However note that nominal-datatype declarations do \emph{not} define 
urbanc
parents: 18628
diff changeset
    72
  ``classical" constructor-based datatypes, but rather define $\alpha$-equivalence 
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
    73
  classes---we can for example show that $\alpha$-equivalent @{typ "ty"}s 
18577
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
    74
  and @{typ "trm"}s are equal: *}
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
    75
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
    76
lemma alpha_illustration:
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    77
  shows "(\<forall>X<:T. Tvar X) = (\<forall>Y<:T. Tvar Y)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    78
  and   "(\<lambda>x:T. Var x) = (\<lambda>y:T. Var y)"
18577
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
    79
  by (simp_all add: ty.inject trm.inject alpha calc_atm fresh_atm)
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
    80
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
    81
section {* SubTyping Contexts *}
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    82
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    83
nominal_datatype binding = 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    84
    VarB vrs ty 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    85
  | TVarB tyvrs ty
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    86
41798
c3aa3c87ef21 modernized specifications;
wenzelm
parents: 39246
diff changeset
    87
type_synonym env = "binding list"
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    88
18650
urbanc
parents: 18628
diff changeset
    89
text {* Typing contexts are represented as lists that ``grow" on the left; we
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
    90
  thereby deviating from the convention in the POPLmark-paper. The lists contain
18650
urbanc
parents: 18628
diff changeset
    91
  pairs of type-variables and types (this is sufficient for Part 1A). *}
18577
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
    92
18628
urbanc
parents: 18621
diff changeset
    93
text {* In order to state validity-conditions for typing-contexts, the notion of
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
    94
  a @{text "dom"} of a typing-context is handy. *}
18577
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
    95
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    96
nominal_primrec
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    97
  "tyvrs_of" :: "binding \<Rightarrow> tyvrs set"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    98
where
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
    99
  "tyvrs_of (VarB  x y) = {}"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   100
| "tyvrs_of (TVarB x y) = {x}"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   101
by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   102
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   103
nominal_primrec
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   104
  "vrs_of" :: "binding \<Rightarrow> vrs set"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   105
where
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   106
  "vrs_of (VarB  x y) = {x}"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   107
| "vrs_of (TVarB x y) = {}"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   108
by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   109
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
   110
primrec
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   111
  "ty_dom" :: "env \<Rightarrow> tyvrs set"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
   112
where
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   113
  "ty_dom [] = {}"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
   114
| "ty_dom (X#\<Gamma>) = (tyvrs_of X)\<union>(ty_dom \<Gamma>)" 
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   115
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
   116
primrec
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   117
  "trm_dom" :: "env \<Rightarrow> vrs set"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
   118
where
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   119
  "trm_dom [] = {}"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
   120
| "trm_dom (X#\<Gamma>) = (vrs_of X)\<union>(trm_dom \<Gamma>)" 
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   121
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   122
lemma vrs_of_eqvt[eqvt]:
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   123
  fixes pi ::"tyvrs prm"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   124
  and   pi'::"vrs   prm"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   125
  shows "pi \<bullet>(tyvrs_of x) = tyvrs_of (pi\<bullet>x)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   126
  and   "pi'\<bullet>(tyvrs_of x) = tyvrs_of (pi'\<bullet>x)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   127
  and   "pi \<bullet>(vrs_of x)   = vrs_of   (pi\<bullet>x)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   128
  and   "pi'\<bullet>(vrs_of x)   = vrs_of   (pi'\<bullet>x)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   129
by (nominal_induct x rule: binding.strong_induct) (simp_all add: tyvrs_of.simps eqvts)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   130
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   131
lemma doms_eqvt[eqvt]:
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   132
  fixes pi::"tyvrs prm"
22537
c55f5631a4ec adapted to nominal_inductive infrastructure
urbanc
parents: 22436
diff changeset
   133
  and   pi'::"vrs prm"
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   134
  shows "pi \<bullet>(ty_dom \<Gamma>)  = ty_dom  (pi\<bullet>\<Gamma>)"
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   135
  and   "pi'\<bullet>(ty_dom \<Gamma>)  = ty_dom  (pi'\<bullet>\<Gamma>)"
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   136
  and   "pi \<bullet>(trm_dom \<Gamma>) = trm_dom (pi\<bullet>\<Gamma>)"
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   137
  and   "pi'\<bullet>(trm_dom \<Gamma>) = trm_dom (pi'\<bullet>\<Gamma>)"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   138
by (induct \<Gamma>) (simp_all add: eqvts)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   139
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   140
lemma finite_vrs:
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   141
  shows "finite (tyvrs_of x)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   142
  and   "finite (vrs_of x)"
49171
3d7a695385f1 tuned proofs;
wenzelm
parents: 46182
diff changeset
   143
by (nominal_induct rule:binding.strong_induct) auto
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   144
 
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   145
lemma finite_doms:
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   146
  shows "finite (ty_dom \<Gamma>)"
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   147
  and   "finite (trm_dom \<Gamma>)"
49171
3d7a695385f1 tuned proofs;
wenzelm
parents: 46182
diff changeset
   148
by (induct \<Gamma>) (auto simp add: finite_vrs)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   149
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   150
lemma ty_dom_supp:
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   151
  shows "(supp (ty_dom  \<Gamma>)) = (ty_dom  \<Gamma>)"
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   152
  and   "(supp (trm_dom \<Gamma>)) = (trm_dom \<Gamma>)"
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   153
by (simp only: at_fin_set_supp at_tyvrs_inst at_vrs_inst finite_doms)+
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   154
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   155
lemma ty_dom_inclusion:
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   156
  assumes a: "(TVarB X T)\<in>set \<Gamma>" 
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   157
  shows "X\<in>(ty_dom \<Gamma>)"
49171
3d7a695385f1 tuned proofs;
wenzelm
parents: 46182
diff changeset
   158
using a by (induct \<Gamma>) (auto)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   159
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   160
lemma ty_binding_existence:
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   161
  assumes "X \<in> (tyvrs_of a)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   162
  shows "\<exists>T.(TVarB X T=a)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   163
  using assms
49171
3d7a695385f1 tuned proofs;
wenzelm
parents: 46182
diff changeset
   164
by (nominal_induct a rule: binding.strong_induct) (auto)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   165
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   166
lemma ty_dom_existence:
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   167
  assumes a: "X\<in>(ty_dom \<Gamma>)" 
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   168
  shows "\<exists>T.(TVarB X T)\<in>set \<Gamma>"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   169
  using a 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   170
  apply (induct \<Gamma>, auto) 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   171
  apply (subgoal_tac "\<exists>T.(TVarB X T=a)")
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   172
  apply (auto)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   173
  apply (auto simp add: ty_binding_existence)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   174
done
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   175
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   176
lemma doms_append:
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   177
  shows "ty_dom (\<Gamma>@\<Delta>) = ((ty_dom \<Gamma>) \<union> (ty_dom \<Delta>))"
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   178
  and   "trm_dom (\<Gamma>@\<Delta>) = ((trm_dom \<Gamma>) \<union> (trm_dom \<Delta>))"
49171
3d7a695385f1 tuned proofs;
wenzelm
parents: 46182
diff changeset
   179
  by (induct \<Gamma>) (auto)
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   180
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   181
lemma ty_vrs_prm_simp:
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   182
  fixes pi::"vrs prm"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   183
  and   S::"ty"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   184
  shows "pi\<bullet>S = S"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   185
by (induct S rule: ty.induct) (auto simp add: calc_atm)
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   186
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   187
lemma fresh_ty_dom_cons:
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   188
  fixes X::"tyvrs"
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   189
  shows "X\<sharp>(ty_dom (Y#\<Gamma>)) = (X\<sharp>(tyvrs_of Y) \<and> X\<sharp>(ty_dom \<Gamma>))"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   190
  apply (nominal_induct rule:binding.strong_induct)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   191
  apply (auto)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   192
  apply (simp add: fresh_def supp_def eqvts)
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   193
  apply (simp add: fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] finite_doms)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   194
  apply (simp add: fresh_def supp_def eqvts)
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   195
  apply (simp add: fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] finite_doms)+
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   196
  done
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   197
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   198
lemma tyvrs_fresh:
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   199
  fixes   X::"tyvrs"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   200
  assumes "X \<sharp> a" 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   201
  shows   "X \<sharp> tyvrs_of a"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   202
  and     "X \<sharp> vrs_of a"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   203
  using assms
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   204
  apply (nominal_induct a rule:binding.strong_induct)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   205
  apply (auto)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   206
  apply (fresh_guess)+
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   207
done
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   208
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   209
lemma fresh_dom:
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   210
  fixes X::"tyvrs"
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   211
  assumes a: "X\<sharp>\<Gamma>" 
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   212
  shows "X\<sharp>(ty_dom \<Gamma>)"
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   213
using a
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   214
apply(induct \<Gamma>)
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   215
apply(simp add: fresh_set_empty) 
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   216
apply(simp only: fresh_ty_dom_cons)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   217
apply(auto simp add: fresh_prod fresh_list_cons tyvrs_fresh) 
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   218
done
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   219
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   220
text {* Not all lists of type @{typ "env"} are well-formed. One condition
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   221
  requires that in @{term "TVarB X S#\<Gamma>"} all free variables of @{term "S"} must be 
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   222
  in the @{term "ty_dom"} of @{term "\<Gamma>"}, that is @{term "S"} must be @{text "closed"} 
18650
urbanc
parents: 18628
diff changeset
   223
  in @{term "\<Gamma>"}. The set of free variables of @{term "S"} is the 
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   224
  @{text "support"} of @{term "S"}. *}
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   225
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34915
diff changeset
   226
definition "closed_in" :: "ty \<Rightarrow> env \<Rightarrow> bool" ("_ closed'_in _" [100,100] 100) where
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   227
  "S closed_in \<Gamma> \<equiv> (supp S)\<subseteq>(ty_dom \<Gamma>)"
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   228
22537
c55f5631a4ec adapted to nominal_inductive infrastructure
urbanc
parents: 22436
diff changeset
   229
lemma closed_in_eqvt[eqvt]:
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   230
  fixes pi::"tyvrs prm"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   231
  assumes a: "S closed_in \<Gamma>" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   232
  shows "(pi\<bullet>S) closed_in (pi\<bullet>\<Gamma>)"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   233
  using a
26091
f31d4fe763aa updated
urbanc
parents: 23760
diff changeset
   234
proof -
f31d4fe763aa updated
urbanc
parents: 23760
diff changeset
   235
  from a have "pi\<bullet>(S closed_in \<Gamma>)" by (simp add: perm_bool)
f31d4fe763aa updated
urbanc
parents: 23760
diff changeset
   236
  then show "(pi\<bullet>S) closed_in (pi\<bullet>\<Gamma>)" by (simp add: closed_in_def eqvts)
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   237
qed
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   238
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   239
lemma tyvrs_vrs_prm_simp:
22537
c55f5631a4ec adapted to nominal_inductive infrastructure
urbanc
parents: 22436
diff changeset
   240
  fixes pi::"vrs prm"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   241
  shows "tyvrs_of (pi\<bullet>a) = tyvrs_of a"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   242
  apply (nominal_induct rule:binding.strong_induct) 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   243
  apply (simp_all add: eqvts)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   244
  apply (simp add: dj_perm_forget[OF dj_tyvrs_vrs])
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   245
  done
22537
c55f5631a4ec adapted to nominal_inductive infrastructure
urbanc
parents: 22436
diff changeset
   246
30986
047fa04a9fe8 deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
Christian Urban <urbanc@in.tum.de>
parents: 30091
diff changeset
   247
lemma ty_vrs_fresh:
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   248
  fixes x::"vrs"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   249
  and   T::"ty"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   250
  shows "x \<sharp> T"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   251
by (simp add: fresh_def supp_def ty_vrs_prm_simp)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   252
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   253
lemma ty_dom_vrs_prm_simp:
22537
c55f5631a4ec adapted to nominal_inductive infrastructure
urbanc
parents: 22436
diff changeset
   254
  fixes pi::"vrs prm"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   255
  and   \<Gamma>::"env"
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   256
  shows "(ty_dom (pi\<bullet>\<Gamma>)) = (ty_dom \<Gamma>)"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   257
  apply(induct \<Gamma>) 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   258
  apply (simp add: eqvts)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   259
  apply(simp add:  tyvrs_vrs_prm_simp)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   260
done
22537
c55f5631a4ec adapted to nominal_inductive infrastructure
urbanc
parents: 22436
diff changeset
   261
c55f5631a4ec adapted to nominal_inductive infrastructure
urbanc
parents: 22436
diff changeset
   262
lemma closed_in_eqvt'[eqvt]:
c55f5631a4ec adapted to nominal_inductive infrastructure
urbanc
parents: 22436
diff changeset
   263
  fixes pi::"vrs prm"
c55f5631a4ec adapted to nominal_inductive infrastructure
urbanc
parents: 22436
diff changeset
   264
  assumes a: "S closed_in \<Gamma>" 
c55f5631a4ec adapted to nominal_inductive infrastructure
urbanc
parents: 22436
diff changeset
   265
  shows "(pi\<bullet>S) closed_in (pi\<bullet>\<Gamma>)"
c55f5631a4ec adapted to nominal_inductive infrastructure
urbanc
parents: 22436
diff changeset
   266
using a
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   267
by (simp add: closed_in_def ty_dom_vrs_prm_simp  ty_vrs_prm_simp)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   268
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   269
lemma fresh_vrs_of: 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   270
  fixes x::"vrs"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   271
  shows "x\<sharp>vrs_of b = x\<sharp>b"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   272
  by (nominal_induct b rule: binding.strong_induct)
46182
b4aa5e39f944 Removed strange hack introduced in b27e93132603, since equivariance
berghofe
parents: 45971
diff changeset
   273
    (simp_all add: fresh_singleton fresh_set_empty ty_vrs_fresh fresh_atm)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   274
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   275
lemma fresh_trm_dom: 
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   276
  fixes x::"vrs"
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   277
  shows "x\<sharp> trm_dom \<Gamma> = x\<sharp>\<Gamma>"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   278
  by (induct \<Gamma>)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   279
    (simp_all add: fresh_set_empty fresh_list_cons
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   280
     fresh_fin_union [OF pt_vrs_inst at_vrs_inst fs_vrs_inst]
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   281
     finite_doms finite_vrs fresh_vrs_of fresh_list_nil)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   282
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   283
lemma closed_in_fresh: "(X::tyvrs) \<sharp> ty_dom \<Gamma> \<Longrightarrow> T closed_in \<Gamma> \<Longrightarrow> X \<sharp> T"
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   284
  by (auto simp add: closed_in_def fresh_def ty_dom_supp)
22537
c55f5631a4ec adapted to nominal_inductive infrastructure
urbanc
parents: 22436
diff changeset
   285
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   286
text {* Now validity of a context is a straightforward inductive definition. *}
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   287
  
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   288
inductive
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   289
  valid_rel :: "env \<Rightarrow> bool" ("\<turnstile> _ ok" [100] 100)
22436
c9e384a956df Adapted to new inductive definition package.
berghofe
parents: 22418
diff changeset
   290
where
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   291
  valid_nil[simp]:   "\<turnstile> [] ok"
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   292
| valid_consT[simp]: "\<lbrakk>\<turnstile> \<Gamma> ok; X\<sharp>(ty_dom  \<Gamma>); T closed_in \<Gamma>\<rbrakk>  \<Longrightarrow>  \<turnstile> (TVarB X T#\<Gamma>) ok"
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   293
| valid_cons [simp]: "\<lbrakk>\<turnstile> \<Gamma> ok; x\<sharp>(trm_dom \<Gamma>); T closed_in \<Gamma>\<rbrakk>  \<Longrightarrow>  \<turnstile> (VarB  x T#\<Gamma>) ok"
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   294
22537
c55f5631a4ec adapted to nominal_inductive infrastructure
urbanc
parents: 22436
diff changeset
   295
equivariance valid_rel
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   296
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   297
declare binding.inject [simp add]
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   298
declare trm.inject     [simp add]
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   299
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   300
inductive_cases validE[elim]: 
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   301
  "\<turnstile> (TVarB X T#\<Gamma>) ok" 
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   302
  "\<turnstile> (VarB  x T#\<Gamma>) ok" 
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   303
  "\<turnstile> (b#\<Gamma>) ok" 
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   304
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   305
declare binding.inject [simp del]
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   306
declare trm.inject     [simp del]
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   307
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   308
lemma validE_append:
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   309
  assumes a: "\<turnstile> (\<Delta>@\<Gamma>) ok" 
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   310
  shows "\<turnstile> \<Gamma> ok"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   311
  using a 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   312
proof (induct \<Delta>)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   313
  case (Cons a \<Gamma>')
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   314
  then show ?case 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   315
    by (nominal_induct a rule:binding.strong_induct)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   316
       (auto elim: validE)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   317
qed (auto)
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   318
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   319
lemma replace_type:
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   320
  assumes a: "\<turnstile> (\<Delta>@(TVarB X T)#\<Gamma>) ok"
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   321
  and     b: "S closed_in \<Gamma>"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   322
  shows "\<turnstile> (\<Delta>@(TVarB X S)#\<Gamma>) ok"
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   323
using a b
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   324
proof(induct \<Delta>)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   325
  case Nil
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   326
  then show ?case by (auto elim: validE intro: valid_cons simp add: doms_append closed_in_def)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   327
next
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   328
  case (Cons a \<Gamma>')
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   329
  then show ?case 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   330
    by (nominal_induct a rule:binding.strong_induct)
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   331
       (auto elim: validE intro!: valid_cons simp add: doms_append closed_in_def)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   332
qed
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   333
18650
urbanc
parents: 18628
diff changeset
   334
text {* Well-formed contexts have a unique type-binding for a type-variable. *} 
urbanc
parents: 18628
diff changeset
   335
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   336
lemma uniqueness_of_ctxt:
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   337
  fixes \<Gamma>::"env"
18412
9f6b3e1da352 tuned the proof of transitivity/narrowing
urbanc
parents: 18410
diff changeset
   338
  assumes a: "\<turnstile> \<Gamma> ok"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   339
  and     b: "(TVarB X T)\<in>set \<Gamma>"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   340
  and     c: "(TVarB X S)\<in>set \<Gamma>"
18412
9f6b3e1da352 tuned the proof of transitivity/narrowing
urbanc
parents: 18410
diff changeset
   341
  shows "T=S"
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   342
using a b c
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   343
proof (induct)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   344
  case (valid_consT \<Gamma> X' T')
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   345
  moreover
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   346
  { fix \<Gamma>'::"env"
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   347
    assume a: "X'\<sharp>(ty_dom \<Gamma>')" 
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   348
    have "\<not>(\<exists>T.(TVarB X' T)\<in>(set \<Gamma>'))" using a 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   349
    proof (induct \<Gamma>')
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   350
      case (Cons Y \<Gamma>')
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   351
      thus "\<not> (\<exists>T.(TVarB X' T)\<in>set(Y#\<Gamma>'))"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   352
        by (simp add:  fresh_ty_dom_cons 
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   353
                       fresh_fin_union[OF pt_tyvrs_inst  at_tyvrs_inst fs_tyvrs_inst]  
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   354
                       finite_vrs finite_doms, 
46182
b4aa5e39f944 Removed strange hack introduced in b27e93132603, since equivariance
berghofe
parents: 45971
diff changeset
   355
            auto simp add: fresh_atm fresh_singleton)
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   356
    qed (simp)
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   357
  }
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   358
  ultimately show "T=S" by (auto simp add: binding.inject)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   359
qed (auto)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   360
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   361
lemma uniqueness_of_ctxt':
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   362
  fixes \<Gamma>::"env"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   363
  assumes a: "\<turnstile> \<Gamma> ok"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   364
  and     b: "(VarB x T)\<in>set \<Gamma>"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   365
  and     c: "(VarB x S)\<in>set \<Gamma>"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   366
  shows "T=S"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   367
using a b c
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   368
proof (induct)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   369
  case (valid_cons \<Gamma> x' T')
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   370
  moreover
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   371
  { fix \<Gamma>'::"env"
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   372
    assume a: "x'\<sharp>(trm_dom \<Gamma>')" 
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   373
    have "\<not>(\<exists>T.(VarB x' T)\<in>(set \<Gamma>'))" using a 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   374
    proof (induct \<Gamma>')
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   375
      case (Cons y \<Gamma>')
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   376
      thus "\<not> (\<exists>T.(VarB x' T)\<in>set(y#\<Gamma>'))" 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   377
        by (simp add:  fresh_fin_union[OF pt_vrs_inst  at_vrs_inst fs_vrs_inst]  
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   378
                       finite_vrs finite_doms, 
46182
b4aa5e39f944 Removed strange hack introduced in b27e93132603, since equivariance
berghofe
parents: 45971
diff changeset
   379
            auto simp add: fresh_atm fresh_singleton)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   380
    qed (simp)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   381
  }
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   382
  ultimately show "T=S" by (auto simp add: binding.inject)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   383
qed (auto)
18577
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
   384
18628
urbanc
parents: 18621
diff changeset
   385
section {* Size and Capture-Avoiding Substitution for Types *}
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   386
21554
0625898865a9 Adapted to new nominal_primrec command.
berghofe
parents: 21377
diff changeset
   387
nominal_primrec
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 26966
diff changeset
   388
  size_ty :: "ty \<Rightarrow> nat"
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 26966
diff changeset
   389
where
21554
0625898865a9 Adapted to new nominal_primrec command.
berghofe
parents: 21377
diff changeset
   390
  "size_ty (Tvar X) = 1"
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 26966
diff changeset
   391
| "size_ty (Top) = 1"
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 26966
diff changeset
   392
| "size_ty (T1 \<rightarrow> T2) = (size_ty T1) + (size_ty T2) + 1"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   393
| "X \<sharp> T1 \<Longrightarrow> size_ty (\<forall>X<:T1. T2) = (size_ty T1) + (size_ty T2) + 1"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21554
diff changeset
   394
  apply (finite_guess)+
21554
0625898865a9 Adapted to new nominal_primrec command.
berghofe
parents: 21377
diff changeset
   395
  apply (rule TrueI)+
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21554
diff changeset
   396
  apply (simp add: fresh_nat)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21554
diff changeset
   397
  apply (fresh_guess)+
21554
0625898865a9 Adapted to new nominal_primrec command.
berghofe
parents: 21377
diff changeset
   398
  done
20395
9a60e3151244 added definition for size and substitution using the recursion
urbanc
parents: 19972
diff changeset
   399
21554
0625898865a9 Adapted to new nominal_primrec command.
berghofe
parents: 21377
diff changeset
   400
nominal_primrec
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   401
  subst_ty :: "ty \<Rightarrow> tyvrs \<Rightarrow> ty \<Rightarrow> ty" ("_[_ \<mapsto> _]\<^sub>\<tau>" [300, 0, 0] 300)
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 26966
diff changeset
   402
where
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   403
  "(Tvar X)[Y \<mapsto> T]\<^sub>\<tau> = (if X=Y then T else Tvar X)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   404
| "(Top)[Y \<mapsto> T]\<^sub>\<tau> = Top"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   405
| "(T\<^isub>1 \<rightarrow> T\<^isub>2)[Y \<mapsto> T]\<^sub>\<tau> = T\<^isub>1[Y \<mapsto> T]\<^sub>\<tau> \<rightarrow> T\<^isub>2[Y \<mapsto> T]\<^sub>\<tau>"
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   406
| "X\<sharp>(Y,T,T\<^isub>1) \<Longrightarrow> (\<forall>X<:T\<^isub>1. T\<^isub>2)[Y \<mapsto> T]\<^sub>\<tau> = (\<forall>X<:T\<^isub>1[Y \<mapsto> T]\<^sub>\<tau>. T\<^isub>2[Y \<mapsto> T]\<^sub>\<tau>)"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21554
diff changeset
   407
  apply (finite_guess)+
21554
0625898865a9 Adapted to new nominal_primrec command.
berghofe
parents: 21377
diff changeset
   408
  apply (rule TrueI)+
0625898865a9 Adapted to new nominal_primrec command.
berghofe
parents: 21377
diff changeset
   409
  apply (simp add: abs_fresh)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21554
diff changeset
   410
  apply (fresh_guess)+
21554
0625898865a9 Adapted to new nominal_primrec command.
berghofe
parents: 21377
diff changeset
   411
  done
18577
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
   412
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   413
lemma subst_eqvt[eqvt]:
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   414
  fixes pi::"tyvrs prm" 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   415
  and   T::"ty"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   416
  shows "pi\<bullet>(T[X \<mapsto> T']\<^sub>\<tau>) = (pi\<bullet>T)[(pi\<bullet>X) \<mapsto> (pi\<bullet>T')]\<^sub>\<tau>"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   417
  by (nominal_induct T avoiding: X T' rule: ty.strong_induct)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   418
     (perm_simp add: fresh_bij)+
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   419
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   420
lemma subst_eqvt'[eqvt]:
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   421
  fixes pi::"vrs prm" 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   422
  and   T::"ty"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   423
  shows "pi\<bullet>(T[X \<mapsto> T']\<^sub>\<tau>) = (pi\<bullet>T)[(pi\<bullet>X) \<mapsto> (pi\<bullet>T')]\<^sub>\<tau>"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   424
  by (nominal_induct T avoiding: X T' rule: ty.strong_induct)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   425
     (perm_simp add: fresh_left)+
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   426
30986
047fa04a9fe8 deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
Christian Urban <urbanc@in.tum.de>
parents: 30091
diff changeset
   427
lemma type_subst_fresh:
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   428
  fixes X::"tyvrs"
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   429
  assumes "X\<sharp>T" and "X\<sharp>P"
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   430
  shows   "X\<sharp>T[Y \<mapsto> P]\<^sub>\<tau>"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   431
using assms
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   432
by (nominal_induct T avoiding: X Y P rule:ty.strong_induct)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   433
   (auto simp add: abs_fresh)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   434
30986
047fa04a9fe8 deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
Christian Urban <urbanc@in.tum.de>
parents: 30091
diff changeset
   435
lemma fresh_type_subst_fresh:
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   436
    assumes "X\<sharp>T'"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   437
    shows "X\<sharp>T[X \<mapsto> T']\<^sub>\<tau>"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   438
using assms 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   439
by (nominal_induct T avoiding: X T' rule: ty.strong_induct)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   440
   (auto simp add: fresh_atm abs_fresh fresh_nat) 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   441
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   442
lemma type_subst_identity: 
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   443
  "X\<sharp>T \<Longrightarrow> T[X \<mapsto> U]\<^sub>\<tau> = T"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   444
  by (nominal_induct T avoiding: X U rule: ty.strong_induct)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   445
    (simp_all add: fresh_atm abs_fresh)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   446
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   447
lemma type_substitution_lemma:  
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   448
  "X \<noteq> Y \<Longrightarrow> X\<sharp>L \<Longrightarrow> M[X \<mapsto> N]\<^sub>\<tau>[Y \<mapsto> L]\<^sub>\<tau> = M[Y \<mapsto> L]\<^sub>\<tau>[X \<mapsto> N[Y \<mapsto> L]\<^sub>\<tau>]\<^sub>\<tau>"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   449
  by (nominal_induct M avoiding: X Y N L rule: ty.strong_induct)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   450
    (auto simp add: type_subst_fresh type_subst_identity)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   451
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   452
lemma type_subst_rename:
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   453
  "Y\<sharp>T \<Longrightarrow> ([(Y,X)]\<bullet>T)[Y \<mapsto> U]\<^sub>\<tau> = T[X \<mapsto> U]\<^sub>\<tau>"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   454
  by (nominal_induct T avoiding: X Y U rule: ty.strong_induct)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   455
    (simp_all add: fresh_atm calc_atm abs_fresh fresh_aux)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   456
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   457
nominal_primrec
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   458
  subst_tyb :: "binding \<Rightarrow> tyvrs \<Rightarrow> ty \<Rightarrow> binding" ("_[_ \<mapsto> _]\<^sub>b" [100,100,100] 100)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   459
where
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   460
  "(TVarB X U)[Y \<mapsto> T]\<^sub>b = TVarB X (U[Y \<mapsto> T]\<^sub>\<tau>)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   461
| "(VarB  X U)[Y \<mapsto> T]\<^sub>b =  VarB X (U[Y \<mapsto> T]\<^sub>\<tau>)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   462
by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   463
30986
047fa04a9fe8 deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
Christian Urban <urbanc@in.tum.de>
parents: 30091
diff changeset
   464
lemma binding_subst_fresh:
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   465
  fixes X::"tyvrs"
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   466
  assumes "X\<sharp>a"
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   467
  and     "X\<sharp>P"
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   468
  shows "X\<sharp>a[Y \<mapsto> P]\<^sub>b"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   469
using assms
30986
047fa04a9fe8 deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
Christian Urban <urbanc@in.tum.de>
parents: 30091
diff changeset
   470
by (nominal_induct a rule: binding.strong_induct)
047fa04a9fe8 deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
Christian Urban <urbanc@in.tum.de>
parents: 30091
diff changeset
   471
   (auto simp add: type_subst_fresh)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   472
30986
047fa04a9fe8 deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
Christian Urban <urbanc@in.tum.de>
parents: 30091
diff changeset
   473
lemma binding_subst_identity: 
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   474
  shows "X\<sharp>B \<Longrightarrow> B[X \<mapsto> U]\<^sub>b = B"
30986
047fa04a9fe8 deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
Christian Urban <urbanc@in.tum.de>
parents: 30091
diff changeset
   475
by (induct B rule: binding.induct)
047fa04a9fe8 deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
Christian Urban <urbanc@in.tum.de>
parents: 30091
diff changeset
   476
   (simp_all add: fresh_atm type_subst_identity)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   477
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
   478
primrec subst_tyc :: "env \<Rightarrow> tyvrs \<Rightarrow> ty \<Rightarrow> env" ("_[_ \<mapsto> _]\<^sub>e" [100,100,100] 100) where
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
   479
  "([])[Y \<mapsto> T]\<^sub>e= []"
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
   480
| "(B#\<Gamma>)[Y \<mapsto> T]\<^sub>e = (B[Y \<mapsto> T]\<^sub>b)#(\<Gamma>[Y \<mapsto> T]\<^sub>e)"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   481
30986
047fa04a9fe8 deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
Christian Urban <urbanc@in.tum.de>
parents: 30091
diff changeset
   482
lemma ctxt_subst_fresh':
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   483
  fixes X::"tyvrs"
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   484
  assumes "X\<sharp>\<Gamma>"
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   485
  and     "X\<sharp>P"
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   486
  shows   "X\<sharp>\<Gamma>[Y \<mapsto> P]\<^sub>e"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   487
using assms
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   488
by (induct \<Gamma>)
30986
047fa04a9fe8 deleted thm-attributes "fresh" and "bij" (not used); same features can later be implemented by simpler means
Christian Urban <urbanc@in.tum.de>
parents: 30091
diff changeset
   489
   (auto simp add: fresh_list_cons binding_subst_fresh)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   490
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   491
lemma ctxt_subst_mem_TVarB: "TVarB X T \<in> set \<Gamma> \<Longrightarrow> TVarB X (T[Y \<mapsto> U]\<^sub>\<tau>) \<in> set (\<Gamma>[Y \<mapsto> U]\<^sub>e)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   492
  by (induct \<Gamma>) auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   493
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   494
lemma ctxt_subst_mem_VarB: "VarB x T \<in> set \<Gamma> \<Longrightarrow> VarB x (T[Y \<mapsto> U]\<^sub>\<tau>) \<in> set (\<Gamma>[Y \<mapsto> U]\<^sub>e)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   495
  by (induct \<Gamma>) auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   496
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   497
lemma ctxt_subst_identity: "X\<sharp>\<Gamma> \<Longrightarrow> \<Gamma>[X \<mapsto> U]\<^sub>e = \<Gamma>"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   498
  by (induct \<Gamma>) (simp_all add: fresh_list_cons binding_subst_identity)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   499
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   500
lemma ctxt_subst_append: "(\<Delta> @ \<Gamma>)[X \<mapsto> T]\<^sub>e = \<Delta>[X \<mapsto> T]\<^sub>e @ \<Gamma>[X \<mapsto> T]\<^sub>e"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   501
  by (induct \<Delta>) simp_all
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   502
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   503
nominal_primrec
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   504
   subst_trm :: "trm \<Rightarrow> vrs \<Rightarrow> trm \<Rightarrow> trm"  ("_[_ \<mapsto> _]" [300, 0, 0] 300)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   505
where
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   506
  "(Var x)[y \<mapsto> t'] = (if x=y then t' else (Var x))"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   507
| "(t1 \<cdot> t2)[y \<mapsto> t'] = t1[y \<mapsto> t'] \<cdot> t2[y \<mapsto> t']"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   508
| "(t \<cdot>\<^sub>\<tau> T)[y \<mapsto> t'] = t[y \<mapsto> t'] \<cdot>\<^sub>\<tau> T"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   509
| "X\<sharp>(T,t') \<Longrightarrow> (\<lambda>X<:T. t)[y \<mapsto> t'] = (\<lambda>X<:T. t[y \<mapsto> t'])" 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   510
| "x\<sharp>(y,t') \<Longrightarrow> (\<lambda>x:T. t)[y \<mapsto> t'] = (\<lambda>x:T. t[y \<mapsto> t'])"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   511
apply(finite_guess)+
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   512
apply(rule TrueI)+
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   513
apply(simp add: abs_fresh)+
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   514
apply(fresh_guess add: ty_vrs_fresh abs_fresh)+
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   515
done
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   516
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   517
lemma subst_trm_fresh_tyvar:
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   518
  fixes X::"tyvrs"
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   519
  shows "X\<sharp>t \<Longrightarrow> X\<sharp>u \<Longrightarrow> X\<sharp>t[x \<mapsto> u]"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   520
  by (nominal_induct t avoiding: x u rule: trm.strong_induct)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   521
    (auto simp add: trm.fresh abs_fresh)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   522
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   523
lemma subst_trm_fresh_var: 
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   524
  "x\<sharp>u \<Longrightarrow> x\<sharp>t[x \<mapsto> u]"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   525
  by (nominal_induct t avoiding: x u rule: trm.strong_induct)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   526
    (simp_all add: abs_fresh fresh_atm ty_vrs_fresh)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   527
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   528
lemma subst_trm_eqvt[eqvt]:
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   529
  fixes pi::"tyvrs prm" 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   530
  and   t::"trm"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   531
  shows "pi\<bullet>(t[x \<mapsto> u]) = (pi\<bullet>t)[(pi\<bullet>x) \<mapsto> (pi\<bullet>u)]"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   532
  by (nominal_induct t avoiding: x u rule: trm.strong_induct)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   533
     (perm_simp add: fresh_left)+
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   534
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   535
lemma subst_trm_eqvt'[eqvt]:
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   536
  fixes pi::"vrs prm" 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   537
  and   t::"trm"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   538
  shows "pi\<bullet>(t[x \<mapsto> u]) = (pi\<bullet>t)[(pi\<bullet>x) \<mapsto> (pi\<bullet>u)]"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   539
  by (nominal_induct t avoiding: x u rule: trm.strong_induct)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   540
     (perm_simp add: fresh_left)+
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   541
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   542
lemma subst_trm_rename:
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   543
  "y\<sharp>t \<Longrightarrow> ([(y, x)] \<bullet> t)[y \<mapsto> u] = t[x \<mapsto> u]"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   544
  by (nominal_induct t avoiding: x y u rule: trm.strong_induct)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   545
    (simp_all add: fresh_atm calc_atm abs_fresh fresh_aux ty_vrs_fresh perm_fresh_fresh)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   546
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   547
nominal_primrec (freshness_context: "T2::ty")
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   548
  subst_trm_ty :: "trm \<Rightarrow> tyvrs \<Rightarrow> ty \<Rightarrow> trm"  ("_[_ \<mapsto>\<^sub>\<tau> _]" [300, 0, 0] 300)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   549
where
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   550
  "(Var x)[Y \<mapsto>\<^sub>\<tau> T2] = Var x"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   551
| "(t1 \<cdot> t2)[Y \<mapsto>\<^sub>\<tau> T2] = t1[Y \<mapsto>\<^sub>\<tau> T2] \<cdot> t2[Y \<mapsto>\<^sub>\<tau> T2]"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   552
| "(t1 \<cdot>\<^sub>\<tau> T)[Y \<mapsto>\<^sub>\<tau> T2] = t1[Y \<mapsto>\<^sub>\<tau> T2] \<cdot>\<^sub>\<tau> T[Y \<mapsto> T2]\<^sub>\<tau>"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   553
| "X\<sharp>(Y,T,T2) \<Longrightarrow> (\<lambda>X<:T. t)[Y \<mapsto>\<^sub>\<tau> T2] = (\<lambda>X<:T[Y \<mapsto> T2]\<^sub>\<tau>. t[Y \<mapsto>\<^sub>\<tau> T2])" 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   554
| "(\<lambda>x:T. t)[Y \<mapsto>\<^sub>\<tau> T2] = (\<lambda>x:T[Y \<mapsto> T2]\<^sub>\<tau>. t[Y \<mapsto>\<^sub>\<tau> T2])"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   555
apply(finite_guess)+
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   556
apply(rule TrueI)+
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   557
apply(simp add: abs_fresh ty_vrs_fresh)+
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   558
apply(simp add: type_subst_fresh)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   559
apply(fresh_guess add: ty_vrs_fresh abs_fresh)+
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   560
done
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   561
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   562
lemma subst_trm_ty_fresh:
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   563
  fixes X::"tyvrs"
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   564
  shows "X\<sharp>t \<Longrightarrow> X\<sharp>T \<Longrightarrow> X\<sharp>t[Y \<mapsto>\<^sub>\<tau> T]"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   565
  by (nominal_induct t avoiding: Y T rule: trm.strong_induct)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   566
    (auto simp add: abs_fresh type_subst_fresh)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   567
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   568
lemma subst_trm_ty_fresh':
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   569
  "X\<sharp>T \<Longrightarrow> X\<sharp>t[X \<mapsto>\<^sub>\<tau> T]"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   570
  by (nominal_induct t avoiding: X T rule: trm.strong_induct)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   571
    (simp_all add: abs_fresh fresh_type_subst_fresh fresh_atm)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   572
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   573
lemma subst_trm_ty_eqvt[eqvt]:
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   574
  fixes pi::"tyvrs prm" 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   575
  and   t::"trm"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   576
  shows "pi\<bullet>(t[X \<mapsto>\<^sub>\<tau> T]) = (pi\<bullet>t)[(pi\<bullet>X) \<mapsto>\<^sub>\<tau> (pi\<bullet>T)]"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   577
  by (nominal_induct t avoiding: X T rule: trm.strong_induct)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   578
     (perm_simp add: fresh_bij subst_eqvt)+
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   579
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   580
lemma subst_trm_ty_eqvt'[eqvt]:
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   581
  fixes pi::"vrs prm" 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   582
  and   t::"trm"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   583
  shows "pi\<bullet>(t[X \<mapsto>\<^sub>\<tau> T]) = (pi\<bullet>t)[(pi\<bullet>X) \<mapsto>\<^sub>\<tau> (pi\<bullet>T)]"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   584
  by (nominal_induct t avoiding: X T rule: trm.strong_induct)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   585
     (perm_simp add: fresh_left subst_eqvt')+
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   586
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   587
lemma subst_trm_ty_rename:
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   588
  "Y\<sharp>t \<Longrightarrow> ([(Y, X)] \<bullet> t)[Y \<mapsto>\<^sub>\<tau> U] = t[X \<mapsto>\<^sub>\<tau> U]"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   589
  by (nominal_induct t avoiding: X Y U rule: trm.strong_induct)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   590
    (simp_all add: fresh_atm calc_atm abs_fresh fresh_aux type_subst_rename)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   591
18577
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
   592
section {* Subtyping-Relation *}
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   593
18650
urbanc
parents: 18628
diff changeset
   594
text {* The definition for the subtyping-relation follows quite closely what is written 
urbanc
parents: 18628
diff changeset
   595
  in the POPLmark-paper, except for the premises dealing with well-formed contexts and 
urbanc
parents: 18628
diff changeset
   596
  the freshness constraint @{term "X\<sharp>\<Gamma>"} in the @{text "S_Forall"}-rule. (The freshness
urbanc
parents: 18628
diff changeset
   597
  constraint is specific to the \emph{nominal approach}. Note, however, that the constraint
urbanc
parents: 18628
diff changeset
   598
  does \emph{not} make the subtyping-relation ``partial"\ldots because we work over
urbanc
parents: 18628
diff changeset
   599
  $\alpha$-equivalence classes.) *}
18628
urbanc
parents: 18621
diff changeset
   600
23760
aca2c7f80e2f Renamed inductive2 to inductive.
berghofe
parents: 23393
diff changeset
   601
inductive 
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   602
  subtype_of :: "env \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> bool"   ("_\<turnstile>_<:_" [100,100,100] 100)
22436
c9e384a956df Adapted to new inductive definition package.
berghofe
parents: 22418
diff changeset
   603
where
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   604
  SA_Top[intro]:    "\<lbrakk>\<turnstile> \<Gamma> ok; S closed_in \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> S <: Top"
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   605
| SA_refl_TVar[intro]:   "\<lbrakk>\<turnstile> \<Gamma> ok; X \<in> ty_dom \<Gamma>\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> Tvar X <: Tvar X"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   606
| SA_trans_TVar[intro]:    "\<lbrakk>(TVarB X S) \<in> set \<Gamma>; \<Gamma> \<turnstile> S <: T\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> (Tvar X) <: T"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   607
| SA_arrow[intro]:  "\<lbrakk>\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1; \<Gamma> \<turnstile> S\<^isub>2 <: T\<^isub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> (S\<^isub>1 \<rightarrow> S\<^isub>2) <: (T\<^isub>1 \<rightarrow> T\<^isub>2)" 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   608
| SA_all[intro]: "\<lbrakk>\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1; ((TVarB X T\<^isub>1)#\<Gamma>) \<turnstile> S\<^isub>2 <: T\<^isub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> (\<forall>X<:S\<^isub>1. S\<^isub>2) <: (\<forall>X<:T\<^isub>1. T\<^isub>2)"
22537
c55f5631a4ec adapted to nominal_inductive infrastructure
urbanc
parents: 22436
diff changeset
   609
c55f5631a4ec adapted to nominal_inductive infrastructure
urbanc
parents: 22436
diff changeset
   610
lemma subtype_implies_ok:
c55f5631a4ec adapted to nominal_inductive infrastructure
urbanc
parents: 22436
diff changeset
   611
  fixes X::"tyvrs"
c55f5631a4ec adapted to nominal_inductive infrastructure
urbanc
parents: 22436
diff changeset
   612
  assumes a: "\<Gamma> \<turnstile> S <: T"
c55f5631a4ec adapted to nominal_inductive infrastructure
urbanc
parents: 22436
diff changeset
   613
  shows "\<turnstile> \<Gamma> ok"  
c55f5631a4ec adapted to nominal_inductive infrastructure
urbanc
parents: 22436
diff changeset
   614
using a by (induct) (auto)
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   615
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   616
lemma subtype_implies_closed:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   617
  assumes a: "\<Gamma> \<turnstile> S <: T"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   618
  shows "S closed_in \<Gamma> \<and> T closed_in \<Gamma>"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   619
using a
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   620
proof (induct)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   621
  case (SA_Top \<Gamma> S)
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   622
  have "Top closed_in \<Gamma>" by (simp add: closed_in_def ty.supp)
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   623
  moreover
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   624
  have "S closed_in \<Gamma>" by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   625
  ultimately show "S closed_in \<Gamma> \<and> Top closed_in \<Gamma>" by simp
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   626
next
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   627
  case (SA_trans_TVar X S \<Gamma> T)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   628
  have "(TVarB X S)\<in>set \<Gamma>" by fact
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   629
  hence "X \<in> ty_dom \<Gamma>" by (rule ty_dom_inclusion)
18577
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
   630
  hence "(Tvar X) closed_in \<Gamma>" by (simp add: closed_in_def ty.supp supp_atm)
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   631
  moreover
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   632
  have "S closed_in \<Gamma> \<and> T closed_in \<Gamma>" by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   633
  hence "T closed_in \<Gamma>" by force
18577
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
   634
  ultimately show "(Tvar X) closed_in \<Gamma> \<and> T closed_in \<Gamma>" by simp
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   635
qed (auto simp add: closed_in_def ty.supp supp_atm abs_supp)
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   636
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   637
lemma subtype_implies_fresh:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   638
  fixes X::"tyvrs"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   639
  assumes a1: "\<Gamma> \<turnstile> S <: T"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   640
  and     a2: "X\<sharp>\<Gamma>"
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   641
  shows "X\<sharp>S \<and> X\<sharp>T"  
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   642
proof -
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   643
  from a1 have "\<turnstile> \<Gamma> ok" by (rule subtype_implies_ok)
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   644
  with a2 have "X\<sharp>ty_dom(\<Gamma>)" by (simp add: fresh_dom)
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   645
  moreover
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   646
  from a1 have "S closed_in \<Gamma> \<and> T closed_in \<Gamma>" by (rule subtype_implies_closed)
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   647
  hence "supp S \<subseteq> ((supp (ty_dom \<Gamma>))::tyvrs set)" 
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   648
    and "supp T \<subseteq> ((supp (ty_dom \<Gamma>))::tyvrs set)" by (simp_all add: ty_dom_supp closed_in_def)
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   649
  ultimately show "X\<sharp>S \<and> X\<sharp>T" by (force simp add: supp_prod fresh_def)
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   650
qed
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   651
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   652
lemma valid_ty_dom_fresh:
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   653
  fixes X::"tyvrs"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   654
  assumes valid: "\<turnstile> \<Gamma> ok"
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   655
  shows "X\<sharp>(ty_dom \<Gamma>) = X\<sharp>\<Gamma>" 
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   656
  using valid
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   657
  apply induct
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   658
  apply (simp add: fresh_list_nil fresh_set_empty)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   659
  apply (simp_all add: binding.fresh fresh_list_cons
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   660
     fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] finite_doms fresh_atm)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   661
  apply (auto simp add: closed_in_fresh)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   662
  done
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   663
22730
8bcc8809ed3b nominal_inductive no longer proves equivariance.
berghofe
parents: 22542
diff changeset
   664
equivariance subtype_of
8bcc8809ed3b nominal_inductive no longer proves equivariance.
berghofe
parents: 22542
diff changeset
   665
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   666
nominal_inductive subtype_of
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   667
  apply (simp_all add: abs_fresh)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 41798
diff changeset
   668
  apply (fastforce simp add: valid_ty_dom_fresh dest: subtype_implies_ok)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   669
  apply (force simp add: closed_in_fresh dest: subtype_implies_closed subtype_implies_ok)+
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   670
  done
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   671
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   672
section {* Reflexivity of Subtyping *}
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   673
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   674
lemma subtype_reflexivity:
18353
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   675
  assumes a: "\<turnstile> \<Gamma> ok"
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   676
  and b: "T closed_in \<Gamma>"
18353
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   677
  shows "\<Gamma> \<turnstile> T <: T"
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   678
using a b
26966
071f40487734 made the naming of the induction principles consistent: weak_induct is
urbanc
parents: 26091
diff changeset
   679
proof(nominal_induct T avoiding: \<Gamma> rule: ty.strong_induct)
18577
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
   680
  case (Forall X T\<^isub>1 T\<^isub>2)
18747
7dd9aa160b6c no essential changes
urbanc
parents: 18660
diff changeset
   681
  have ih_T\<^isub>1: "\<And>\<Gamma>. \<lbrakk>\<turnstile> \<Gamma> ok; T\<^isub>1 closed_in \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> T\<^isub>1 <: T\<^isub>1" by fact 
7dd9aa160b6c no essential changes
urbanc
parents: 18660
diff changeset
   682
  have ih_T\<^isub>2: "\<And>\<Gamma>. \<lbrakk>\<turnstile> \<Gamma> ok; T\<^isub>2 closed_in \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> T\<^isub>2 <: T\<^isub>2" by fact
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   683
  have fresh_cond: "X\<sharp>\<Gamma>" by fact
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   684
  hence fresh_ty_dom: "X\<sharp>(ty_dom \<Gamma>)" by (simp add: fresh_dom)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   685
  have "(\<forall>X<:T\<^isub>2. T\<^isub>1) closed_in \<Gamma>" by fact
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
   686
  hence closed\<^isub>T2: "T\<^isub>2 closed_in \<Gamma>" and closed\<^isub>T1: "T\<^isub>1 closed_in ((TVarB  X T\<^isub>2)#\<Gamma>)" 
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   687
    by (auto simp add: closed_in_def ty.supp abs_supp)
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   688
  have ok: "\<turnstile> \<Gamma> ok" by fact  
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
   689
  hence ok': "\<turnstile> ((TVarB X T\<^isub>2)#\<Gamma>) ok" using closed\<^isub>T2 fresh_ty_dom by simp
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
   690
  have "\<Gamma> \<turnstile> T\<^isub>2 <: T\<^isub>2" using ih_T\<^isub>2 closed\<^isub>T2 ok by simp
18353
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   691
  moreover
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
   692
  have "((TVarB X T\<^isub>2)#\<Gamma>) \<turnstile> T\<^isub>1 <: T\<^isub>1" using ih_T\<^isub>1 closed\<^isub>T1 ok' by simp
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   693
  ultimately show "\<Gamma> \<turnstile> (\<forall>X<:T\<^isub>2. T\<^isub>1) <: (\<forall>X<:T\<^isub>2. T\<^isub>1)" using fresh_cond 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   694
    by (simp add: subtype_of.SA_all)
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   695
qed (auto simp add: closed_in_def ty.supp supp_atm)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   696
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   697
lemma subtype_reflexivity_semiautomated:
18305
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   698
  assumes a: "\<turnstile> \<Gamma> ok"
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   699
  and     b: "T closed_in \<Gamma>"
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   700
  shows "\<Gamma> \<turnstile> T <: T"
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   701
using a b
26966
071f40487734 made the naming of the induction principles consistent: weak_induct is
urbanc
parents: 26091
diff changeset
   702
apply(nominal_induct T avoiding: \<Gamma> rule: ty.strong_induct)
18747
7dd9aa160b6c no essential changes
urbanc
parents: 18660
diff changeset
   703
apply(auto simp add: ty.supp abs_supp supp_atm closed_in_def)
18577
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
   704
  --{* Too bad that this instantiation cannot be found automatically by
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   705
  \isakeyword{auto}; \isakeyword{blast} would find it if we had not used 
18628
urbanc
parents: 18621
diff changeset
   706
  an explicit definition for @{text "closed_in_def"}. *}
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   707
apply(drule_tac x="(TVarB tyvrs ty2)#\<Gamma>" in meta_spec)
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   708
apply(force dest: fresh_dom simp add: closed_in_def)
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   709
done
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   710
18628
urbanc
parents: 18621
diff changeset
   711
section {* Weakening *}
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   712
18628
urbanc
parents: 18621
diff changeset
   713
text {* In order to prove weakening we introduce the notion of a type-context extending 
urbanc
parents: 18621
diff changeset
   714
  another. This generalization seems to make the proof for weakening to be
urbanc
parents: 18621
diff changeset
   715
  smoother than if we had strictly adhered to the version in the POPLmark-paper. *}
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   716
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 34915
diff changeset
   717
definition extends :: "env \<Rightarrow> env \<Rightarrow> bool" ("_ extends _" [100,100] 100) where
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   718
  "\<Delta> extends \<Gamma> \<equiv> \<forall>X Q. (TVarB X Q)\<in>set \<Gamma> \<longrightarrow> (TVarB X Q)\<in>set \<Delta>"
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   719
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   720
lemma extends_ty_dom:
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   721
  assumes a: "\<Delta> extends \<Gamma>"
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   722
  shows "ty_dom \<Gamma> \<subseteq> ty_dom \<Delta>"
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   723
  using a 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   724
  apply (auto simp add: extends_def)
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   725
  apply (drule ty_dom_existence)
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   726
  apply (force simp add: ty_dom_inclusion)
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   727
  done
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   728
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   729
lemma extends_closed:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   730
  assumes a1: "T closed_in \<Gamma>"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   731
  and     a2: "\<Delta> extends \<Gamma>"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   732
  shows "T closed_in \<Delta>"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   733
  using a1 a2
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   734
  by (auto dest: extends_ty_dom simp add: closed_in_def)
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   735
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   736
lemma extends_memb:
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   737
  assumes a: "\<Delta> extends \<Gamma>"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   738
  and b: "(TVarB X T) \<in> set \<Gamma>"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   739
  shows "(TVarB X T) \<in> set \<Delta>"
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   740
  using a b by (simp add: extends_def)
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   741
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   742
lemma weakening:
18353
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   743
  assumes a: "\<Gamma> \<turnstile> S <: T"
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   744
  and b: "\<turnstile> \<Delta> ok"
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   745
  and c: "\<Delta> extends \<Gamma>"
18353
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   746
  shows "\<Delta> \<turnstile> S <: T"
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   747
  using a b c 
22537
c55f5631a4ec adapted to nominal_inductive infrastructure
urbanc
parents: 22436
diff changeset
   748
proof (nominal_induct \<Gamma> S T avoiding: \<Delta> rule: subtype_of.strong_induct)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   749
  case (SA_Top \<Gamma> S) 
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   750
  have lh_drv_prem: "S closed_in \<Gamma>" by fact
18353
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   751
  have "\<turnstile> \<Delta> ok" by fact
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   752
  moreover
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   753
  have "\<Delta> extends \<Gamma>" by fact
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   754
  hence "S closed_in \<Delta>" using lh_drv_prem by (simp only: extends_closed)
18353
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   755
  ultimately show "\<Delta> \<turnstile> S <: Top" by force
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   756
next 
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   757
  case (SA_trans_TVar X S \<Gamma> T)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   758
  have lh_drv_prem: "(TVarB X S) \<in> set \<Gamma>" by fact
18353
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   759
  have ih: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<Longrightarrow> \<Delta> extends \<Gamma> \<Longrightarrow> \<Delta> \<turnstile> S <: T" by fact
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   760
  have ok: "\<turnstile> \<Delta> ok" by fact
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   761
  have extends: "\<Delta> extends \<Gamma>" by fact
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   762
  have "(TVarB X S) \<in> set \<Delta>" using lh_drv_prem extends by (simp only: extends_memb)
18353
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   763
  moreover
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   764
  have "\<Delta> \<turnstile> S <: T" using ok extends ih by simp
18577
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
   765
  ultimately show "\<Delta> \<turnstile> Tvar X <: T" using ok by force
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   766
next
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   767
  case (SA_refl_TVar \<Gamma> X)
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   768
  have lh_drv_prem: "X \<in> ty_dom \<Gamma>" by fact
18353
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   769
  have "\<turnstile> \<Delta> ok" by fact
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   770
  moreover
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   771
  have "\<Delta> extends \<Gamma>" by fact
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   772
  hence "X \<in> ty_dom \<Delta>" using lh_drv_prem by (force dest: extends_ty_dom)
18577
a636846a02c7 added more documentation; will now try out a modification
urbanc
parents: 18424
diff changeset
   773
  ultimately show "\<Delta> \<turnstile> Tvar X <: Tvar X" by force
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   774
next 
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   775
  case (SA_arrow \<Gamma> T\<^isub>1 S\<^isub>1 S\<^isub>2 T\<^isub>2) thus "\<Delta> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T\<^isub>1 \<rightarrow> T\<^isub>2" by blast
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   776
next
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   777
  case (SA_all \<Gamma> T\<^isub>1 S\<^isub>1 X S\<^isub>2 T\<^isub>2)
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   778
  have fresh_cond: "X\<sharp>\<Delta>" by fact
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   779
  hence fresh_dom: "X\<sharp>(ty_dom \<Delta>)" by (simp add: fresh_dom)
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   780
  have ih\<^isub>1: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<Longrightarrow> \<Delta> extends \<Gamma> \<Longrightarrow> \<Delta> \<turnstile> T\<^isub>1 <: S\<^isub>1" by fact
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   781
  have ih\<^isub>2: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<Longrightarrow> \<Delta> extends ((TVarB X T\<^isub>1)#\<Gamma>) \<Longrightarrow> \<Delta> \<turnstile> S\<^isub>2 <: T\<^isub>2" by fact
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   782
  have lh_drv_prem: "\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1" by fact
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
   783
  hence closed\<^isub>T1: "T\<^isub>1 closed_in \<Gamma>" by (simp add: subtype_implies_closed) 
18353
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   784
  have ok: "\<turnstile> \<Delta> ok" by fact
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   785
  have ext: "\<Delta> extends \<Gamma>" by fact
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
   786
  have "T\<^isub>1 closed_in \<Delta>" using ext closed\<^isub>T1 by (simp only: extends_closed)
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   787
  hence "\<turnstile> ((TVarB X T\<^isub>1)#\<Delta>) ok" using fresh_dom ok by force   
18353
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   788
  moreover 
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   789
  have "((TVarB X T\<^isub>1)#\<Delta>) extends ((TVarB X T\<^isub>1)#\<Gamma>)" using ext by (force simp add: extends_def)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   790
  ultimately have "((TVarB X T\<^isub>1)#\<Delta>) \<turnstile> S\<^isub>2 <: T\<^isub>2" using ih\<^isub>2 by simp
18353
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   791
  moreover
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   792
  have "\<Delta> \<turnstile> T\<^isub>1 <: S\<^isub>1" using ok ext ih\<^isub>1 by simp 
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   793
  ultimately show "\<Delta> \<turnstile> (\<forall>X<:S\<^isub>1. S\<^isub>2) <: (\<forall>X<:T\<^isub>1. T\<^isub>2)" using ok by (force intro: SA_all)
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   794
qed
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   795
18650
urbanc
parents: 18628
diff changeset
   796
text {* In fact all ``non-binding" cases can be solved automatically: *}
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   797
18628
urbanc
parents: 18621
diff changeset
   798
lemma weakening_more_automated:
18353
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   799
  assumes a: "\<Gamma> \<turnstile> S <: T"
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   800
  and b: "\<turnstile> \<Delta> ok"
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   801
  and c: "\<Delta> extends \<Gamma>"
18353
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   802
  shows "\<Delta> \<turnstile> S <: T"
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   803
  using a b c 
22537
c55f5631a4ec adapted to nominal_inductive infrastructure
urbanc
parents: 22436
diff changeset
   804
proof (nominal_induct \<Gamma> S T avoiding: \<Delta> rule: subtype_of.strong_induct)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   805
  case (SA_all \<Gamma> T\<^isub>1 S\<^isub>1 X S\<^isub>2 T\<^isub>2)
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   806
  have fresh_cond: "X\<sharp>\<Delta>" by fact
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   807
  hence fresh_dom: "X\<sharp>(ty_dom \<Delta>)" by (simp add: fresh_dom)
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   808
  have ih\<^isub>1: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<Longrightarrow> \<Delta> extends \<Gamma> \<Longrightarrow> \<Delta> \<turnstile> T\<^isub>1 <: S\<^isub>1" by fact
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   809
  have ih\<^isub>2: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<Longrightarrow> \<Delta> extends ((TVarB X T\<^isub>1)#\<Gamma>) \<Longrightarrow> \<Delta> \<turnstile> S\<^isub>2 <: T\<^isub>2" by fact
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   810
  have lh_drv_prem: "\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1" by fact
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
   811
  hence closed\<^isub>T1: "T\<^isub>1 closed_in \<Gamma>" by (simp add: subtype_implies_closed) 
18353
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   812
  have ok: "\<turnstile> \<Delta> ok" by fact
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   813
  have ext: "\<Delta> extends \<Gamma>" by fact
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
   814
  have "T\<^isub>1 closed_in \<Delta>" using ext closed\<^isub>T1 by (simp only: extends_closed)
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   815
  hence "\<turnstile> ((TVarB X T\<^isub>1)#\<Delta>) ok" using fresh_dom ok by force   
18628
urbanc
parents: 18621
diff changeset
   816
  moreover
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   817
  have "((TVarB X T\<^isub>1)#\<Delta>) extends ((TVarB X T\<^isub>1)#\<Gamma>)" using ext by (force simp add: extends_def)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   818
  ultimately have "((TVarB X T\<^isub>1)#\<Delta>) \<turnstile> S\<^isub>2 <: T\<^isub>2" using ih\<^isub>2 by simp
18353
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   819
  moreover
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   820
  have "\<Delta> \<turnstile> T\<^isub>1 <: S\<^isub>1" using ok ext ih\<^isub>1 by simp 
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   821
  ultimately show "\<Delta> \<turnstile> (\<forall>X<:S\<^isub>1. S\<^isub>2) <: (\<forall>X<:T\<^isub>1. T\<^isub>2)" using ok by (force intro: SA_all)
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   822
qed (blast intro: extends_closed extends_memb dest: extends_ty_dom)+
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   823
18628
urbanc
parents: 18621
diff changeset
   824
section {* Transitivity and Narrowing *}
urbanc
parents: 18621
diff changeset
   825
18650
urbanc
parents: 18628
diff changeset
   826
text {* Some inversion lemmas that are needed in the transitivity and narrowing proof.*}
urbanc
parents: 18628
diff changeset
   827
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   828
declare ty.inject [simp add]
18650
urbanc
parents: 18628
diff changeset
   829
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   830
inductive_cases S_TopE: "\<Gamma> \<turnstile> Top <: T"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   831
inductive_cases S_ArrowE_left: "\<Gamma> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T" 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   832
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   833
declare ty.inject [simp del]
18650
urbanc
parents: 18628
diff changeset
   834
urbanc
parents: 18628
diff changeset
   835
lemma S_ForallE_left:
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   836
  shows "\<lbrakk>\<Gamma> \<turnstile> (\<forall>X<:S\<^isub>1. S\<^isub>2) <: T; X\<sharp>\<Gamma>; X\<sharp>S\<^isub>1; X\<sharp>T\<rbrakk>
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   837
         \<Longrightarrow> T = Top \<or> (\<exists>T\<^isub>1 T\<^isub>2. T = (\<forall>X<:T\<^isub>1. T\<^isub>2) \<and> \<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1 \<and> ((TVarB X T\<^isub>1)#\<Gamma>) \<turnstile> S\<^isub>2 <: T\<^isub>2)"
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   838
apply(erule subtype_of.strong_cases[where X="X"])
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   839
apply(auto simp add: abs_fresh ty.inject alpha)
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   840
done
18650
urbanc
parents: 18628
diff changeset
   841
urbanc
parents: 18628
diff changeset
   842
text {* Next we prove the transitivity and narrowing for the subtyping-relation. 
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   843
The POPLmark-paper says the following:
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   844
18650
urbanc
parents: 18628
diff changeset
   845
\begin{quote}
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   846
\begin{lemma}[Transitivity and Narrowing] \
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   847
\begin{enumerate}
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   848
\item If @{term "\<Gamma> \<turnstile> S<:Q"} and @{term "\<Gamma> \<turnstile> Q<:T"}, then @{term "\<Gamma> \<turnstile> S<:T"}.
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   849
\item If @{text "\<Gamma>,X<:Q,\<Delta> \<turnstile> M<:N"} and @{term "\<Gamma> \<turnstile> P<:Q"} then @{text "\<Gamma>,X<:P,\<Delta> \<turnstile> M<:N"}.
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   850
\end{enumerate}
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   851
\end{lemma}
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   852
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   853
The two parts are proved simultaneously, by induction on the size
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   854
of @{term "Q"}.  The argument for part (2) assumes that part (1) has 
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   855
been established already for the @{term "Q"} in question; part (1) uses 
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   856
part (2) only for strictly smaller @{term "Q"}.
18650
urbanc
parents: 18628
diff changeset
   857
\end{quote}
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   858
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   859
For the induction on the size of @{term "Q"}, we use the induction-rule 
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   860
@{text "measure_induct_rule"}:
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   861
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   862
\begin{center}
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   863
@{thm measure_induct_rule[of "size_ty",no_vars]}
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   864
\end{center}
18410
73bb08d2823c made further tunings
urbanc
parents: 18353
diff changeset
   865
18628
urbanc
parents: 18621
diff changeset
   866
That means in order to show a property @{term "P a"} for all @{term "a"}, 
18650
urbanc
parents: 18628
diff changeset
   867
the induct-rule requires to prove that for all @{term x} @{term "P x"} holds using the 
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   868
assumption that for all @{term y} whose size is strictly smaller than 
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   869
that of @{term x} the property @{term "P y"} holds. *}
18353
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   870
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   871
lemma 
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   872
  shows subtype_transitivity: "\<Gamma>\<turnstile>S<:Q \<Longrightarrow> \<Gamma>\<turnstile>Q<:T \<Longrightarrow> \<Gamma>\<turnstile>S<:T" 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   873
  and subtype_narrow: "(\<Delta>@[(TVarB X Q)]@\<Gamma>)\<turnstile>M<:N \<Longrightarrow> \<Gamma>\<turnstile>P<:Q \<Longrightarrow> (\<Delta>@[(TVarB X P)]@\<Gamma>)\<turnstile>M<:N"
20503
503ac4c5ef91 induct method: renamed 'fixing' to 'arbitrary';
wenzelm
parents: 20399
diff changeset
   874
proof (induct Q arbitrary: \<Gamma> S T \<Delta> X P M N taking: "size_ty" rule: measure_induct_rule)
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   875
  case (less Q)
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   876
  have IH_trans:  
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   877
    "\<And>Q' \<Gamma> S T. \<lbrakk>size_ty Q' < size_ty Q; \<Gamma>\<turnstile>S<:Q'; \<Gamma>\<turnstile>Q'<:T\<rbrakk> \<Longrightarrow> \<Gamma>\<turnstile>S<:T" by fact
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   878
  have IH_narrow:
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   879
    "\<And>Q' \<Delta> \<Gamma> X M N P. \<lbrakk>size_ty Q' < size_ty Q; (\<Delta>@[(TVarB X Q')]@\<Gamma>)\<turnstile>M<:N; \<Gamma>\<turnstile>P<:Q'\<rbrakk> 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   880
    \<Longrightarrow> (\<Delta>@[(TVarB X P)]@\<Gamma>)\<turnstile>M<:N" by fact
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   881
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   882
  { fix \<Gamma> S T
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   883
    have "\<lbrakk>\<Gamma> \<turnstile> S <: Q; \<Gamma> \<turnstile> Q <: T\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> S <: T" 
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   884
    proof (induct \<Gamma> S Q\<equiv>Q rule: subtype_of.induct) 
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   885
      case (SA_Top \<Gamma> S) 
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   886
      then have rh_drv: "\<Gamma> \<turnstile> Top <: T" by simp
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   887
      then have T_inst: "T = Top" by (auto elim: S_TopE)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   888
      from `\<turnstile> \<Gamma> ok` and `S closed_in \<Gamma>`
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   889
      have "\<Gamma> \<turnstile> S <: Top" by auto
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   890
      then show "\<Gamma> \<turnstile> S <: T" using T_inst by simp 
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   891
    next
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   892
      case (SA_trans_TVar Y U \<Gamma>) 
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   893
      then have IH_inner: "\<Gamma> \<turnstile> U <: T" by simp
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   894
      have "(TVarB Y U) \<in> set \<Gamma>" by fact
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   895
      with IH_inner show "\<Gamma> \<turnstile> Tvar Y <: T" by auto
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   896
    next
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   897
      case (SA_refl_TVar \<Gamma> X) 
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   898
      then show "\<Gamma> \<turnstile> Tvar X <: T" by simp
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   899
    next
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   900
      case (SA_arrow \<Gamma> Q\<^isub>1 S\<^isub>1 S\<^isub>2 Q\<^isub>2) 
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   901
      then have rh_drv: "\<Gamma> \<turnstile> Q\<^isub>1 \<rightarrow> Q\<^isub>2 <: T" by simp
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   902
      from `Q\<^isub>1 \<rightarrow> Q\<^isub>2 = Q` 
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
   903
      have Q\<^isub>12_less: "size_ty Q\<^isub>1 < size_ty Q" "size_ty Q\<^isub>2 < size_ty Q" by auto
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   904
      have lh_drv_prm\<^isub>1: "\<Gamma> \<turnstile> Q\<^isub>1 <: S\<^isub>1" by fact
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   905
      have lh_drv_prm\<^isub>2: "\<Gamma> \<turnstile> S\<^isub>2 <: Q\<^isub>2" by fact      
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   906
      from rh_drv have "T=Top \<or> (\<exists>T\<^isub>1 T\<^isub>2. T=T\<^isub>1\<rightarrow>T\<^isub>2 \<and> \<Gamma>\<turnstile>T\<^isub>1<:Q\<^isub>1 \<and> \<Gamma>\<turnstile>Q\<^isub>2<:T\<^isub>2)" 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   907
        by (auto elim: S_ArrowE_left)
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   908
      moreover
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   909
      have "S\<^isub>1 closed_in \<Gamma>" and "S\<^isub>2 closed_in \<Gamma>" 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   910
        using lh_drv_prm\<^isub>1 lh_drv_prm\<^isub>2 by (simp_all add: subtype_implies_closed)
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   911
      hence "(S\<^isub>1 \<rightarrow> S\<^isub>2) closed_in \<Gamma>" by (simp add: closed_in_def ty.supp)
18353
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   912
      moreover
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   913
      have "\<turnstile> \<Gamma> ok" using rh_drv by (rule subtype_implies_ok)
18353
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   914
      moreover
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   915
      { assume "\<exists>T\<^isub>1 T\<^isub>2. T = T\<^isub>1\<rightarrow>T\<^isub>2 \<and> \<Gamma> \<turnstile> T\<^isub>1 <: Q\<^isub>1 \<and> \<Gamma> \<turnstile> Q\<^isub>2 <: T\<^isub>2"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   916
        then obtain T\<^isub>1 T\<^isub>2 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   917
          where T_inst: "T = T\<^isub>1 \<rightarrow> T\<^isub>2" 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   918
          and   rh_drv_prm\<^isub>1: "\<Gamma> \<turnstile> T\<^isub>1 <: Q\<^isub>1"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   919
          and   rh_drv_prm\<^isub>2: "\<Gamma> \<turnstile> Q\<^isub>2 <: T\<^isub>2" by force
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   920
        from IH_trans[of "Q\<^isub>1"] 
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
   921
        have "\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1" using Q\<^isub>12_less rh_drv_prm\<^isub>1 lh_drv_prm\<^isub>1 by simp 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   922
        moreover
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   923
        from IH_trans[of "Q\<^isub>2"] 
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
   924
        have "\<Gamma> \<turnstile> S\<^isub>2 <: T\<^isub>2" using Q\<^isub>12_less rh_drv_prm\<^isub>2 lh_drv_prm\<^isub>2 by simp
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   925
        ultimately have "\<Gamma> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T\<^isub>1 \<rightarrow> T\<^isub>2" by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   926
        then have "\<Gamma> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T" using T_inst by simp
18353
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   927
      }
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   928
      ultimately show "\<Gamma> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T" by blast
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   929
    next
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   930
      case (SA_all \<Gamma> Q\<^isub>1 S\<^isub>1 X S\<^isub>2 Q\<^isub>2) 
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   931
      then have rh_drv: "\<Gamma> \<turnstile> (\<forall>X<:Q\<^isub>1. Q\<^isub>2) <: T" by simp
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   932
      have lh_drv_prm\<^isub>1: "\<Gamma> \<turnstile> Q\<^isub>1 <: S\<^isub>1" by fact
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   933
      have lh_drv_prm\<^isub>2: "((TVarB X Q\<^isub>1)#\<Gamma>) \<turnstile> S\<^isub>2 <: Q\<^isub>2" by fact
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   934
      then have "X\<sharp>\<Gamma>" by (force dest: subtype_implies_ok simp add: valid_ty_dom_fresh)
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   935
      then have fresh_cond: "X\<sharp>\<Gamma>" "X\<sharp>Q\<^isub>1" "X\<sharp>T" using rh_drv lh_drv_prm\<^isub>1 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   936
        by (simp_all add: subtype_implies_fresh)
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   937
      from rh_drv 
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   938
      have "T = Top \<or> 
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   939
          (\<exists>T\<^isub>1 T\<^isub>2. T = (\<forall>X<:T\<^isub>1. T\<^isub>2) \<and> \<Gamma> \<turnstile> T\<^isub>1 <: Q\<^isub>1 \<and> ((TVarB X T\<^isub>1)#\<Gamma>) \<turnstile> Q\<^isub>2 <: T\<^isub>2)" 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   940
        using fresh_cond by (simp add: S_ForallE_left)
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   941
      moreover
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   942
      have "S\<^isub>1 closed_in \<Gamma>" and "S\<^isub>2 closed_in ((TVarB X Q\<^isub>1)#\<Gamma>)" 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   943
        using lh_drv_prm\<^isub>1 lh_drv_prm\<^isub>2 by (simp_all add: subtype_implies_closed)
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   944
      then have "(\<forall>X<:S\<^isub>1. S\<^isub>2) closed_in \<Gamma>" by (force simp add: closed_in_def ty.supp abs_supp)
18353
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   945
      moreover
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
   946
      have "\<turnstile> \<Gamma> ok" using rh_drv by (rule subtype_implies_ok)
18353
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   947
      moreover
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   948
      { assume "\<exists>T\<^isub>1 T\<^isub>2. T=(\<forall>X<:T\<^isub>1. T\<^isub>2) \<and> \<Gamma>\<turnstile>T\<^isub>1<:Q\<^isub>1 \<and> ((TVarB X T\<^isub>1)#\<Gamma>)\<turnstile>Q\<^isub>2<:T\<^isub>2"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   949
        then obtain T\<^isub>1 T\<^isub>2 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   950
          where T_inst: "T = (\<forall>X<:T\<^isub>1. T\<^isub>2)" 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   951
          and   rh_drv_prm\<^isub>1: "\<Gamma> \<turnstile> T\<^isub>1 <: Q\<^isub>1" 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   952
          and   rh_drv_prm\<^isub>2:"((TVarB X T\<^isub>1)#\<Gamma>) \<turnstile> Q\<^isub>2 <: T\<^isub>2" by force
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   953
        have "(\<forall>X<:Q\<^isub>1. Q\<^isub>2) = Q" by fact 
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
   954
        then have Q\<^isub>12_less: "size_ty Q\<^isub>1 < size_ty Q" "size_ty Q\<^isub>2 < size_ty Q" 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   955
          using fresh_cond by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   956
        from IH_trans[of "Q\<^isub>1"] 
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
   957
        have "\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1" using lh_drv_prm\<^isub>1 rh_drv_prm\<^isub>1 Q\<^isub>12_less by blast
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   958
        moreover
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   959
        from IH_narrow[of "Q\<^isub>1" "[]"] 
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
   960
        have "((TVarB X T\<^isub>1)#\<Gamma>) \<turnstile> S\<^isub>2 <: Q\<^isub>2" using Q\<^isub>12_less lh_drv_prm\<^isub>2 rh_drv_prm\<^isub>1 by simp
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   961
        with IH_trans[of "Q\<^isub>2"] 
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
   962
        have "((TVarB X T\<^isub>1)#\<Gamma>) \<turnstile> S\<^isub>2 <: T\<^isub>2" using Q\<^isub>12_less rh_drv_prm\<^isub>2 by simp
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   963
        ultimately have "\<Gamma> \<turnstile> (\<forall>X<:S\<^isub>1. S\<^isub>2) <: (\<forall>X<:T\<^isub>1. T\<^isub>2)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   964
          using fresh_cond by (simp add: subtype_of.SA_all)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   965
        hence "\<Gamma> \<turnstile> (\<forall>X<:S\<^isub>1. S\<^isub>2) <: T" using T_inst by simp
18353
4dd468ccfdf7 transitivity should be now in a reasonable state. But
urbanc
parents: 18306
diff changeset
   966
      }
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   967
      ultimately show "\<Gamma> \<turnstile> (\<forall>X<:S\<^isub>1. S\<^isub>2) <: T" by blast
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   968
    qed
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   969
  } note transitivity_lemma = this  
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   970
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   971
  { --{* The transitivity proof is now by the auxiliary lemma. *}
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   972
    case 1 
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   973
    from `\<Gamma> \<turnstile> S <: Q` and `\<Gamma> \<turnstile> Q <: T`
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   974
    show "\<Gamma> \<turnstile> S <: T" by (rule transitivity_lemma) 
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   975
  next 
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   976
    case 2
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   977
    from `(\<Delta>@[(TVarB X Q)]@\<Gamma>) \<turnstile> M <: N` 
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   978
      and `\<Gamma> \<turnstile> P<:Q` 
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   979
    show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> M <: N" 
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
   980
    proof (induct "\<Delta>@[(TVarB X Q)]@\<Gamma>" M N arbitrary: \<Gamma> X \<Delta> rule: subtype_of.induct) 
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
   981
      case (SA_Top S \<Gamma> X \<Delta>)
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
   982
      from `\<Gamma> \<turnstile> P <: Q`
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
   983
      have "P closed_in \<Gamma>" by (simp add: subtype_implies_closed)
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
   984
      with `\<turnstile> (\<Delta>@[(TVarB X Q)]@\<Gamma>) ok` have "\<turnstile> (\<Delta>@[(TVarB X P)]@\<Gamma>) ok"
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
   985
        by (simp add: replace_type)
18412
9f6b3e1da352 tuned the proof of transitivity/narrowing
urbanc
parents: 18410
diff changeset
   986
      moreover
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
   987
      from `S closed_in (\<Delta>@[(TVarB X Q)]@\<Gamma>)` have "S closed_in (\<Delta>@[(TVarB X P)]@\<Gamma>)" 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   988
        by (simp add: closed_in_def doms_append)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   989
      ultimately show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> S <: Top" by (simp add: subtype_of.SA_Top)
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   990
    next
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
   991
      case (SA_trans_TVar Y S N \<Gamma> X \<Delta>) 
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   992
      then have IH_inner: "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> S <: N"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   993
        and lh_drv_prm: "(TVarB Y S) \<in> set (\<Delta>@[(TVarB X Q)]@\<Gamma>)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   994
        and rh_drv: "\<Gamma> \<turnstile> P<:Q"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   995
        and ok\<^isub>Q: "\<turnstile> (\<Delta>@[(TVarB X Q)]@\<Gamma>) ok" by (simp_all add: subtype_implies_ok)
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
   996
      then have ok\<^isub>P: "\<turnstile> (\<Delta>@[(TVarB X P)]@\<Gamma>) ok" by (simp add: subtype_implies_ok) 
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
   997
      show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Tvar Y <: N"
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
   998
      proof (cases "X=Y")
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
   999
        case False
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
  1000
        have "X\<noteq>Y" by fact
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
  1001
        hence "(TVarB Y S)\<in>set (\<Delta>@[(TVarB X P)]@\<Gamma>)" using lh_drv_prm by (simp add:binding.inject)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
  1002
        with IH_inner show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Tvar Y <: N" by (simp add: subtype_of.SA_trans_TVar)
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
  1003
      next
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
  1004
        case True
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1005
        have memb\<^isub>XQ: "(TVarB X Q)\<in>set (\<Delta>@[(TVarB X Q)]@\<Gamma>)" by simp
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1006
        have memb\<^isub>XP: "(TVarB X P)\<in>set (\<Delta>@[(TVarB X P)]@\<Gamma>)" by simp
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
  1007
        have eq: "X=Y" by fact 
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1008
        hence "S=Q" using ok\<^isub>Q lh_drv_prm memb\<^isub>XQ by (simp only: uniqueness_of_ctxt)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
  1009
        hence "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Q <: N" using IH_inner by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
  1010
        moreover
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
  1011
        have "(\<Delta>@[(TVarB X P)]@\<Gamma>) extends \<Gamma>" by (simp add: extends_def)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
  1012
        hence "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> P <: Q" using rh_drv ok\<^isub>P by (simp only: weakening)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
  1013
        ultimately have "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> P <: N" by (simp add: transitivity_lemma) 
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1014
        then show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Tvar Y <: N" using memb\<^isub>XP eq by auto
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
  1015
      qed
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
  1016
    next
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1017
      case (SA_refl_TVar Y \<Gamma> X \<Delta>)
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1018
      from `\<Gamma> \<turnstile> P <: Q`
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1019
      have "P closed_in \<Gamma>" by (simp add: subtype_implies_closed)
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1020
      with `\<turnstile> (\<Delta>@[(TVarB X Q)]@\<Gamma>) ok` have "\<turnstile> (\<Delta>@[(TVarB X P)]@\<Gamma>) ok"
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1021
        by (simp add: replace_type)
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
  1022
      moreover
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1023
      from `Y \<in> ty_dom (\<Delta>@[(TVarB X Q)]@\<Gamma>)` have "Y \<in> ty_dom (\<Delta>@[(TVarB X P)]@\<Gamma>)"
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1024
        by (simp add: doms_append)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1025
      ultimately show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Tvar Y <: Tvar Y" by (simp add: subtype_of.SA_refl_TVar)
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
  1026
    next
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1027
      case (SA_arrow S\<^isub>1 Q\<^isub>1 Q\<^isub>2 S\<^isub>2 \<Gamma> X \<Delta>) 
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1028
      then show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Q\<^isub>1 \<rightarrow> Q\<^isub>2 <: S\<^isub>1 \<rightarrow> S\<^isub>2" by blast 
18424
a37f06555c07 tuned more proofs
urbanc
parents: 18417
diff changeset
  1029
    next
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1030
      case (SA_all T\<^isub>1 S\<^isub>1 Y S\<^isub>2 T\<^isub>2 \<Gamma> X \<Delta>)
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1031
      have IH_inner\<^isub>1: "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> T\<^isub>1 <: S\<^isub>1" 
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1032
        and IH_inner\<^isub>2: "(((TVarB Y T\<^isub>1)#\<Delta>)@[(TVarB X P)]@\<Gamma>) \<turnstile> S\<^isub>2 <: T\<^isub>2"
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 41798
diff changeset
  1033
        by (fastforce intro: SA_all)+
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1034
      then show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> (\<forall>Y<:S\<^isub>1. S\<^isub>2) <: (\<forall>Y<:T\<^isub>1. T\<^isub>2)" by auto
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
  1035
    qed
18621
4a3806b41d29 commented the transitivity and narrowing proof
urbanc
parents: 18577
diff changeset
  1036
  } 
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
  1037
qed
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
  1038
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1039
section {* Typing *}
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1040
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1041
inductive
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1042
  typing :: "env \<Rightarrow> trm \<Rightarrow> ty \<Rightarrow> bool" ("_ \<turnstile> _ : _" [60,60,60] 60) 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1043
where
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1044
  T_Var[intro]: "\<lbrakk> VarB x T \<in> set \<Gamma>; \<turnstile> \<Gamma> ok \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1045
| T_App[intro]: "\<lbrakk> \<Gamma> \<turnstile> t\<^isub>1 : T\<^isub>1 \<rightarrow> T\<^isub>2; \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>1 \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> t\<^isub>1 \<cdot> t\<^isub>2 : T\<^isub>2"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1046
| T_Abs[intro]: "\<lbrakk> VarB x T\<^isub>1 # \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2 \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> (\<lambda>x:T\<^isub>1. t\<^isub>2) : T\<^isub>1 \<rightarrow> T\<^isub>2"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1047
| T_Sub[intro]: "\<lbrakk> \<Gamma> \<turnstile> t : S; \<Gamma> \<turnstile> S <: T \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> t : T"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1048
| T_TAbs[intro]:"\<lbrakk> TVarB X T\<^isub>1 # \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2 \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> (\<lambda>X<:T\<^isub>1. t\<^isub>2) : (\<forall>X<:T\<^isub>1. T\<^isub>2)"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1049
| T_TApp[intro]:"\<lbrakk>X\<sharp>(\<Gamma>,t\<^isub>1,T\<^isub>2); \<Gamma> \<turnstile> t\<^isub>1 : (\<forall>X<:T\<^isub>11. T\<^isub>12); \<Gamma> \<turnstile> T\<^isub>2 <: T\<^isub>11\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> t\<^isub>1 \<cdot>\<^sub>\<tau> T\<^isub>2 : (T\<^isub>12[X \<mapsto> T\<^isub>2]\<^sub>\<tau>)" 
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1050
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1051
equivariance typing
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1052
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1053
lemma better_T_TApp:
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1054
  assumes H1: "\<Gamma> \<turnstile> t\<^isub>1 : (\<forall>X<:T11. T12)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1055
  and H2: "\<Gamma> \<turnstile> T2 <: T11"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1056
  shows "\<Gamma> \<turnstile> t\<^isub>1 \<cdot>\<^sub>\<tau> T2 : (T12[X \<mapsto> T2]\<^sub>\<tau>)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1057
proof -
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1058
  obtain Y::tyvrs where Y: "Y \<sharp> (X, T12, \<Gamma>, t\<^isub>1, T2)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1059
    by (rule exists_fresh) (rule fin_supp)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1060
  then have "Y \<sharp> (\<Gamma>, t\<^isub>1, T2)" by simp
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1061
  moreover from Y have "(\<forall>X<:T11. T12) = (\<forall>Y<:T11. [(Y, X)] \<bullet> T12)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1062
    by (auto simp add: ty.inject alpha' fresh_prod fresh_atm)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1063
  with H1 have "\<Gamma> \<turnstile> t\<^isub>1 : (\<forall>Y<:T11. [(Y, X)] \<bullet> T12)" by simp
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1064
  ultimately have "\<Gamma> \<turnstile> t\<^isub>1 \<cdot>\<^sub>\<tau> T2 : (([(Y, X)] \<bullet> T12)[Y \<mapsto> T2]\<^sub>\<tau>)" using H2
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1065
    by (rule T_TApp)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1066
  with Y show ?thesis by (simp add: type_subst_rename)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1067
qed
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1068
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1069
lemma typing_ok:
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1070
  assumes "\<Gamma> \<turnstile> t : T"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1071
  shows   "\<turnstile> \<Gamma> ok"
49171
3d7a695385f1 tuned proofs;
wenzelm
parents: 46182
diff changeset
  1072
using assms by (induct) (auto)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1073
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1074
nominal_inductive typing
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1075
by (auto dest!: typing_ok intro: closed_in_fresh fresh_dom type_subst_fresh
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1076
    simp: abs_fresh fresh_type_subst_fresh ty_vrs_fresh valid_ty_dom_fresh fresh_trm_dom)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1077
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1078
lemma ok_imp_VarB_closed_in:
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1079
  assumes ok: "\<turnstile> \<Gamma> ok"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1080
  shows "VarB x T \<in> set \<Gamma> \<Longrightarrow> T closed_in \<Gamma>" using ok
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1081
  by induct (auto simp add: binding.inject closed_in_def)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1082
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1083
lemma tyvrs_of_subst: "tyvrs_of (B[X \<mapsto> T]\<^sub>b) = tyvrs_of B"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1084
  by (nominal_induct B rule: binding.strong_induct) simp_all
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1085
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1086
lemma ty_dom_subst: "ty_dom (\<Gamma>[X \<mapsto> T]\<^sub>e) = ty_dom \<Gamma>"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1087
  by (induct \<Gamma>) (simp_all add: tyvrs_of_subst)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1088
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1089
lemma vrs_of_subst: "vrs_of (B[X \<mapsto> T]\<^sub>b) = vrs_of B"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1090
  by (nominal_induct B rule: binding.strong_induct) simp_all
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1091
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1092
lemma trm_dom_subst: "trm_dom (\<Gamma>[X \<mapsto> T]\<^sub>e) = trm_dom \<Gamma>"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1093
  by (induct \<Gamma>) (simp_all add: vrs_of_subst)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1094
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1095
lemma subst_closed_in:
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1096
  "T closed_in (\<Delta> @ TVarB X S # \<Gamma>) \<Longrightarrow> U closed_in \<Gamma> \<Longrightarrow> T[X \<mapsto> U]\<^sub>\<tau> closed_in (\<Delta>[X \<mapsto> U]\<^sub>e @ \<Gamma>)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1097
  apply (nominal_induct T avoiding: X U \<Gamma> rule: ty.strong_induct)
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1098
  apply (simp add: closed_in_def ty.supp supp_atm doms_append ty_dom_subst)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1099
  apply blast
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1100
  apply (simp add: closed_in_def ty.supp)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1101
  apply (simp add: closed_in_def ty.supp)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1102
  apply (simp add: closed_in_def ty.supp abs_supp)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1103
  apply (drule_tac x = X in meta_spec)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1104
  apply (drule_tac x = U in meta_spec)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1105
  apply (drule_tac x = "(TVarB tyvrs ty2) # \<Gamma>" in meta_spec)
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1106
  apply (simp add: doms_append ty_dom_subst)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1107
  apply blast
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1108
  done
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1109
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1110
lemmas subst_closed_in' = subst_closed_in [where \<Delta>="[]", simplified]
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1111
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1112
lemma typing_closed_in:
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1113
  assumes "\<Gamma> \<turnstile> t : T"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1114
  shows   "T closed_in \<Gamma>"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1115
using assms
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1116
proof induct
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1117
  case (T_Var x T \<Gamma>)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1118
  from `\<turnstile> \<Gamma> ok` and `VarB x T \<in> set \<Gamma>`
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1119
  show ?case by (rule ok_imp_VarB_closed_in)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1120
next
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1121
  case (T_App \<Gamma> t\<^isub>1 T\<^isub>1 T\<^isub>2 t\<^isub>2)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1122
  then show ?case by (auto simp add: ty.supp closed_in_def)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1123
next
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1124
  case (T_Abs x T\<^isub>1 \<Gamma> t\<^isub>2 T\<^isub>2)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1125
  from `VarB x T\<^isub>1 # \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2`
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1126
  have "T\<^isub>1 closed_in \<Gamma>" by (auto dest: typing_ok)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1127
  with T_Abs show ?case by (auto simp add: ty.supp closed_in_def)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1128
next
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1129
  case (T_Sub \<Gamma> t S T)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1130
  from `\<Gamma> \<turnstile> S <: T` show ?case by (simp add: subtype_implies_closed)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1131
next
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1132
  case (T_TAbs X T\<^isub>1 \<Gamma> t\<^isub>2 T\<^isub>2)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1133
  from `TVarB X T\<^isub>1 # \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2`
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1134
  have "T\<^isub>1 closed_in \<Gamma>" by (auto dest: typing_ok)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1135
  with T_TAbs show ?case by (auto simp add: ty.supp closed_in_def abs_supp)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1136
next
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1137
  case (T_TApp X \<Gamma> t\<^isub>1 T2 T11 T12)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1138
  then have "T12 closed_in (TVarB X T11 # \<Gamma>)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1139
    by (auto simp add: closed_in_def ty.supp abs_supp)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1140
  moreover from T_TApp have "T2 closed_in \<Gamma>"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1141
    by (simp add: subtype_implies_closed)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1142
  ultimately show ?case by (rule subst_closed_in')
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1143
qed
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1144
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1145
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1146
subsection {* Evaluation *}
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1147
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1148
inductive
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1149
  val :: "trm \<Rightarrow> bool"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1150
where
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1151
  Abs[intro]:  "val (\<lambda>x:T. t)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1152
| TAbs[intro]: "val (\<lambda>X<:T. t)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1153
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1154
equivariance val
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1155
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1156
inductive_cases val_inv_auto[elim]: 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1157
  "val (Var x)" 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1158
  "val (t1 \<cdot> t2)" 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1159
  "val (t1 \<cdot>\<^sub>\<tau> t2)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1160
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1161
inductive 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1162
  eval :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<longmapsto> _" [60,60] 60)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1163
where
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1164
  E_Abs         : "\<lbrakk> x \<sharp> v\<^isub>2; val v\<^isub>2 \<rbrakk> \<Longrightarrow> (\<lambda>x:T\<^isub>11. t\<^isub>12) \<cdot> v\<^isub>2 \<longmapsto> t\<^isub>12[x \<mapsto> v\<^isub>2]"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1165
| E_App1 [intro]: "t \<longmapsto> t' \<Longrightarrow> t \<cdot> u \<longmapsto> t' \<cdot> u"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1166
| E_App2 [intro]: "\<lbrakk> val v; t \<longmapsto> t' \<rbrakk> \<Longrightarrow> v \<cdot> t \<longmapsto> v \<cdot> t'"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1167
| E_TAbs        : "X \<sharp> (T\<^isub>11, T\<^isub>2) \<Longrightarrow> (\<lambda>X<:T\<^isub>11. t\<^isub>12) \<cdot>\<^sub>\<tau> T\<^isub>2 \<longmapsto> t\<^isub>12[X \<mapsto>\<^sub>\<tau> T\<^isub>2]"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1168
| E_TApp [intro]: "t \<longmapsto> t' \<Longrightarrow> t \<cdot>\<^sub>\<tau> T \<longmapsto> t' \<cdot>\<^sub>\<tau> T"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1169
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1170
lemma better_E_Abs[intro]:
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1171
  assumes H: "val v2"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1172
  shows "(\<lambda>x:T11. t12) \<cdot> v2 \<longmapsto> t12[x \<mapsto> v2]"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1173
proof -
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1174
  obtain y::vrs where y: "y \<sharp> (x, t12, v2)" by (rule exists_fresh) (rule fin_supp)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1175
  then have "y \<sharp> v2" by simp
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1176
  then have "(\<lambda>y:T11. [(y, x)] \<bullet> t12) \<cdot> v2 \<longmapsto> ([(y, x)] \<bullet> t12)[y \<mapsto> v2]" using H
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1177
    by (rule E_Abs)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1178
  moreover from y have "(\<lambda>x:T11. t12) \<cdot> v2 = (\<lambda>y:T11. [(y, x)] \<bullet> t12) \<cdot> v2"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1179
    by (auto simp add: trm.inject alpha' fresh_prod fresh_atm)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1180
  ultimately have "(\<lambda>x:T11. t12) \<cdot> v2 \<longmapsto> ([(y, x)] \<bullet> t12)[y \<mapsto> v2]"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1181
    by simp
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1182
  with y show ?thesis by (simp add: subst_trm_rename)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1183
qed
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1184
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1185
lemma better_E_TAbs[intro]: "(\<lambda>X<:T11. t12) \<cdot>\<^sub>\<tau> T2 \<longmapsto> t12[X \<mapsto>\<^sub>\<tau> T2]"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1186
proof -
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1187
  obtain Y::tyvrs where Y: "Y \<sharp> (X, t12, T11, T2)" by (rule exists_fresh) (rule fin_supp)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1188
  then have "Y \<sharp> (T11, T2)" by simp
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1189
  then have "(\<lambda>Y<:T11. [(Y, X)] \<bullet> t12) \<cdot>\<^sub>\<tau> T2 \<longmapsto> ([(Y, X)] \<bullet> t12)[Y \<mapsto>\<^sub>\<tau> T2]"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1190
    by (rule E_TAbs)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1191
  moreover from Y have "(\<lambda>X<:T11. t12) \<cdot>\<^sub>\<tau> T2 = (\<lambda>Y<:T11. [(Y, X)] \<bullet> t12) \<cdot>\<^sub>\<tau> T2"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1192
    by (auto simp add: trm.inject alpha' fresh_prod fresh_atm)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1193
  ultimately have "(\<lambda>X<:T11. t12) \<cdot>\<^sub>\<tau> T2 \<longmapsto> ([(Y, X)] \<bullet> t12)[Y \<mapsto>\<^sub>\<tau> T2]"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1194
    by simp
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1195
  with Y show ?thesis by (simp add: subst_trm_ty_rename)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1196
qed
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1197
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1198
equivariance eval
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1199
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1200
nominal_inductive eval
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1201
  by (simp_all add: abs_fresh ty_vrs_fresh subst_trm_fresh_tyvar
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1202
    subst_trm_fresh_var subst_trm_ty_fresh')
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1203
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1204
inductive_cases eval_inv_auto[elim]: 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1205
  "Var x \<longmapsto> t'" 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1206
  "(\<lambda>x:T. t) \<longmapsto> t'" 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1207
  "(\<lambda>X<:T. t) \<longmapsto> t'" 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1208
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1209
lemma ty_dom_cons:
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1210
  shows "ty_dom (\<Gamma>@[VarB X Q]@\<Delta>) = ty_dom (\<Gamma>@\<Delta>)"
49171
3d7a695385f1 tuned proofs;
wenzelm
parents: 46182
diff changeset
  1211
by (induct \<Gamma>) (auto)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1212
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1213
lemma closed_in_cons: 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1214
  assumes "S closed_in (\<Gamma> @ VarB X Q # \<Delta>)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1215
  shows "S closed_in (\<Gamma>@\<Delta>)"
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1216
using assms ty_dom_cons closed_in_def by auto
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1217
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1218
lemma closed_in_weaken: "T closed_in (\<Delta> @ \<Gamma>) \<Longrightarrow> T closed_in (\<Delta> @ B # \<Gamma>)"
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1219
  by (auto simp add: closed_in_def doms_append)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1220
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1221
lemma closed_in_weaken': "T closed_in \<Gamma> \<Longrightarrow> T closed_in (\<Delta> @ \<Gamma>)"
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1222
  by (auto simp add: closed_in_def doms_append)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1223
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1224
lemma valid_subst:
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1225
  assumes ok: "\<turnstile> (\<Delta> @ TVarB X Q # \<Gamma>) ok"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1226
  and closed: "P closed_in \<Gamma>"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1227
  shows "\<turnstile> (\<Delta>[X \<mapsto> P]\<^sub>e @ \<Gamma>) ok" using ok closed
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1228
  apply (induct \<Delta>)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1229
  apply simp_all
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1230
  apply (erule validE)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1231
  apply assumption
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1232
  apply (erule validE)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1233
  apply simp
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1234
  apply (rule valid_consT)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1235
  apply assumption
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1236
  apply (simp add: doms_append ty_dom_subst)
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1237
  apply (simp add: fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] finite_doms)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1238
  apply (rule_tac S=Q in subst_closed_in')
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1239
  apply (simp add: closed_in_def doms_append ty_dom_subst)
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1240
  apply (simp add: closed_in_def doms_append)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1241
  apply blast
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1242
  apply simp
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1243
  apply (rule valid_cons)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1244
  apply assumption
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1245
  apply (simp add: doms_append trm_dom_subst)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1246
  apply (rule_tac S=Q in subst_closed_in')
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1247
  apply (simp add: closed_in_def doms_append ty_dom_subst)
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1248
  apply (simp add: closed_in_def doms_append)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1249
  apply blast
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1250
  done
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1251
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1252
lemma ty_dom_vrs:
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1253
  shows "ty_dom (G @ [VarB x Q] @ D) = ty_dom (G @ D)"
49171
3d7a695385f1 tuned proofs;
wenzelm
parents: 46182
diff changeset
  1254
by (induct G) (auto)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1255
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1256
lemma valid_cons':
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1257
  assumes "\<turnstile> (\<Gamma> @ VarB x Q # \<Delta>) ok"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1258
  shows "\<turnstile> (\<Gamma> @ \<Delta>) ok"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1259
  using assms
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1260
proof (induct "\<Gamma> @ VarB x Q # \<Delta>" arbitrary: \<Gamma> \<Delta>)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1261
  case valid_nil
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1262
  have "[] = \<Gamma> @ VarB x Q # \<Delta>" by fact
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1263
  then have "False" by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1264
  then show ?case by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1265
next
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1266
  case (valid_consT G X T)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1267
  then show ?case
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1268
  proof (cases \<Gamma>)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1269
    case Nil
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1270
    with valid_consT show ?thesis by simp
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1271
  next
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1272
    case (Cons b bs)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1273
    with valid_consT
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1274
    have "\<turnstile> (bs @ \<Delta>) ok" by simp
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1275
    moreover from Cons and valid_consT have "X \<sharp> ty_dom (bs @ \<Delta>)"
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1276
      by (simp add: doms_append)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1277
    moreover from Cons and valid_consT have "T closed_in (bs @ \<Delta>)"
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1278
      by (simp add: closed_in_def doms_append)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1279
    ultimately have "\<turnstile> (TVarB X T # bs @ \<Delta>) ok"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1280
      by (rule valid_rel.valid_consT)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1281
    with Cons and valid_consT show ?thesis by simp
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1282
  qed
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1283
next
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1284
  case (valid_cons G x T)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1285
  then show ?case
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1286
  proof (cases \<Gamma>)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1287
    case Nil
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1288
    with valid_cons show ?thesis by simp
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1289
  next
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1290
    case (Cons b bs)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1291
    with valid_cons
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1292
    have "\<turnstile> (bs @ \<Delta>) ok" by simp
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1293
    moreover from Cons and valid_cons have "x \<sharp> trm_dom (bs @ \<Delta>)"
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1294
      by (simp add: doms_append finite_doms
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
  1295
        fresh_fin_insert [OF pt_vrs_inst at_vrs_inst fs_vrs_inst])
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1296
    moreover from Cons and valid_cons have "T closed_in (bs @ \<Delta>)"
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1297
      by (simp add: closed_in_def doms_append)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1298
    ultimately have "\<turnstile> (VarB x T # bs @ \<Delta>) ok"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1299
      by (rule valid_rel.valid_cons)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1300
    with Cons and valid_cons show ?thesis by simp
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1301
  qed
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1302
qed
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1303
  
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1304
text {* A.5(6) *}
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1305
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1306
lemma type_weaken:
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1307
  assumes "(\<Delta>@\<Gamma>) \<turnstile> t : T"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1308
  and     "\<turnstile> (\<Delta> @ B # \<Gamma>) ok"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1309
  shows   "(\<Delta> @ B # \<Gamma>) \<turnstile> t : T"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1310
using assms
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1311
proof(nominal_induct "\<Delta> @ \<Gamma>" t T avoiding: \<Delta> \<Gamma> B rule: typing.strong_induct)
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1312
  case (T_Var x T)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1313
  then show ?case by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1314
next
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1315
  case (T_App X t\<^isub>1 T\<^isub>2 T\<^isub>11 T\<^isub>12)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1316
  then show ?case by force
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1317
next
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1318
  case (T_Abs y T\<^isub>1 t\<^isub>2 T\<^isub>2 \<Delta> \<Gamma>)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1319
  then have "VarB y T\<^isub>1 # \<Delta> @ \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2" by simp
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1320
  then have closed: "T\<^isub>1 closed_in (\<Delta> @ \<Gamma>)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1321
    by (auto dest: typing_ok)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1322
  have "\<turnstile> (VarB y T\<^isub>1 # \<Delta> @ B # \<Gamma>) ok"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1323
    apply (rule valid_cons)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1324
    apply (rule T_Abs)
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1325
    apply (simp add: doms_append
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1326
      fresh_fin_insert [OF pt_vrs_inst at_vrs_inst fs_vrs_inst]
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1327
      fresh_fin_union [OF pt_vrs_inst at_vrs_inst fs_vrs_inst]
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1328
      finite_doms finite_vrs fresh_vrs_of T_Abs fresh_trm_dom)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1329
    apply (rule closed_in_weaken)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1330
    apply (rule closed)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1331
    done
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1332
  then have "\<turnstile> ((VarB y T\<^isub>1 # \<Delta>) @ B # \<Gamma>) ok" by simp
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1333
  with _ have "(VarB y T\<^isub>1 # \<Delta>) @ B # \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2"
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1334
    by (rule T_Abs) simp
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1335
  then have "VarB y T\<^isub>1 # \<Delta> @ B # \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2" by simp
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1336
  then show ?case by (rule typing.T_Abs)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1337
next
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1338
  case (T_Sub t S T \<Delta> \<Gamma>)
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1339
  from refl and `\<turnstile> (\<Delta> @ B # \<Gamma>) ok`
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1340
  have "\<Delta> @ B # \<Gamma> \<turnstile> t : S" by (rule T_Sub)
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1341
  moreover from  `(\<Delta> @ \<Gamma>)\<turnstile>S<:T` and `\<turnstile> (\<Delta> @ B # \<Gamma>) ok`
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1342
  have "(\<Delta> @ B # \<Gamma>)\<turnstile>S<:T"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1343
    by (rule weakening) (simp add: extends_def T_Sub)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1344
  ultimately show ?case by (rule typing.T_Sub)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1345
next
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1346
  case (T_TAbs X T\<^isub>1 t\<^isub>2 T\<^isub>2 \<Delta> \<Gamma>)
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1347
  from `TVarB X T\<^isub>1 # \<Delta> @ \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2`
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1348
  have closed: "T\<^isub>1 closed_in (\<Delta> @ \<Gamma>)"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1349
    by (auto dest: typing_ok)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1350
  have "\<turnstile> (TVarB X T\<^isub>1 # \<Delta> @ B # \<Gamma>) ok"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1351
    apply (rule valid_consT)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1352
    apply (rule T_TAbs)
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1353
    apply (simp add: doms_append
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1354
      fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst]
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1355
      fresh_fin_union [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst]
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1356
      finite_doms finite_vrs tyvrs_fresh T_TAbs fresh_dom)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1357
    apply (rule closed_in_weaken)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1358
    apply (rule closed)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1359
    done
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1360
  then have "\<turnstile> ((TVarB X T\<^isub>1 # \<Delta>) @ B # \<Gamma>) ok" by simp
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1361
  with _ have "(TVarB X T\<^isub>1 # \<Delta>) @ B # \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2"
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1362
    by (rule T_TAbs) simp
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1363
  then have "TVarB X T\<^isub>1 # \<Delta> @ B # \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2" by simp
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1364
  then show ?case by (rule typing.T_TAbs)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1365
next
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1366
  case (T_TApp X t\<^isub>1 T2 T11 T12 \<Delta> \<Gamma>)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1367
  have "\<Delta> @ B # \<Gamma> \<turnstile> t\<^isub>1 : (\<forall>X<:T11. T12)"
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1368
    by (rule T_TApp refl)+
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1369
  moreover from `(\<Delta> @ \<Gamma>)\<turnstile>T2<:T11` and `\<turnstile> (\<Delta> @ B # \<Gamma>) ok`
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1370
  have "(\<Delta> @ B # \<Gamma>)\<turnstile>T2<:T11"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1371
    by (rule weakening) (simp add: extends_def T_TApp)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1372
  ultimately show ?case by (rule better_T_TApp)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1373
qed
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1374
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1375
lemma type_weaken':
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1376
 "\<Gamma> \<turnstile> t : T \<Longrightarrow>  \<turnstile> (\<Delta>@\<Gamma>) ok \<Longrightarrow> (\<Delta>@\<Gamma>) \<turnstile> t : T"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1377
  apply (induct \<Delta>)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1378
  apply simp_all
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1379
  apply (erule validE)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1380
  apply (insert type_weaken [of "[]", simplified])
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1381
  apply simp_all
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1382
  done
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1383
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1384
text {* A.6 *}
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1385
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1386
lemma strengthening:
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1387
  assumes "(\<Gamma> @ VarB x Q # \<Delta>) \<turnstile> S <: T"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1388
  shows  "(\<Gamma>@\<Delta>) \<turnstile> S <: T"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1389
  using assms
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1390
proof (induct "\<Gamma> @ VarB x Q # \<Delta>" S T arbitrary: \<Gamma>)
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1391
  case (SA_Top S)
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1392
  then have "\<turnstile> (\<Gamma> @ \<Delta>) ok" by (auto dest: valid_cons')
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1393
  moreover have "S closed_in (\<Gamma> @ \<Delta>)" using SA_Top by (auto dest: closed_in_cons)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1394
  ultimately show ?case using subtype_of.SA_Top by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1395
next
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1396
  case (SA_refl_TVar X)
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1397
  from `\<turnstile> (\<Gamma> @ VarB x Q # \<Delta>) ok`
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1398
  have h1:"\<turnstile> (\<Gamma> @ \<Delta>) ok" by (auto dest: valid_cons')
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1399
  have "X \<in> ty_dom (\<Gamma> @ VarB x Q # \<Delta>)" using SA_refl_TVar by auto
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1400
  then have h2:"X \<in> ty_dom (\<Gamma> @ \<Delta>)" using ty_dom_vrs by auto
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1401
  show ?case using h1 h2 by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1402
next
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1403
  case (SA_all T1 S1 X S2 T2)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 41798
diff changeset
  1404
  have h1:"((TVarB X T1 # \<Gamma>) @ \<Delta>)\<turnstile>S2<:T2" by (fastforce intro: SA_all)
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1405
  have h2:"(\<Gamma> @ \<Delta>)\<turnstile>T1<:S1" using SA_all by auto
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1406
  then show ?case using h1 h2 by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1407
qed (auto)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1408
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1409
lemma narrow_type: -- {* A.7 *}
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1410
  assumes H: "\<Delta> @ (TVarB X Q) # \<Gamma> \<turnstile> t : T"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1411
  shows "\<Gamma> \<turnstile> P <: Q \<Longrightarrow> \<Delta> @ (TVarB X P) # \<Gamma> \<turnstile> t : T"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1412
  using H
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1413
  proof (nominal_induct "\<Delta> @ (TVarB X Q) # \<Gamma>" t T avoiding: P arbitrary: \<Delta> rule: typing.strong_induct)
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1414
    case (T_Var x T P D)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1415
    then have "VarB x T \<in> set (D @ TVarB X P # \<Gamma>)" 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1416
      and "\<turnstile>  (D @ TVarB X P # \<Gamma>) ok"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1417
      by (auto intro: replace_type dest!: subtype_implies_closed)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1418
    then show ?case by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1419
  next
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1420
    case (T_App t1 T1 T2 t2 P D)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1421
    then show ?case by force
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1422
  next
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1423
    case (T_Abs x T1 t2 T2 P D)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 41798
diff changeset
  1424
    then show ?case by (fastforce dest: typing_ok)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1425
  next
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1426
    case (T_Sub t S T P D)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 41798
diff changeset
  1427
    then show ?case using subtype_narrow by fastforce
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1428
  next
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1429
    case (T_TAbs X' T1 t2 T2 P D)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 41798
diff changeset
  1430
    then show ?case by (fastforce dest: typing_ok)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1431
  next
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1432
    case (T_TApp X' t1 T2 T11 T12 P D)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 41798
diff changeset
  1433
    then have "D @ TVarB X P # \<Gamma> \<turnstile> t1 : Forall X' T12 T11" by fastforce
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1434
    moreover have "(D @ [TVarB X Q] @ \<Gamma>) \<turnstile> T2<:T11" using T_TApp by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1435
    then have "(D @ [TVarB X P] @ \<Gamma>) \<turnstile> T2<:T11" using `\<Gamma>\<turnstile>P<:Q`
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1436
      by (rule subtype_narrow)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1437
    moreover from T_TApp have "X' \<sharp> (D @ TVarB X P # \<Gamma>, t1, T2)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1438
      by (simp add: fresh_list_append fresh_list_cons fresh_prod)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1439
    ultimately show ?case by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1440
qed
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1441
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1442
subsection {* Substitution lemmas *}
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1443
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1444
subsubsection {* Substition Preserves Typing *}
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1445
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1446
theorem subst_type: -- {* A.8 *}
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1447
  assumes H: "(\<Delta> @ (VarB x U) # \<Gamma>) \<turnstile> t : T"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1448
  shows "\<Gamma> \<turnstile> u : U \<Longrightarrow> \<Delta> @ \<Gamma> \<turnstile> t[x \<mapsto> u] : T" using H
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1449
 proof (nominal_induct "\<Delta> @ (VarB x U) # \<Gamma>" t T avoiding: x u arbitrary: \<Delta> rule: typing.strong_induct)
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1450
   case (T_Var y T x u D)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1451
   show ?case
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1452
   proof (cases "x = y")
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1453
     assume eq:"x=y"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1454
     then have "T=U" using T_Var uniqueness_of_ctxt' by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1455
     then show ?case using eq T_Var
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1456
       by (auto intro: type_weaken' dest: valid_cons')
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1457
   next
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1458
     assume "x\<noteq>y"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1459
     then show ?case using T_Var
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1460
       by (auto simp add:binding.inject dest: valid_cons')
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1461
   qed
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1462
 next
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1463
   case (T_App t1 T1 T2 t2 x u D)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1464
   then show ?case by force
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1465
 next
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1466
   case (T_Abs y T1 t2 T2 x u D)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1467
   then show ?case by force
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1468
 next
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1469
   case (T_Sub t S T x u D)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1470
   then have "D @ \<Gamma> \<turnstile> t[x \<mapsto> u] : S" by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1471
   moreover have "(D @ \<Gamma>) \<turnstile> S<:T" using T_Sub by (auto dest: strengthening)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1472
   ultimately show ?case by auto 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1473
 next
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1474
   case (T_TAbs X T1 t2 T2 x u D)
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1475
   from `TVarB X T1 # D @ VarB x U # \<Gamma> \<turnstile> t2 : T2` have "X \<sharp> T1"
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1476
     by (auto simp add: valid_ty_dom_fresh dest: typing_ok intro!: closed_in_fresh)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 41798
diff changeset
  1477
   with `X \<sharp> u` and T_TAbs show ?case by fastforce
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1478
 next
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1479
   case (T_TApp X t1 T2 T11 T12 x u D)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1480
   then have "(D@\<Gamma>) \<turnstile>T2<:T11" using T_TApp by (auto dest: strengthening)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1481
   then show "((D @ \<Gamma>) \<turnstile> ((t1 \<cdot>\<^sub>\<tau> T2)[x \<mapsto> u]) : (T12[X \<mapsto> T2]\<^sub>\<tau>))" using T_TApp
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1482
     by (force simp add: fresh_prod fresh_list_append fresh_list_cons subst_trm_fresh_tyvar)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1483
qed
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1484
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1485
subsubsection {* Type Substitution Preserves Subtyping *}
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1486
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1487
lemma substT_subtype: -- {* A.10 *}
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1488
  assumes H: "(\<Delta> @ ((TVarB X Q) # \<Gamma>)) \<turnstile> S <: T"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1489
  shows "\<Gamma> \<turnstile> P <: Q \<Longrightarrow> (\<Delta>[X \<mapsto> P]\<^sub>e @ \<Gamma>) \<turnstile> S[X \<mapsto> P]\<^sub>\<tau> <: T[X \<mapsto> P]\<^sub>\<tau>" 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1490
  using H
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1491
proof (nominal_induct "\<Delta> @ TVarB X Q # \<Gamma>" S T avoiding: X P arbitrary: \<Delta> rule: subtype_of.strong_induct)
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1492
  case (SA_Top S X P D)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1493
  then have "\<turnstile> (D @ TVarB X Q # \<Gamma>) ok" by simp
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1494
  moreover have closed: "P closed_in \<Gamma>" using SA_Top subtype_implies_closed by auto 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1495
  ultimately have "\<turnstile> (D[X \<mapsto> P]\<^sub>e @ \<Gamma>) ok" by (rule valid_subst)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1496
  moreover from SA_Top have "S closed_in (D @ TVarB X Q # \<Gamma>)" by simp
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1497
  then have "S[X \<mapsto> P]\<^sub>\<tau> closed_in  (D[X \<mapsto> P]\<^sub>e @ \<Gamma>)" using closed by (rule subst_closed_in)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1498
  ultimately show ?case by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1499
next
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1500
  case (SA_trans_TVar Y S T X P D)
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1501
  have h:"(D @ TVarB X Q # \<Gamma>)\<turnstile>S<:T" by fact
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1502
  then have ST: "(D[X \<mapsto> P]\<^sub>e @ \<Gamma>) \<turnstile> S[X \<mapsto> P]\<^sub>\<tau> <: T[X \<mapsto> P]\<^sub>\<tau>" using SA_trans_TVar by auto
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1503
  from h have G_ok: "\<turnstile> (D @ TVarB X Q # \<Gamma>) ok" by (rule subtype_implies_ok)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 49171
diff changeset
  1504
  from G_ok and SA_trans_TVar have X_\<Gamma>_ok: "\<turnstile> (TVarB X Q # \<Gamma>) ok"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1505
    by (auto intro: validE_append)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1506
  show "(D[X \<mapsto> P]\<^sub>e @ \<Gamma>) \<turnstile> Tvar Y[X \<mapsto> P]\<^sub>\<tau><:T[X \<mapsto> P]\<^sub>\<tau>"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1507
  proof (cases "X = Y")
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1508
    assume eq: "X = Y"
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1509
    from eq and SA_trans_TVar have "TVarB Y Q \<in> set (D @ TVarB X Q # \<Gamma>)" by simp
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1510
    with G_ok have QS: "Q = S" using `TVarB Y S \<in> set (D @ TVarB X Q # \<Gamma>)`
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1511
      by (rule uniqueness_of_ctxt)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 49171
diff changeset
  1512
    from X_\<Gamma>_ok have "X \<sharp> ty_dom \<Gamma>" and "Q closed_in \<Gamma>" by auto
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1513
    then have XQ: "X \<sharp> Q" by (rule closed_in_fresh)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1514
    note `\<Gamma>\<turnstile>P<:Q`
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1515
    moreover from ST have "\<turnstile> (D[X \<mapsto> P]\<^sub>e @ \<Gamma>) ok" by (rule subtype_implies_ok)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1516
    moreover have "(D[X \<mapsto> P]\<^sub>e @ \<Gamma>) extends \<Gamma>" by (simp add: extends_def)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1517
    ultimately have "(D[X \<mapsto> P]\<^sub>e @ \<Gamma>) \<turnstile> P<:Q" by (rule weakening)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1518
    with QS have "(D[X \<mapsto> P]\<^sub>e @ \<Gamma>) \<turnstile> P<:S" by simp
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1519
    moreover from XQ and ST and QS have "(D[X \<mapsto> P]\<^sub>e @ \<Gamma>) \<turnstile> S<:T[X \<mapsto> P]\<^sub>\<tau>"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1520
      by (simp add: type_subst_identity)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1521
    ultimately have "(D[X \<mapsto> P]\<^sub>e @ \<Gamma>) \<turnstile> P<:T[X \<mapsto> P]\<^sub>\<tau>"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1522
      by (rule subtype_transitivity)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1523
    with eq show ?case by simp
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1524
  next
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1525
    assume neq: "X \<noteq> Y"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1526
    with SA_trans_TVar have "TVarB Y S \<in> set D \<or> TVarB Y S \<in> set \<Gamma>"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1527
      by (simp add: binding.inject)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1528
    then show ?case
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1529
    proof
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1530
      assume "TVarB Y S \<in> set D"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1531
      then have "TVarB Y (S[X \<mapsto> P]\<^sub>\<tau>) \<in> set (D[X \<mapsto> P]\<^sub>e)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
  1532
        by (rule ctxt_subst_mem_TVarB)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1533
      then have "TVarB Y (S[X \<mapsto> P]\<^sub>\<tau>) \<in> set (D[X \<mapsto> P]\<^sub>e @ \<Gamma>)" by simp
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1534
      with neq and ST show ?thesis by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1535
    next
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1536
      assume Y: "TVarB Y S \<in> set \<Gamma>"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 49171
diff changeset
  1537
      from X_\<Gamma>_ok have "X \<sharp> ty_dom \<Gamma>" and "\<turnstile> \<Gamma> ok" by auto
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1538
      then have "X \<sharp> \<Gamma>" by (simp add: valid_ty_dom_fresh)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1539
      with Y have "X \<sharp> S"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
  1540
        by (induct \<Gamma>) (auto simp add: fresh_list_nil fresh_list_cons)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1541
      with ST have "(D[X \<mapsto> P]\<^sub>e @ \<Gamma>)\<turnstile>S<:T[X \<mapsto> P]\<^sub>\<tau>"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32011
diff changeset
  1542
        by (simp add: type_subst_identity)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1543
      moreover from Y have "TVarB Y S \<in> set (D[X \<mapsto> P]\<^sub>e @ \<Gamma>)" by simp
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1544
      ultimately show ?thesis using neq by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1545
    qed
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1546
  qed
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1547
next
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1548
  case (SA_refl_TVar Y X P D)
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1549
  note `\<turnstile> (D @ TVarB X Q # \<Gamma>) ok`
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1550
  moreover from SA_refl_TVar have closed: "P closed_in \<Gamma>"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1551
    by (auto dest: subtype_implies_closed)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1552
  ultimately have ok: "\<turnstile> (D[X \<mapsto> P]\<^sub>e @ \<Gamma>) ok" using valid_subst by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1553
  from closed have closed': "P closed_in (D[X \<mapsto> P]\<^sub>e @ \<Gamma>)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1554
    by (simp add: closed_in_weaken')
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1555
  show ?case
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1556
  proof (cases "X = Y")
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1557
    assume "X = Y"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1558
    with closed' and ok show ?thesis
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1559
      by (auto intro: subtype_reflexivity)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1560
  next
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1561
    assume neq: "X \<noteq> Y"
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1562
    with SA_refl_TVar have "Y \<in> ty_dom (D[X \<mapsto> P]\<^sub>e @ \<Gamma>)"
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1563
      by (simp add: ty_dom_subst doms_append)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1564
    with neq and ok show ?thesis by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1565
  qed
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1566
next
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1567
  case (SA_arrow T1 S1 S2 T2 X P D)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1568
  then have h1:"(D[X \<mapsto> P]\<^sub>e @ \<Gamma>)\<turnstile>T1[X \<mapsto> P]\<^sub>\<tau><:S1[X \<mapsto> P]\<^sub>\<tau>" using SA_arrow by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1569
  from SA_arrow have h2:"(D[X \<mapsto> P]\<^sub>e @ \<Gamma>)\<turnstile>S2[X \<mapsto> P]\<^sub>\<tau><:T2[X \<mapsto> P]\<^sub>\<tau>" using SA_arrow by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1570
  show ?case using subtype_of.SA_arrow h1 h2 by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1571
next
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1572
  case (SA_all T1 S1 Y S2 T2 X P D)
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1573
  then have Y: "Y \<sharp> ty_dom (D @ TVarB X Q # \<Gamma>)"
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1574
    by (auto dest: subtype_implies_ok intro: fresh_dom)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1575
  moreover from SA_all have "S1 closed_in (D @ TVarB X Q # \<Gamma>)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1576
    by (auto dest: subtype_implies_closed)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1577
  ultimately have S1: "Y \<sharp> S1" by (rule closed_in_fresh)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1578
  from SA_all have "T1 closed_in (D @ TVarB X Q # \<Gamma>)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1579
    by (auto dest: subtype_implies_closed)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1580
  with Y have T1: "Y \<sharp> T1" by (rule closed_in_fresh)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1581
  with SA_all and S1 show ?case by force
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1582
qed
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1583
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1584
subsubsection {* Type Substitution Preserves Typing *}
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1585
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1586
theorem substT_type: -- {* A.11 *}
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1587
  assumes H: "(D @ TVarB X Q # G) \<turnstile> t : T"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1588
  shows "G \<turnstile> P <: Q \<Longrightarrow>
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1589
    (D[X \<mapsto> P]\<^sub>e @ G) \<turnstile> t[X \<mapsto>\<^sub>\<tau> P] : T[X \<mapsto> P]\<^sub>\<tau>" using H
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1590
proof (nominal_induct "D @ TVarB X Q # G" t T avoiding: X P arbitrary: D rule: typing.strong_induct)
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1591
  case (T_Var x T X P D')
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1592
  have "G\<turnstile>P<:Q" by fact
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1593
  then have "P closed_in G" using subtype_implies_closed by auto
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1594
  moreover note `\<turnstile> (D' @ TVarB X Q # G) ok`
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1595
  ultimately have "\<turnstile> (D'[X \<mapsto> P]\<^sub>e @ G) ok" using valid_subst by auto
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1596
  moreover note `VarB x T \<in> set (D' @ TVarB X Q # G)`
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1597
  then have "VarB x T \<in> set D' \<or> VarB x T \<in> set G" by simp
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1598
  then have "(VarB x (T[X \<mapsto> P]\<^sub>\<tau>)) \<in> set (D'[X \<mapsto> P]\<^sub>e @ G)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1599
  proof
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1600
    assume "VarB x T \<in> set D'"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1601
    then have "VarB x (T[X \<mapsto> P]\<^sub>\<tau>) \<in> set (D'[X \<mapsto> P]\<^sub>e)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1602
      by (rule ctxt_subst_mem_VarB)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1603
    then show ?thesis by simp
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1604
  next
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1605
    assume x: "VarB x T \<in> set G"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1606
    from T_Var have ok: "\<turnstile> G ok" by (auto dest: subtype_implies_ok)
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1607
    then have "X \<sharp> ty_dom G" using T_Var by (auto dest: validE_append)
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1608
    with ok have "X \<sharp> G" by (simp add: valid_ty_dom_fresh)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1609
    moreover from x have "VarB x T \<in> set (D' @ G)" by simp
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1610
    then have "VarB x (T[X \<mapsto> P]\<^sub>\<tau>) \<in> set ((D' @ G)[X \<mapsto> P]\<^sub>e)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1611
      by (rule ctxt_subst_mem_VarB)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1612
    ultimately show ?thesis
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1613
      by (simp add: ctxt_subst_append ctxt_subst_identity)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1614
  qed
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1615
  ultimately show ?case by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1616
next
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1617
  case (T_App t1 T1 T2 t2 X P D')
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1618
  then have "D'[X \<mapsto> P]\<^sub>e @ G \<turnstile> t1[X \<mapsto>\<^sub>\<tau> P] : (T1 \<rightarrow> T2)[X \<mapsto> P]\<^sub>\<tau>" by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1619
  moreover from T_App have "D'[X \<mapsto> P]\<^sub>e @ G \<turnstile> t2[X \<mapsto>\<^sub>\<tau> P] : T1[X \<mapsto> P]\<^sub>\<tau>" by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1620
  ultimately show ?case by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1621
next
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1622
  case (T_Abs x T1 t2 T2 X P D')
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1623
  then show ?case by force
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1624
next
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1625
  case (T_Sub t S T X P D')
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1626
  then show ?case using substT_subtype by force
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1627
next
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1628
  case (T_TAbs X' T1 t2 T2 X P D')
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1629
  then have "X' \<sharp> ty_dom (D' @ TVarB X Q # G)"
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1630
  and "T1 closed_in (D' @ TVarB X Q # G)"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1631
    by (auto dest: typing_ok)
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1632
  then have "X' \<sharp> T1" by (rule closed_in_fresh)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1633
  with T_TAbs show ?case by force
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1634
next
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1635
  case (T_TApp X' t1 T2 T11 T12 X P D')
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1636
  then have "X' \<sharp> ty_dom (D' @ TVarB X Q # G)"
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1637
    by (simp add: fresh_dom)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1638
  moreover from T_TApp have "T11 closed_in (D' @ TVarB X Q # G)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1639
    by (auto dest: subtype_implies_closed)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1640
  ultimately have X': "X' \<sharp> T11" by (rule closed_in_fresh)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1641
  from T_TApp have "D'[X \<mapsto> P]\<^sub>e @ G \<turnstile> t1[X \<mapsto>\<^sub>\<tau> P] : (\<forall>X'<:T11. T12)[X \<mapsto> P]\<^sub>\<tau>"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1642
    by simp
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1643
  with X' and T_TApp show ?case
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1644
    by (auto simp add: fresh_atm type_substitution_lemma
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1645
      fresh_list_append fresh_list_cons
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1646
      ctxt_subst_fresh' type_subst_fresh subst_trm_ty_fresh
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1647
      intro: substT_subtype)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1648
qed
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1649
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1650
lemma Abs_type: -- {* A.13(1) *}
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1651
  assumes H: "\<Gamma> \<turnstile> (\<lambda>x:S. s) : T"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1652
  and H': "\<Gamma> \<turnstile> T <: U \<rightarrow> U'"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1653
  and H'': "x \<sharp> \<Gamma>"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1654
  obtains S' where "\<Gamma> \<turnstile> U <: S"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1655
             and   "(VarB x S) # \<Gamma> \<turnstile> s : S'"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1656
             and   "\<Gamma> \<turnstile> S' <: U'"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1657
  using H H' H''
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1658
proof (nominal_induct \<Gamma> t \<equiv> "\<lambda>x:S. s" T avoiding: x arbitrary: U U' S s rule: typing.strong_induct)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1659
  case (T_Abs y T\<^isub>1 \<Gamma> t\<^isub>2 T\<^isub>2)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1660
  from `\<Gamma> \<turnstile> T\<^isub>1 \<rightarrow> T\<^isub>2 <: U \<rightarrow> U'`
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1661
  obtain ty1: "\<Gamma> \<turnstile> U <: S" and ty2: "\<Gamma> \<turnstile> T\<^isub>2 <: U'" using T_Abs
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1662
    by cases (simp_all add: ty.inject trm.inject alpha fresh_atm)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1663
  from T_Abs have "VarB y S # \<Gamma> \<turnstile> [(y, x)] \<bullet> s : T\<^isub>2"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1664
    by (simp add: trm.inject alpha fresh_atm)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1665
  then have "[(y, x)] \<bullet> (VarB y S # \<Gamma>) \<turnstile> [(y, x)] \<bullet> [(y, x)] \<bullet> s : [(y, x)] \<bullet> T\<^isub>2"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1666
    by (rule typing.eqvt)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1667
  moreover from T_Abs have "y \<sharp> \<Gamma>"
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1668
    by (auto dest!: typing_ok simp add: fresh_trm_dom)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1669
  ultimately have "VarB x S # \<Gamma> \<turnstile> s : T\<^isub>2" using T_Abs
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1670
    by (perm_simp add: ty_vrs_prm_simp)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1671
  with ty1 show ?case using ty2 by (rule T_Abs)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1672
next
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1673
  case (T_Sub \<Gamma> t S T)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1674
  then show ?case using subtype_transitivity by blast
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1675
qed simp_all
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1676
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1677
lemma subtype_reflexivity_from_typing:
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1678
  assumes "\<Gamma> \<turnstile> t : T"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1679
  shows "\<Gamma> \<turnstile> T <: T"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1680
using assms subtype_reflexivity typing_ok typing_closed_in by simp
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1681
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1682
lemma Abs_type':
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1683
  assumes H: "\<Gamma> \<turnstile> (\<lambda>x:S. s) : U \<rightarrow> U'"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1684
  and H': "x \<sharp> \<Gamma>"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1685
  obtains S'
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1686
  where "\<Gamma> \<turnstile> U <: S"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1687
  and "(VarB x S) # \<Gamma> \<turnstile> s : S'"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1688
  and "\<Gamma> \<turnstile> S' <: U'"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1689
  using H subtype_reflexivity_from_typing [OF H] H'
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1690
  by (rule Abs_type)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1691
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1692
lemma TAbs_type: -- {* A.13(2) *}
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1693
  assumes H: "\<Gamma> \<turnstile> (\<lambda>X<:S. s) : T"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1694
  and H': "\<Gamma> \<turnstile> T <: (\<forall>X<:U. U')"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1695
  and fresh: "X \<sharp> \<Gamma>" "X \<sharp> S" "X \<sharp> U"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1696
  obtains S'
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1697
  where "\<Gamma> \<turnstile> U <: S"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1698
  and   "(TVarB X U # \<Gamma>) \<turnstile> s : S'"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1699
  and   "(TVarB X U # \<Gamma>) \<turnstile> S' <: U'"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1700
  using H H' fresh
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1701
proof (nominal_induct \<Gamma> t \<equiv> "\<lambda>X<:S. s" T avoiding: X U U' S arbitrary: s rule: typing.strong_induct)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1702
  case (T_TAbs Y T\<^isub>1 \<Gamma> t\<^isub>2 T\<^isub>2)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1703
  from `TVarB Y T\<^isub>1 # \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2` have Y: "Y \<sharp> \<Gamma>"
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1704
    by (auto dest!: typing_ok simp add: valid_ty_dom_fresh)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1705
  from `Y \<sharp> U'` and `Y \<sharp> X`
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1706
  have "(\<forall>X<:U. U') = (\<forall>Y<:U. [(Y, X)] \<bullet> U')"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1707
    by (simp add: ty.inject alpha' fresh_atm)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1708
  with T_TAbs have "\<Gamma> \<turnstile> (\<forall>Y<:S. T\<^isub>2) <: (\<forall>Y<:U. [(Y, X)] \<bullet> U')" by (simp add: trm.inject)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1709
  then obtain ty1: "\<Gamma> \<turnstile> U <: S" and ty2: "(TVarB Y U # \<Gamma>) \<turnstile> T\<^isub>2 <: ([(Y, X)] \<bullet> U')" using T_TAbs Y
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1710
    by (cases rule: subtype_of.strong_cases [where X=Y]) (simp_all add: ty.inject alpha abs_fresh)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1711
  note ty1
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1712
  moreover from T_TAbs have "TVarB Y S # \<Gamma> \<turnstile> ([(Y, X)] \<bullet> s) : T\<^isub>2"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1713
    by (simp add: trm.inject alpha fresh_atm)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1714
  then have "[(Y, X)] \<bullet> (TVarB Y S # \<Gamma>) \<turnstile> [(Y, X)] \<bullet> [(Y, X)] \<bullet> s : [(Y, X)] \<bullet> T\<^isub>2"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1715
    by (rule typing.eqvt)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1716
  with `X \<sharp> \<Gamma>` `X \<sharp> S` Y `Y \<sharp> S` have "TVarB X S # \<Gamma> \<turnstile> s : [(Y, X)] \<bullet> T\<^isub>2"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1717
    by perm_simp
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1718
  then have "TVarB X U # \<Gamma> \<turnstile> s : [(Y, X)] \<bullet> T\<^isub>2" using ty1
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1719
    by (rule narrow_type [of "[]", simplified])
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1720
  moreover from ty2 have "([(Y, X)] \<bullet> (TVarB Y U # \<Gamma>)) \<turnstile> ([(Y, X)] \<bullet> T\<^isub>2) <: ([(Y, X)] \<bullet> [(Y, X)] \<bullet> U')"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1721
    by (rule subtype_of.eqvt)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1722
  with `X \<sharp> \<Gamma>` `X \<sharp> U` Y `Y \<sharp> U` have "(TVarB X U # \<Gamma>) \<turnstile> ([(Y, X)] \<bullet> T\<^isub>2) <: U'"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1723
    by perm_simp
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1724
  ultimately show ?case by (rule T_TAbs)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1725
next
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1726
  case (T_Sub \<Gamma> t S T)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1727
  then show ?case using subtype_transitivity by blast 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1728
qed simp_all
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1729
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1730
lemma TAbs_type':
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1731
  assumes H: "\<Gamma> \<turnstile> (\<lambda>X<:S. s) : (\<forall>X<:U. U')"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1732
  and fresh: "X \<sharp> \<Gamma>" "X \<sharp> S" "X \<sharp> U"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1733
  obtains S'
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1734
  where "\<Gamma> \<turnstile> U <: S"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1735
  and "(TVarB X U # \<Gamma>) \<turnstile> s : S'"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1736
  and "(TVarB X U # \<Gamma>) \<turnstile> S' <: U'"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1737
  using H subtype_reflexivity_from_typing [OF H] fresh
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1738
  by (rule TAbs_type)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1739
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1740
theorem preservation: -- {* A.20 *}
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1741
  assumes H: "\<Gamma> \<turnstile> t : T"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1742
  shows "t \<longmapsto> t' \<Longrightarrow> \<Gamma> \<turnstile> t' : T" using H
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1743
proof (nominal_induct avoiding: t' rule: typing.strong_induct)
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1744
  case (T_App \<Gamma> t\<^isub>1 T\<^isub>11 T\<^isub>12 t\<^isub>2 t')
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1745
  obtain x::vrs where x_fresh: "x \<sharp> (\<Gamma>, t\<^isub>1 \<cdot> t\<^isub>2, t')"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1746
    by (rule exists_fresh) (rule fin_supp)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1747
  obtain X::tyvrs where "X \<sharp> (t\<^isub>1 \<cdot> t\<^isub>2, t')"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1748
    by (rule exists_fresh) (rule fin_supp)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1749
  with `t\<^isub>1 \<cdot> t\<^isub>2 \<longmapsto> t'` show ?case
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1750
  proof (cases rule: eval.strong_cases [where x=x and X=X])
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1751
    case (E_Abs v\<^isub>2 T\<^isub>11' t\<^isub>12)
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1752
    with T_App and x_fresh have h: "\<Gamma> \<turnstile> (\<lambda>x:T\<^isub>11'. t\<^isub>12) : T\<^isub>11 \<rightarrow> T\<^isub>12"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1753
      by (simp add: trm.inject fresh_prod)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1754
    moreover from x_fresh have "x \<sharp> \<Gamma>" by simp
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1755
    ultimately obtain S'
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1756
      where T\<^isub>11: "\<Gamma> \<turnstile> T\<^isub>11 <: T\<^isub>11'"
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1757
      and t\<^isub>12: "(VarB x T\<^isub>11') # \<Gamma> \<turnstile> t\<^isub>12 : S'"
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1758
      and S': "\<Gamma> \<turnstile> S' <: T\<^isub>12"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1759
      by (rule Abs_type') blast
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1760
    from `\<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>11`
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1761
    have "\<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>11'" using T\<^isub>11 by (rule T_Sub)
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1762
    with t\<^isub>12 have "\<Gamma> \<turnstile> t\<^isub>12[x \<mapsto> t\<^isub>2] : S'" 
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1763
      by (rule subst_type [where \<Delta>="[]", simplified])
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1764
    hence "\<Gamma> \<turnstile> t\<^isub>12[x \<mapsto> t\<^isub>2] : T\<^isub>12" using S' by (rule T_Sub)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1765
    with E_Abs and x_fresh show ?thesis by (simp add: trm.inject fresh_prod)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1766
  next
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1767
    case (E_App1 t''' t'' u)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1768
    hence "t\<^isub>1 \<longmapsto> t''" by (simp add:trm.inject) 
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1769
    hence "\<Gamma> \<turnstile> t'' : T\<^isub>11 \<rightarrow> T\<^isub>12" by (rule T_App)
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1770
    hence "\<Gamma> \<turnstile> t'' \<cdot> t\<^isub>2 : T\<^isub>12" using `\<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>11`
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1771
      by (rule typing.T_App)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1772
    with E_App1 show ?thesis by (simp add:trm.inject)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1773
  next
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1774
    case (E_App2 v t''' t'')
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1775
    hence "t\<^isub>2 \<longmapsto> t''" by (simp add:trm.inject) 
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1776
    hence "\<Gamma> \<turnstile> t'' : T\<^isub>11" by (rule T_App)
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1777
    with T_App(1) have "\<Gamma> \<turnstile> t\<^isub>1 \<cdot> t'' : T\<^isub>12"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1778
      by (rule typing.T_App)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1779
    with E_App2 show ?thesis by (simp add:trm.inject) 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1780
  qed (simp_all add: fresh_prod)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1781
next
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1782
  case (T_TApp X \<Gamma> t\<^isub>1 T\<^isub>2  T\<^isub>11  T\<^isub>12 t')
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1783
  obtain x::vrs where "x \<sharp> (t\<^isub>1 \<cdot>\<^sub>\<tau> T\<^isub>2, t')"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1784
    by (rule exists_fresh) (rule fin_supp)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1785
  with `t\<^isub>1 \<cdot>\<^sub>\<tau> T\<^isub>2 \<longmapsto> t'`
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1786
  show ?case
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1787
  proof (cases rule: eval.strong_cases [where X=X and x=x])
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1788
    case (E_TAbs T\<^isub>11' T\<^isub>2' t\<^isub>12)
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1789
    with T_TApp have "\<Gamma> \<turnstile> (\<lambda>X<:T\<^isub>11'. t\<^isub>12) : (\<forall>X<:T\<^isub>11. T\<^isub>12)" and "X \<sharp> \<Gamma>" and "X \<sharp> T\<^isub>11'"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1790
      by (simp_all add: trm.inject)
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1791
    moreover from `\<Gamma>\<turnstile>T\<^isub>2<:T\<^isub>11` and `X \<sharp> \<Gamma>` have "X \<sharp> T\<^isub>11"
32011
01da62fb4a20 simplified proofs
Christian Urban <urbanc@in.tum.de>
parents: 30986
diff changeset
  1792
      by (blast intro: closed_in_fresh fresh_dom dest: subtype_implies_closed)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1793
    ultimately obtain S'
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1794
      where "TVarB X T\<^isub>11 # \<Gamma> \<turnstile> t\<^isub>12 : S'"
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1795
      and "(TVarB X T\<^isub>11 # \<Gamma>) \<turnstile> S' <: T\<^isub>12"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1796
      by (rule TAbs_type') blast
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1797
    hence "TVarB X T\<^isub>11 # \<Gamma> \<turnstile> t\<^isub>12 : T\<^isub>12" by (rule T_Sub)
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1798
    hence "\<Gamma> \<turnstile> t\<^isub>12[X \<mapsto>\<^sub>\<tau> T\<^isub>2] : T\<^isub>12[X \<mapsto> T\<^isub>2]\<^sub>\<tau>" using `\<Gamma> \<turnstile> T\<^isub>2 <: T\<^isub>11`
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1799
      by (rule substT_type [where D="[]", simplified])
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1800
    with T_TApp and E_TAbs show ?thesis by (simp add: trm.inject)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1801
  next
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1802
    case (E_TApp t''' t'' T)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1803
    from E_TApp have "t\<^isub>1 \<longmapsto> t''" by (simp add: trm.inject)
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1804
    then have "\<Gamma> \<turnstile> t'' : (\<forall>X<:T\<^isub>11. T\<^isub>12)" by (rule T_TApp)
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1805
    then have "\<Gamma> \<turnstile> t'' \<cdot>\<^sub>\<tau> T\<^isub>2 : T\<^isub>12[X \<mapsto> T\<^isub>2]\<^sub>\<tau>" using `\<Gamma> \<turnstile> T\<^isub>2 <: T\<^isub>11`
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1806
      by (rule better_T_TApp)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1807
    with E_TApp show ?thesis by (simp add: trm.inject)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1808
  qed (simp_all add: fresh_prod)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1809
next
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1810
  case (T_Sub \<Gamma> t S T t')
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1811
  have "t \<longmapsto> t'" by fact
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1812
  hence "\<Gamma> \<turnstile> t' : S" by (rule T_Sub)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1813
  moreover have "\<Gamma> \<turnstile> S <: T" by fact
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1814
  ultimately show ?case by (rule typing.T_Sub)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1815
qed (auto)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1816
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1817
lemma Fun_canonical: -- {* A.14(1) *}
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1818
  assumes ty: "[] \<turnstile> v : T\<^isub>1 \<rightarrow> T\<^isub>2"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1819
  shows "val v \<Longrightarrow> \<exists>x t S. v = (\<lambda>x:S. t)" using ty
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1820
proof (induct "[]::env" v "T\<^isub>1 \<rightarrow> T\<^isub>2" arbitrary: T\<^isub>1 T\<^isub>2)
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1821
  case (T_Sub t S)
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1822
  from `[] \<turnstile> S <: T\<^isub>1 \<rightarrow> T\<^isub>2`
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1823
  obtain S\<^isub>1 S\<^isub>2 where S: "S = S\<^isub>1 \<rightarrow> S\<^isub>2" 
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1824
    by cases (auto simp add: T_Sub)
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1825
  then show ?case using `val t` by (rule T_Sub)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1826
qed (auto)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1827
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1828
lemma TyAll_canonical: -- {* A.14(3) *}
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1829
  fixes X::tyvrs
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1830
  assumes ty: "[] \<turnstile> v : (\<forall>X<:T\<^isub>1. T\<^isub>2)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1831
  shows "val v \<Longrightarrow> \<exists>X t S. v = (\<lambda>X<:S. t)" using ty
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1832
proof (induct "[]::env" v "\<forall>X<:T\<^isub>1. T\<^isub>2" arbitrary: X T\<^isub>1 T\<^isub>2)
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1833
  case (T_Sub t S)
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1834
  from `[] \<turnstile> S <: (\<forall>X<:T\<^isub>1. T\<^isub>2)`
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1835
  obtain X S\<^isub>1 S\<^isub>2 where S: "S = (\<forall>X<:S\<^isub>1. S\<^isub>2)"
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1836
    by cases (auto simp add: T_Sub)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1837
  then show ?case using T_Sub by auto 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1838
qed (auto)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1839
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1840
theorem progress:
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1841
  assumes "[] \<turnstile> t : T"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1842
  shows "val t \<or> (\<exists>t'. t \<longmapsto> t')" 
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1843
using assms
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 32960
diff changeset
  1844
proof (induct "[]::env" t T)
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1845
  case (T_App t\<^isub>1 T\<^isub>11  T\<^isub>12 t\<^isub>2)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1846
  hence "val t\<^isub>1 \<or> (\<exists>t'. t\<^isub>1 \<longmapsto> t')" by simp
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1847
  thus ?case
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1848
  proof
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1849
    assume t\<^isub>1_val: "val t\<^isub>1"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1850
    with T_App obtain x t3 S where t\<^isub>1: "t\<^isub>1 = (\<lambda>x:S. t3)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1851
      by (auto dest!: Fun_canonical)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1852
    from T_App have "val t\<^isub>2 \<or> (\<exists>t'. t\<^isub>2 \<longmapsto> t')" by simp
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1853
    thus ?case
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1854
    proof
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1855
      assume "val t\<^isub>2"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1856
      with t\<^isub>1 have "t\<^isub>1 \<cdot> t\<^isub>2 \<longmapsto> t3[x \<mapsto> t\<^isub>2]" by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1857
      thus ?case by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1858
    next
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1859
      assume "\<exists>t'. t\<^isub>2 \<longmapsto> t'"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1860
      then obtain t' where "t\<^isub>2 \<longmapsto> t'" by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1861
      with t\<^isub>1_val have "t\<^isub>1 \<cdot> t\<^isub>2 \<longmapsto> t\<^isub>1 \<cdot> t'" by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1862
      thus ?case by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1863
    qed
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1864
  next
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1865
    assume "\<exists>t'. t\<^isub>1 \<longmapsto> t'"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1866
    then obtain t' where "t\<^isub>1 \<longmapsto> t'" by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1867
    hence "t\<^isub>1 \<cdot> t\<^isub>2 \<longmapsto> t' \<cdot> t\<^isub>2" by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1868
    thus ?case by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1869
  qed
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1870
next
39246
9e58f0499f57 modernized primrec
haftmann
parents: 35438
diff changeset
  1871
  case (T_TApp X t\<^isub>1 T\<^isub>2 T\<^isub>11 T\<^isub>12)
30091
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1872
  hence "val t\<^isub>1 \<or> (\<exists>t'. t\<^isub>1 \<longmapsto> t')" by simp
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1873
  thus ?case
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1874
  proof
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1875
    assume "val t\<^isub>1"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1876
    with T_TApp obtain x t S where "t\<^isub>1 = (\<lambda>x<:S. t)"
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1877
      by (auto dest!: TyAll_canonical)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1878
    hence "t\<^isub>1 \<cdot>\<^sub>\<tau> T\<^isub>2 \<longmapsto> t[x \<mapsto>\<^sub>\<tau> T\<^isub>2]" by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1879
    thus ?case by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1880
  next
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1881
    assume "\<exists>t'. t\<^isub>1 \<longmapsto> t'" thus ?case by auto
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1882
  qed
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1883
qed (auto)
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1884
2fb0b721e9c2 Added typing and evaluation relations, together with proofs of preservation
berghofe
parents: 29097
diff changeset
  1885
end