src/HOL/Nominal/Examples/Fsub.thy
author urbanc
Wed, 30 Nov 2005 19:08:51 +0100
changeset 18306 a28269045ff0
parent 18305 a780f9c1538b
child 18353 4dd468ccfdf7
permissions -rw-r--r--
started to change the transitivity/narrowing case: have trouble with Q=Q.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
18266
55c201fe4c95 added an authors section (please let me know if somebody is left out or unhappy)
urbanc
parents: 18263
diff changeset
     1
(* $Id$ *)
55c201fe4c95 added an authors section (please let me know if somebody is left out or unhappy)
urbanc
parents: 18263
diff changeset
     2
18269
3f36e2165e51 some small tuning
urbanc
parents: 18266
diff changeset
     3
theory fsub
3f36e2165e51 some small tuning
urbanc
parents: 18266
diff changeset
     4
imports "../nominal" 
3f36e2165e51 some small tuning
urbanc
parents: 18266
diff changeset
     5
begin
3f36e2165e51 some small tuning
urbanc
parents: 18266
diff changeset
     6
18266
55c201fe4c95 added an authors section (please let me know if somebody is left out or unhappy)
urbanc
parents: 18263
diff changeset
     7
text {* Authors: Christian Urban
55c201fe4c95 added an authors section (please let me know if somebody is left out or unhappy)
urbanc
parents: 18263
diff changeset
     8
                 Benjamin Pierce
55c201fe4c95 added an authors section (please let me know if somebody is left out or unhappy)
urbanc
parents: 18263
diff changeset
     9
                 Steve Zdancewic
55c201fe4c95 added an authors section (please let me know if somebody is left out or unhappy)
urbanc
parents: 18263
diff changeset
    10
                 Stephanie Weihrich
55c201fe4c95 added an authors section (please let me know if somebody is left out or unhappy)
urbanc
parents: 18263
diff changeset
    11
                 Dimitrios Vytiniotis
55c201fe4c95 added an authors section (please let me know if somebody is left out or unhappy)
urbanc
parents: 18263
diff changeset
    12
55c201fe4c95 added an authors section (please let me know if somebody is left out or unhappy)
urbanc
parents: 18263
diff changeset
    13
                 with help from Stefan Berghofer
55c201fe4c95 added an authors section (please let me know if somebody is left out or unhappy)
urbanc
parents: 18263
diff changeset
    14
      *}
55c201fe4c95 added an authors section (please let me know if somebody is left out or unhappy)
urbanc
parents: 18263
diff changeset
    15
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    16
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    17
atom_decl tyvrs vrs
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    18
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    19
section {* Types *}
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    20
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    21
nominal_datatype ty = TyVar "tyvrs"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    22
                    | Top
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    23
                    | Arrow  "ty" "ty"          ("_ \<rightarrow> _" [900,900] 900)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    24
                    | TAll  "\<guillemotleft>tyvrs\<guillemotright>ty" "ty" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    25
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    26
syntax
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    27
  TAll_syn  :: "tyvrs \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> ty" ("\<forall> [_<:_]._" [900,900,900] 900)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    28
translations 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    29
  "(\<forall>[a<:ty1].ty2)" \<rightleftharpoons> "TAll a ty2 ty1"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    30
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    31
section {* typing contexts and their domains *}
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    32
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    33
types ty_context = "(tyvrs\<times>ty) list"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    34
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    35
text {* domain of a context --- (the name "dom" is already used elsewhere) *}
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    36
consts
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    37
  "domain" :: "ty_context \<Rightarrow> tyvrs set"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    38
primrec
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    39
  "domain [] = {}"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    40
  "domain (X#\<Gamma>) = {fst X}\<union>(domain \<Gamma>)" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    41
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    42
lemma domain_eqvt:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    43
  fixes pi::"tyvrs prm"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    44
  shows "pi\<bullet>(domain \<Gamma>) = domain (pi\<bullet>\<Gamma>)" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    45
  by (induct \<Gamma>, auto simp add: perm_set_def)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    46
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    47
lemma finite_domain:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    48
  fixes \<Gamma>::"ty_context"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    49
  shows "finite (domain \<Gamma>)"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    50
  by (induct \<Gamma>, auto)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    51
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    52
lemma domain_inclusion[rule_format]:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    53
  shows "(X,T)\<in>set \<Gamma> \<longrightarrow> X\<in>(domain \<Gamma>)"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    54
  by (induct \<Gamma>, auto)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    55
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    56
lemma domain_existence[rule_format]:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    57
  shows "X\<in>(domain \<Gamma>) \<longrightarrow> (\<exists>T.(X,T)\<in>set \<Gamma>)"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    58
  by (induct \<Gamma>, auto)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    59
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    60
lemma domain_append:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    61
  shows "domain (\<Gamma>@\<Delta>) = ((domain \<Gamma>) \<union> (domain \<Delta>))"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    62
  by (induct \<Gamma>, auto)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    63
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    64
section {* Two functions over types *}
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    65
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    66
text {* they cannot yet be defined conveniently unless we have a recursion combinator *}
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    67
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    68
consts size_ty :: "ty \<Rightarrow> nat"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    69
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    70
lemma size_ty[simp]:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    71
  shows "size_ty (TyVar X) = 1"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    72
  and   "size_ty (Top) = 1"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    73
  and   "\<lbrakk>size_ty t1 = m ; size_ty t2 = n\<rbrakk> \<Longrightarrow> size_ty (t1 \<rightarrow> t2) = m + n + 1"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    74
  and   "\<lbrakk>size_ty t1 = m ; size_ty t2 = n\<rbrakk> \<Longrightarrow> size_ty (\<forall> [a<:t1].t2) = m + n + 1"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    75
sorry
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    76
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    77
consts subst_ty :: "ty \<Rightarrow> tyvrs \<Rightarrow> ty \<Rightarrow> ty"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    78
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    79
lemma subst_ty[simp]:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    80
  shows "subst_ty (TyVar X) Y T = (if X=Y then T else (TyVar X))"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    81
  and   "subst_ty Top Y T = Top"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    82
  and   "subst_ty (T1 \<rightarrow> T2) Y T = (subst_ty T1 Y T) \<rightarrow> (subst_ty T2 Y T)"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    83
  and   "X\<sharp>(Y,T) \<Longrightarrow> subst_ty (\<forall>[X<:T1].T2) Y T = (\<forall>[X<:(subst_ty T1 Y T)].(subst_ty T2 Y T))"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    84
  and   "\<lbrakk>X\<sharp>Y; X\<sharp>T\<rbrakk> \<Longrightarrow> subst_ty (\<forall>[X<:T1].T2) Y T = (\<forall>[X<:(subst_ty T1 Y T)].(subst_ty T2 Y T))"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    85
sorry
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    86
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    87
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    88
consts subst_ctxt :: "ty_context \<Rightarrow> tyvrs \<Rightarrow> ty \<Rightarrow> ty_context"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    89
primrec
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    90
"subst_ctxt [] Y T = []"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    91
"subst_ctxt (XT#Xs) Y T = (fst XT, subst_ty (snd XT) Y T)#(subst_ctxt Xs Y T)"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    92
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    93
text {* valid contexts *}
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    94
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    95
constdefs
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    96
  "closed_in" :: "ty \<Rightarrow> ty_context \<Rightarrow> bool" ("_ closed'_in _" [100,100] 100)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    97
  "S closed_in \<Gamma> \<equiv> (supp S)\<subseteq>(domain \<Gamma>)"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    98
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
    99
lemma closed_in_def2:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   100
  shows "(S closed_in \<Gamma>) = ((supp S)\<subseteq>((supp (domain \<Gamma>)):: tyvrs set))"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   101
apply(simp add: closed_in_def)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   102
apply(simp add: at_fin_set_supp[OF at_tyvrs_inst, OF finite_domain])
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   103
done
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   104
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   105
lemma closed_in_eqvt:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   106
  fixes pi::"tyvrs prm"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   107
  assumes a: "S closed_in \<Gamma>" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   108
  shows "(pi\<bullet>S) closed_in (pi\<bullet>\<Gamma>)"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   109
  using a
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   110
proof (unfold "closed_in_def")
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   111
  assume "supp S \<subseteq> (domain \<Gamma>)"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   112
  hence "pi\<bullet>(supp S) \<subseteq> pi\<bullet>(domain \<Gamma>)"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   113
    by (simp add: pt_subseteq_eqvt[OF pt_tyvrs_inst, OF at_tyvrs_inst])
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   114
  thus "(supp (pi\<bullet>S)) \<subseteq> (domain (pi\<bullet>\<Gamma>))"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   115
    by (simp add: domain_eqvt pt_perm_supp[OF pt_tyvrs_inst, OF at_tyvrs_inst])
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   116
qed
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   117
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   118
consts
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   119
  valid_rel :: "ty_context set" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   120
  valid_sym :: "ty_context \<Rightarrow> bool" ("\<turnstile> _ ok" [100] 100)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   121
translations
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   122
  "\<turnstile> \<Gamma> ok" \<rightleftharpoons> "\<Gamma> \<in> valid_rel"  
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   123
inductive valid_rel
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   124
intros
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   125
v1[intro]: "\<turnstile> [] ok"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   126
v2[intro]: "\<lbrakk>\<turnstile> \<Gamma> ok; X\<sharp>\<Gamma>; T closed_in \<Gamma>\<rbrakk>  \<Longrightarrow>  \<turnstile> ((X,T)#\<Gamma>) ok"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   127
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   128
lemma valid_eqvt:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   129
  fixes pi::"tyvrs prm"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   130
  assumes a: "\<turnstile> \<Gamma> ok" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   131
  shows "\<turnstile> (pi\<bullet>\<Gamma>) ok"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   132
  using a
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   133
proof (induct)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   134
  case v1
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   135
  show ?case by force
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   136
next
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   137
  case (v2 T X \<Gamma>)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   138
  have a1: "\<turnstile> (pi \<bullet> \<Gamma>) ok" by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   139
  have a2: "X\<sharp>\<Gamma>" by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   140
  have a3: "T closed_in \<Gamma>" by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   141
  show ?case
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   142
  proof (simp, rule valid_rel.v2)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   143
    show "\<turnstile> (pi \<bullet> \<Gamma>) ok" using a1 by simp
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   144
  next 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   145
    show "(pi\<bullet>X)\<sharp>(pi\<bullet>\<Gamma>)" using a2 by (simp add: pt_fresh_bij[OF pt_tyvrs_inst, OF at_tyvrs_inst])
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   146
  next
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   147
    show "(pi\<bullet>T) closed_in (pi\<bullet>\<Gamma>)" using a3 by (rule closed_in_eqvt)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   148
  qed
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   149
qed
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   150
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   151
lemma validE:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   152
  assumes a: "\<turnstile> ((X,T)#\<Gamma>) ok"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   153
  shows "\<turnstile> \<Gamma> ok \<and> X\<sharp>\<Gamma> \<and> T closed_in \<Gamma>"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   154
using a by (cases, auto)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   155
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   156
lemma validE_append[rule_format]:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   157
  shows "\<turnstile> (\<Delta>@\<Gamma>) ok \<longrightarrow> \<turnstile> \<Gamma> ok"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   158
  by (induct \<Delta>, auto dest: validE)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   159
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   160
lemma domain_fresh[rule_format]:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   161
  fixes X::"tyvrs"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   162
  shows "\<turnstile> \<Gamma> ok \<longrightarrow> (X\<sharp>(domain \<Gamma>) = X\<sharp>\<Gamma>)"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   163
apply(induct \<Gamma>)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   164
apply(simp add: fresh_list_nil fresh_set_empty)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   165
apply(simp add: fresh_list_cons fresh_prod 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   166
   fresh_fin_insert[OF pt_tyvrs_inst, OF at_tyvrs_inst, OF fs_tyvrs_inst, OF finite_domain])
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   167
apply(rule impI)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   168
apply(clarify)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   169
apply(simp add: fresh_prod)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   170
apply(drule validE)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   171
apply(simp)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   172
apply(simp add: closed_in_def2 fresh_def)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   173
apply(force)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   174
done
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   175
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   176
lemma closed_in_fresh:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   177
  fixes X::"tyvrs"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   178
  assumes a1: "S closed_in \<Gamma>"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   179
  and     a2: "X\<sharp>\<Gamma>" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   180
  and     a3: "\<turnstile> \<Gamma> ok"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   181
  shows "X\<sharp>S"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   182
using a1 a2 a3
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   183
apply(simp add: closed_in_def2)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   184
apply(simp add: domain_fresh[symmetric])
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   185
apply(simp add: fresh_def)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   186
apply(force)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   187
done
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   188
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   189
lemma replace_type[rule_format]:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   190
  shows "\<turnstile> (\<Delta>@(X,T)#\<Gamma>) ok \<longrightarrow> S closed_in \<Gamma> \<longrightarrow> \<turnstile> (\<Delta>@(X,S)#\<Gamma>) ok"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   191
apply(induct \<Delta>)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   192
apply(auto dest!: validE intro!: v2 simp add: fresh_list_append fresh_list_cons fresh_prod)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   193
apply(drule validE_append)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   194
apply(drule validE)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   195
apply(drule_tac S="S" in closed_in_fresh)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   196
apply(simp)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   197
apply(force)+
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   198
apply(simp add: closed_in_def2)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   199
apply(simp add: domain_append)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   200
done
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   201
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   202
lemma fresh_implies_not_member[rule_format]:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   203
  fixes \<Gamma>::"ty_context"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   204
  shows "X\<sharp>\<Gamma> \<longrightarrow> \<not>(\<exists>T.(X,T)\<in>(set \<Gamma>))"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   205
  apply (induct \<Gamma>)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   206
  apply (auto dest: validE simp add: fresh_list_cons fresh_prod fresh_atm)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   207
  done
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   208
 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   209
lemma uniqueness_of_ctxt:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   210
  fixes \<Gamma>::"ty_context"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   211
  shows "\<turnstile> \<Gamma> ok \<longrightarrow> (X,T)\<in>(set \<Gamma>) \<longrightarrow> (X,S)\<in>(set \<Gamma>) \<longrightarrow> (T = S)"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   212
  apply (induct \<Gamma>)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   213
  apply (auto dest!: validE fresh_implies_not_member)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   214
  done
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   215
 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   216
section {* Subtyping *}
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   217
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   218
consts
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   219
  subtype_of_rel :: "(ty_context \<times> ty \<times> ty) set"   
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   220
  subtype_of_sym :: "ty_context \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> bool"   ("_ \<turnstile> _ <: _" [100,100,100] 100)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   221
translations
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   222
  "\<Gamma> \<turnstile> S <: T" \<rightleftharpoons> "(\<Gamma>,S,T) \<in> subtype_of_rel"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   223
inductive subtype_of_rel
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   224
intros
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   225
S_Top[intro]:   "\<lbrakk>\<turnstile> \<Gamma> ok; S closed_in \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> S <: Top"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   226
S_Var[intro]:   "\<lbrakk>\<turnstile> \<Gamma> ok; (X,S) \<in> (set \<Gamma>); \<Gamma> \<turnstile> S <: T\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> (TyVar X) <: T"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   227
S_Refl[intro]:  "\<lbrakk>\<turnstile> \<Gamma> ok; X \<in> (domain \<Gamma>)\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> TyVar X <: TyVar X"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   228
S_Arrow[intro]: "\<lbrakk>\<Gamma> \<turnstile> T1 <: S1; \<Gamma> \<turnstile> S2 <: T2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> (S1 \<rightarrow> S2) <: (T1 \<rightarrow> T2)" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   229
S_All[intro]:   "\<lbrakk>\<Gamma> \<turnstile> T1 <: S1; X\<sharp>\<Gamma>; ((X,T1)#\<Gamma>) \<turnstile> S2 <: T2\<rbrakk> 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   230
                  \<Longrightarrow> \<Gamma> \<turnstile> \<forall> [X<:S1].S2 <: \<forall> [X<:T1].T2"  
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   231
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   232
lemma subtype_implies_closed:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   233
  assumes a: "\<Gamma> \<turnstile> S <: T"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   234
  shows "S closed_in \<Gamma> \<and> T closed_in \<Gamma>"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   235
using a
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   236
proof (induct)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   237
  case (S_Top S \<Gamma>)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   238
  have "Top closed_in \<Gamma>" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   239
    by (simp add: closed_in_def ty.supp)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   240
  moreover
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   241
  have "S closed_in \<Gamma>" by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   242
  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
   243
next
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   244
  case (S_Var S T X \<Gamma>)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   245
  have "(X,S)\<in>set \<Gamma>" by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   246
  hence "X\<in>(domain \<Gamma>)" by (rule domain_inclusion)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   247
  hence "(TyVar X) closed_in \<Gamma>" by (simp add: closed_in_def ty.supp supp_atm)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   248
  moreover
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   249
  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
   250
  hence "T closed_in \<Gamma>" by force
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   251
  ultimately show "(TyVar X) closed_in \<Gamma> \<and> T closed_in \<Gamma>" by simp
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   252
next
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   253
  case S_Refl thus ?case 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   254
    by (simp add: closed_in_def ty.supp supp_atm) 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   255
next
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   256
  case S_Arrow thus ?case by (force simp add: closed_in_def ty.supp)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   257
next 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   258
  case S_All thus ?case by (auto simp add: closed_in_def ty.supp abs_supp)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   259
qed
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   260
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   261
text {* completely automated proof *}
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   262
lemma subtype_implies_closed:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   263
  assumes a: "\<Gamma> \<turnstile> S <: T"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   264
  shows "S closed_in \<Gamma> \<and> T closed_in \<Gamma>"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   265
using a
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   266
apply (induct) 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   267
apply (auto simp add: closed_in_def ty.supp abs_supp domain_inclusion supp_atm)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   268
done
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   269
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   270
lemma subtype_implies_ok:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   271
  fixes X::"tyvrs"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   272
  assumes a1: "\<Gamma> \<turnstile> S <: T"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   273
  shows "\<turnstile> \<Gamma> ok"  
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   274
using a1 by (induct, auto)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   275
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   276
lemma subtype_implies_fresh:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   277
  fixes X::"tyvrs"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   278
  assumes a1: "\<Gamma> \<turnstile> S <: T"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   279
  and     a2: "X\<sharp>\<Gamma>"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   280
  shows "X\<sharp>(S,T)"  
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   281
proof -
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   282
  from a1 have "\<turnstile> \<Gamma> ok" by (rule subtype_implies_ok)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   283
  with a2 have b0: "X\<sharp>domain(\<Gamma>)" by (simp add: domain_fresh)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   284
  from a1 have "S closed_in \<Gamma> \<and> T closed_in \<Gamma>" by (rule subtype_implies_closed)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   285
  hence b1: "supp S \<subseteq> ((supp (domain \<Gamma>))::tyvrs set)" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   286
    and b2: "supp T \<subseteq> ((supp (domain \<Gamma>))::tyvrs set)" by (simp_all add: closed_in_def2)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   287
  thus "X\<sharp>(S,T)" using b0 b1 b2 by (force simp add: supp_prod fresh_def)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   288
qed
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   289
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   290
lemma silly_eqvt1:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   291
  fixes X::"'a::pt_tyvrs"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   292
  and   S::"'b::pt_tyvrs"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   293
  and   pi::"tyvrs prm"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   294
 shows "(X,S) \<in> set \<Gamma> \<Longrightarrow> (pi\<bullet>X,pi\<bullet>S) \<in> set (pi\<bullet>\<Gamma>)"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   295
apply(drule_tac pi="pi" in pt_set_bij2[OF pt_tyvrs_inst, OF at_tyvrs_inst])
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   296
apply(simp add: pt_list_set_pi[OF pt_tyvrs_inst])
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   297
done
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   298
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   299
lemma silly_eqvt2:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   300
  fixes X::"tyvrs"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   301
  and   pi::"tyvrs prm"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   302
 shows "X \<in> (domain \<Gamma>) \<Longrightarrow> (pi\<bullet>X) \<in> (domain (pi\<bullet>\<Gamma>))"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   303
apply(drule_tac pi="pi" in pt_set_bij2[OF pt_tyvrs_inst, OF at_tyvrs_inst])
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   304
apply(simp add: domain_eqvt pt_list_set_pi[OF pt_tyvrs_inst] )
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   305
done
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   306
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   307
lemma subtype_eqvt:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   308
  fixes pi::"tyvrs prm"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   309
  shows "\<Gamma> \<turnstile> S <: T \<Longrightarrow> (pi\<bullet>\<Gamma>) \<turnstile> (pi\<bullet>S) <: (pi\<bullet>T)"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   310
apply(erule subtype_of_rel.induct)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   311
apply(force intro: valid_eqvt closed_in_eqvt)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   312
apply(force intro: closed_in_eqvt valid_eqvt silly_eqvt1)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   313
apply(force intro: valid_eqvt silly_eqvt2)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   314
apply(force)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   315
apply(force intro!: S_All simp add: fresh_prod pt_fresh_bij1[OF pt_tyvrs_inst, OF at_tyvrs_inst])
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   316
done
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   317
18305
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   318
lemma subtype_of_rel_induct[consumes 1, case_names S_Top S_Var S_Refl S_Arrow S_All]:
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   319
  fixes  P :: "'a::fs_tyvrs\<Rightarrow>ty_context \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow>bool"
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   320
  assumes a: "\<Gamma> \<turnstile> S <: T"
18305
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   321
  and a1:    "\<And>\<Gamma> S x. (\<turnstile> \<Gamma> ok) \<Longrightarrow> (S closed_in \<Gamma>) \<Longrightarrow> P x \<Gamma> S Top"
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   322
  and a2:    "\<And>\<Gamma> X S T x. (\<turnstile> \<Gamma> ok) \<Longrightarrow> ((X,S)\<in>set \<Gamma>) \<Longrightarrow> (\<Gamma> \<turnstile> S <: T) \<Longrightarrow> (\<And>z. P z \<Gamma> S T)
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   323
              \<Longrightarrow> P x \<Gamma> (TyVar X) T"
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   324
  and a3:    "\<And>\<Gamma> X x. (\<turnstile> \<Gamma> ok) \<Longrightarrow> X\<in>domain \<Gamma> \<Longrightarrow> P x \<Gamma> (TyVar X) (TyVar X)"  
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   325
  and a4:    "\<And>\<Gamma> S1 S2 T1 T2 x. \<Gamma> \<turnstile> T1 <: S1 \<Longrightarrow> (\<And>z. P z \<Gamma> T1 S1) \<Longrightarrow> \<Gamma> \<turnstile> S2 <: T2 \<Longrightarrow> 
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   326
              (\<And>z. P z \<Gamma> S2 T2) \<Longrightarrow> P x \<Gamma> (S1 \<rightarrow> S2) (T1 \<rightarrow> T2)"
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   327
  and a5:    "\<And>\<Gamma> X S1 S2 T1 T2 x. 
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   328
              X\<sharp>x \<Longrightarrow> X\<sharp>(\<Gamma>,S1,T1) \<Longrightarrow> \<Gamma> \<turnstile> T1 <: S1 \<Longrightarrow> (\<And>z. P z \<Gamma> T1 S1) \<Longrightarrow> ((X,T1)#\<Gamma>) \<turnstile> S2 <: T2 
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   329
              \<Longrightarrow> (\<And>z. P z ((X,T1)#\<Gamma>) S2 T2) \<Longrightarrow> P x \<Gamma> (\<forall> [X<:S1].S2) (\<forall> [X<:T1].T2)"
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   330
  shows "P x \<Gamma> S T"
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   331
proof -
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   332
  from a have "\<And>(pi::tyvrs prm) (x::'a::fs_tyvrs). P x (pi\<bullet>\<Gamma>) (pi\<bullet>S) (pi\<bullet>T)"
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   333
  proof (induct)
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   334
    case (S_Top S \<Gamma>)
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   335
    have b1: "\<turnstile> \<Gamma> ok" by fact 
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   336
    have b2: "S closed_in \<Gamma>" by fact
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   337
    from b1 have "\<turnstile> (pi\<bullet>\<Gamma>) ok" by (rule valid_eqvt)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   338
    moreover
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   339
    from b2 have "(pi\<bullet>S) closed_in (pi\<bullet>\<Gamma>)" by (rule closed_in_eqvt)
18305
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   340
    ultimately show "P x (pi \<bullet> \<Gamma>) (pi \<bullet> S) (pi\<bullet>Top)" by (simp add: a1)
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   341
  next
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   342
    case (S_Var S T X \<Gamma>)
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   343
    have b1: "\<turnstile> \<Gamma> ok" by fact 
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   344
    have b2: "(X,S) \<in> set \<Gamma>" by fact
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   345
    have b3: "\<Gamma> \<turnstile> S <: T" by fact
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   346
    have b4: "\<And>(pi::tyvrs prm) x. P x (pi\<bullet>\<Gamma>) (pi\<bullet>S) (pi\<bullet>T)" by fact
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   347
    from b1 have "\<turnstile> (pi\<bullet>\<Gamma>) ok" by (rule valid_eqvt)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   348
    moreover
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   349
    from b2 have "pi\<bullet>(X,S)\<in>pi\<bullet>(set \<Gamma>)" by (rule pt_set_bij2[OF pt_tyvrs_inst, OF at_tyvrs_inst])
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   350
    hence "(pi\<bullet>X,pi\<bullet>S)\<in>set (pi\<bullet>\<Gamma>)" by (simp add: pt_list_set_pi[OF pt_tyvrs_inst])
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   351
    moreover 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   352
    from b3 have "(pi\<bullet>\<Gamma>) \<turnstile> (pi\<bullet>S) <: (pi\<bullet>T)" by (rule subtype_eqvt)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   353
    moreover 
18305
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   354
    from b4 have "\<And>x. P x (pi\<bullet>\<Gamma>) (pi\<bullet>S) (pi\<bullet>T)" by simp
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   355
    ultimately 
18305
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   356
    show "P x (pi\<bullet>\<Gamma>) (pi\<bullet>(TyVar X)) (pi\<bullet>T)" by (simp add: a2)
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   357
  next
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   358
    case (S_Refl X \<Gamma>)
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   359
    have b1: "\<turnstile> \<Gamma> ok" by fact
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   360
    have b2: "X \<in> domain \<Gamma>" by fact
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   361
    from b1 have "\<turnstile> (pi\<bullet>\<Gamma>) ok" by (rule valid_eqvt)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   362
    moreover
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   363
    from b2 have "(pi\<bullet>X)\<in>pi\<bullet>domain \<Gamma>" by (rule pt_set_bij2[OF pt_tyvrs_inst, OF at_tyvrs_inst])
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   364
    hence "(pi\<bullet>X)\<in>domain (pi\<bullet>\<Gamma>)" by (simp add: domain_eqvt pt_list_set_pi[OF pt_tyvrs_inst])
18305
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   365
    ultimately show "P x (pi\<bullet>\<Gamma>) (pi\<bullet>(TyVar X)) (pi\<bullet>(TyVar X))" by (simp add: a3)
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   366
  next
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   367
    case S_Arrow thus ?case by (simp, auto intro!: a4 subtype_eqvt)
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   368
  next
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   369
    case (S_All S1 S2 T1 T2 X \<Gamma>)
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   370
    have b1: "\<Gamma> \<turnstile> T1 <: S1" by fact 
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   371
    have b2: "\<And>(pi::tyvrs prm) x. P x (pi\<bullet>\<Gamma>) (pi\<bullet>T1) (pi\<bullet>S1)" by fact
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   372
    have b4: "((X,T1)#\<Gamma>) \<turnstile> S2 <: T2" by fact
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   373
    have b5: "\<And>(pi::tyvrs prm) x. P x (pi\<bullet>((X,T1)#\<Gamma>)) (pi\<bullet>S2) (pi\<bullet>T2)" by fact
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   374
    have b3': "X\<sharp>\<Gamma>" by fact
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   375
    have b3'': "X\<sharp>(T1,S1)" using b1 b3' by (rule subtype_implies_fresh)
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   376
    have b3: "X\<sharp>(\<Gamma>,S1,T1)" using b3' b3'' by (simp add: fresh_prod)
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   377
    have "\<exists>C::tyvrs. C\<sharp>(pi\<bullet>X,pi\<bullet>S2,pi\<bullet>T2,pi\<bullet>S1,pi\<bullet>T1,pi\<bullet>\<Gamma>,x)"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   378
      by (rule at_exists_fresh[OF at_tyvrs_inst], simp add: fs_tyvrs1)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   379
    then obtain C::"tyvrs" where  f1: "C\<noteq>(pi\<bullet>X)" and f2: "C\<sharp>(pi\<bullet>S1)" and f3: "C\<sharp>(pi\<bullet>T1)"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   380
      and f4: "C\<sharp>(pi\<bullet>S2)" and f5: "C\<sharp>(pi\<bullet>T2)" and f6: "C\<sharp>(pi\<bullet>\<Gamma>)" and f7: "C\<sharp>x"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   381
      by (auto simp add: fresh_prod fresh_atm)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   382
    let ?pi' = "[(C,pi\<bullet>X)]@pi"
18305
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   383
    from b2 have c1: "\<And>x. P x (?pi'\<bullet>\<Gamma>) (?pi'\<bullet>T1) (?pi'\<bullet>S1)" by simp
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   384
    from b5 have "\<And>x. P x (?pi'\<bullet>((X,T1)#\<Gamma>)) (?pi'\<bullet>S2) (?pi'\<bullet>T2)" by simp
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   385
    hence "\<And>x. P x ((?pi'\<bullet>X,?pi'\<bullet>T1)#(?pi'\<bullet>\<Gamma>)) (?pi'\<bullet>S2) (?pi'\<bullet>T2)" by simp
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   386
    hence c2: "\<And>x. P x ((C,?pi'\<bullet>T1)#(?pi'\<bullet>\<Gamma>)) (?pi'\<bullet>S2) (?pi'\<bullet>T2)" using f1
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   387
      by (simp only: pt_tyvrs2 calc_atm, simp)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   388
    from b3 have "(pi\<bullet>X)\<sharp>(pi\<bullet>\<Gamma>)" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   389
      by (simp add: fresh_prod pt_fresh_bij[OF pt_tyvrs_inst, OF at_tyvrs_inst])
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   390
    with f6 have f6a: "?pi'\<bullet>\<Gamma>=pi\<bullet>\<Gamma>" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   391
      by (simp only: pt_tyvrs2 pt_fresh_fresh[OF pt_tyvrs_inst, OF at_tyvrs_inst])
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   392
    hence f6': "C\<sharp>(?pi'\<bullet>\<Gamma>)" using f6 by simp
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   393
    from b3 have "(pi\<bullet>X)\<sharp>(pi\<bullet>S1)" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   394
      by (simp add: fresh_prod pt_fresh_bij[OF pt_tyvrs_inst, OF at_tyvrs_inst])
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   395
    with f2 have f2a: "?pi'\<bullet>S1=pi\<bullet>S1" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   396
      by (simp only: pt_tyvrs2 pt_fresh_fresh[OF pt_tyvrs_inst, OF at_tyvrs_inst])
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   397
    hence f2': "C\<sharp>(?pi'\<bullet>S1)" using f2 by simp
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   398
    from b3 have "(pi\<bullet>X)\<sharp>(pi\<bullet>T1)" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   399
      by (simp add: fresh_prod pt_fresh_bij[OF pt_tyvrs_inst, OF at_tyvrs_inst])
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   400
    with f3 have f3a: "?pi'\<bullet>T1=pi\<bullet>T1" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   401
      by (simp only: pt_tyvrs2 pt_fresh_fresh[OF pt_tyvrs_inst, OF at_tyvrs_inst])
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   402
    hence f3': "C\<sharp>(?pi'\<bullet>T1)" using f3 by simp
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   403
    from b1 have e1: "(?pi'\<bullet>\<Gamma>) \<turnstile> (?pi'\<bullet>T1) <: (?pi'\<bullet>S1)" by (rule subtype_eqvt)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   404
    from b4 have "(?pi'\<bullet>((X,T1)#\<Gamma>)) \<turnstile> (?pi'\<bullet>S2) <: (?pi'\<bullet>T2)" by (rule subtype_eqvt)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   405
    hence "((?pi'\<bullet>X,?pi'\<bullet>T1)#(?pi'\<bullet>\<Gamma>)) \<turnstile> (?pi'\<bullet>S2) <: (?pi'\<bullet>T2)" by simp
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   406
    hence e2: "((C,?pi'\<bullet>T1)#(?pi'\<bullet>\<Gamma>)) \<turnstile> (?pi'\<bullet>S2) <: (?pi'\<bullet>T2)" using f1
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   407
      by (simp only: pt_tyvrs2 calc_atm, simp)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   408
    have fnew: "C\<sharp>(?pi'\<bullet>\<Gamma>,?pi'\<bullet>S1,?pi'\<bullet>T1)" using f6' f2' f3' by (simp add: fresh_prod)
18305
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   409
    have main: "P x (?pi'\<bullet>\<Gamma>) (\<forall> [C<:(?pi'\<bullet>S1)].(?pi'\<bullet>S2)) (\<forall> [C<:(?pi'\<bullet>T1)].(?pi'\<bullet>T2))"
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   410
      using f7 fnew e1 c1 e2 c2 by (rule a5)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   411
    have alpha1: "(\<forall> [C<:(?pi'\<bullet>S1)].(?pi'\<bullet>S2)) = (pi\<bullet>(\<forall> [X<:S1].S2))"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   412
      using f1 f4 f2a[symmetric] by (simp add: ty.inject alpha pt_tyvrs2[symmetric])
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   413
    have alpha2: "(\<forall> [C<:(?pi'\<bullet>T1)].(?pi'\<bullet>T2)) = (pi\<bullet>(\<forall> [X<:T1].T2))"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   414
      using f1 f5 f3a[symmetric] by (simp add: ty.inject alpha pt_tyvrs2[symmetric])
18305
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   415
    show "P x (pi\<bullet>\<Gamma>) (pi\<bullet>(\<forall> [X<:S1].S2)) (pi\<bullet>(\<forall> [X<:T1].T2))"
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   416
      using alpha1 alpha2 f6a main by simp  
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   417
  qed
18305
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   418
  hence "P x (([]::tyvrs prm)\<bullet>\<Gamma>) (([]::tyvrs prm)\<bullet>S) (([]::tyvrs prm)\<bullet>T)" by blast
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   419
  thus ?thesis by simp
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   420
qed
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   421
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   422
section {* Two proofs for reflexivity of subtyping *}
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   423
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   424
lemma subtype_reflexivity:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   425
  shows "\<turnstile> \<Gamma> ok \<longrightarrow> T closed_in \<Gamma> \<longrightarrow> \<Gamma> \<turnstile> T <: T"
18305
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   426
proof(nominal_induct T fresh: \<Gamma> rule: ty.induct_unsafe)
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   427
  case (TAll X T1 T2)
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   428
  have i1: "\<And>\<Gamma>. \<turnstile> \<Gamma> ok \<longrightarrow> T1 closed_in \<Gamma> \<longrightarrow> \<Gamma> \<turnstile> T1 <: T1" by fact 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   429
  have i2: "\<And>\<Gamma>. \<turnstile> \<Gamma> ok \<longrightarrow> T2 closed_in \<Gamma> \<longrightarrow> \<Gamma> \<turnstile> T2 <: T2" by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   430
  have f: "X\<sharp>\<Gamma>" by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   431
  show "\<turnstile> \<Gamma> ok \<longrightarrow> (\<forall>[X<:T2].T1) closed_in \<Gamma> \<longrightarrow> \<Gamma> \<turnstile>  \<forall>[X<:T2].T1 <: \<forall>[X<:T2].T1"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   432
  proof (intro strip)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   433
    assume "(\<forall>[X<:T2].T1) closed_in \<Gamma>"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   434
    hence b1: "T2 closed_in \<Gamma>" and b2: "T1 closed_in ((X,T2)#\<Gamma>)" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   435
      by (auto simp add: closed_in_def ty.supp abs_supp)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   436
    assume c1: "\<turnstile> \<Gamma> ok"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   437
    hence c2: "\<turnstile> ((X,T2)#\<Gamma>) ok" using b1 f by force
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   438
    have "\<Gamma> \<turnstile> T2 <: T2" using i2 b1 c1 by simp
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   439
    moreover
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   440
    have "((X,T2)#\<Gamma>) \<turnstile> T1 <: T1" using i1 b2 c2 by simp
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   441
    ultimately show "\<Gamma> \<turnstile> \<forall>[X<:T2].T1 <: \<forall>[X<:T2].T1" using f by force
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   442
  qed
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   443
qed (auto simp add: closed_in_def ty.supp supp_atm)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   444
18305
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   445
lemma subtype_reflexivity:
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   446
  assumes a: "\<turnstile> \<Gamma> ok"
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   447
  and     b: "T closed_in \<Gamma>"
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   448
  shows "\<Gamma> \<turnstile> T <: T"
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   449
using a b
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   450
apply(nominal_induct T fresh: \<Gamma> rule: ty.induct_unsafe)
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   451
apply(auto simp add: ty.supp abs_supp closed_in_def supp_atm)
18305
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   452
(* too bad that this cannot be found automatically *)
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   453
apply(drule_tac x="(tyvrs, ty2)#\<Gamma>" in meta_spec)
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   454
apply(force simp add: closed_in_def)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   455
done
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   456
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   457
text {* Inversion lemmas\<dots> *}
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   458
lemma S_TopE:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   459
 shows "\<Gamma> \<turnstile> Top <: T \<Longrightarrow> T = Top" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   460
apply(ind_cases "\<Gamma> \<turnstile> Top <: T", auto)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   461
done
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   462
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   463
lemma S_ArrowE_left:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   464
  assumes a: "\<Gamma> \<turnstile> S1 \<rightarrow> S2 <: T" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   465
  shows "T = Top \<or> (\<exists>T1 T2. T = T1 \<rightarrow> T2 \<and> \<Gamma> \<turnstile> T1 <: S1 \<and> \<Gamma> \<turnstile> S2 <: T2)"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   466
using  a by (cases, auto simp add: ty.inject)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   467
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   468
lemma S_AllE_left:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   469
  shows "\<lbrakk>\<Gamma> \<turnstile> \<forall>[X<:S1].S2 <: T; X\<sharp>(\<Gamma>,S1)\<rbrakk>
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   470
         \<Longrightarrow> T = Top \<or> (\<exists>T1 T2. T = \<forall>[X<:T1].T2 \<and> \<Gamma> \<turnstile> T1 <: S1 \<and> ((X,T1)#\<Gamma>) \<turnstile> S2 <: T2)"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   471
  apply(frule subtype_implies_ok)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   472
  apply(ind_cases "\<Gamma> \<turnstile> \<forall>[X<:S1].S2 <: T")
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   473
  apply(auto simp add: ty.inject alpha)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   474
  apply(rule_tac x="[(X,Xa)]\<bullet>T2" in exI)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   475
  (* term *)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   476
  apply(rule conjI)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   477
  apply(rule sym)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   478
  apply(rule pt_bij2[OF pt_tyvrs_inst, OF at_tyvrs_inst])
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   479
  apply(rule pt_tyvrs3)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   480
  apply(simp)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   481
  apply(rule at_ds5[OF at_tyvrs_inst])
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   482
  (* 1st conjunct *)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   483
  apply(rule conjI)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   484
  apply(simp add: pt_fresh_left[OF pt_tyvrs_inst, OF at_tyvrs_inst] calc_atm)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   485
  apply(simp add: fresh_prod)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   486
  apply(drule_tac \<Gamma>="((Xa,T1)#\<Gamma>)" in  subtype_implies_closed)+
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   487
  apply(simp add: closed_in_def)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   488
  apply(auto simp add: domain_fresh[of "\<Gamma>" "X", symmetric])
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   489
  apply(simp add: fresh_def)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   490
  apply(subgoal_tac "X \<notin> (insert Xa (domain \<Gamma>))")(*A*)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   491
  apply(force)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   492
  (*A*)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   493
  apply(simp add: at_fin_set_supp[OF at_tyvrs_inst, OF finite_domain])
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   494
  (* 2nd conjunct *)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   495
  apply(frule_tac X="X" in subtype_implies_fresh)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   496
  apply(simp add: fresh_prod)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   497
  apply(drule_tac X="Xa" in subtype_implies_fresh)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   498
  apply(assumption)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   499
  apply(simp add: fresh_prod)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   500
  apply(drule_tac pi="[(X,Xa)]" in subtype_eqvt)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   501
  apply(simp add: calc_atm)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   502
  apply(simp add: pt_fresh_fresh[OF pt_tyvrs_inst, OF at_tyvrs_inst])
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   503
  done
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   504
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   505
section {* Type substitution *}
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   506
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   507
lemma subst_ty_fresh[rule_format]:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   508
  fixes P :: "ty"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   509
  and   X :: "tyvrs"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   510
  shows "X\<sharp>(T,P) \<longrightarrow> X\<sharp>(subst_ty T Y P)"
18305
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   511
  apply (nominal_induct T fresh: T Y P rule: ty.induct_unsafe)
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   512
  apply (auto simp add: fresh_prod abs_fresh)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   513
  done
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   514
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   515
lemma subst_ctxt_fresh[rule_format]:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   516
  fixes X::"tyvrs"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   517
  shows "X\<sharp>(\<Gamma>,P) \<longrightarrow> X\<sharp>(subst_ctxt \<Gamma> Y P)"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   518
  apply (induct \<Gamma>)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   519
  apply (auto intro: subst_ty_fresh simp add: fresh_prod fresh_list_cons)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   520
  done
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   521
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   522
(*
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   523
lemma subst_ctxt_closed:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   524
  shows  "subst_ty b X P closed_in (subst_ctxt \<Delta> X P @ \<Gamma>)"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   525
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   526
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   527
lemma substitution_ok:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   528
  shows "\<turnstile> (\<Delta>@(X,Q)#\<Gamma>) ok \<longrightarrow> \<Gamma> \<turnstile> P <: Q \<longrightarrow> \<turnstile> ((subst_ctxt \<Delta> X P)@\<Gamma>)  ok"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   529
  apply (induct \<Delta>)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   530
  apply (auto dest: validE)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   531
  apply (rule v2)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   532
  apply assumption
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   533
  apply (drule validE)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   534
  apply (auto simp add: fresh_list_append)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   535
  apply (rule subst_ctxt_fresh)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   536
  apply (simp add: fresh_prod)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   537
  apply (drule_tac X = "a" in subtype_implies_fresh)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   538
  apply (simp add: fresh_list_cons)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   539
  apply (simp add: fresh_prod)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   540
  apply (simp add: fresh_list_cons)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   541
  apply (drule validE)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   542
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   543
done
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   544
*)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   545
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   546
(* note: What should nominal induct do if the context is compound? *)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   547
(*
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   548
lemma type_substitution:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   549
  assumes a0: "(\<Delta>@(X,Q)#\<Gamma>) \<turnstile> S <: T"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   550
  shows "\<turnstile> (\<Delta>@(X,Q)#\<Gamma>) ok \<longrightarrow> \<Gamma> \<turnstile> P <: Q 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   551
         \<longrightarrow> ((subst_ctxt \<Delta> X P)@\<Gamma>) \<turnstile> (subst_ty S X P) <: (subst_ty T X P)"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   552
  using a0 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   553
  thm subtype_of_rel_induct
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   554
  apply(rule subtype_of_rel_induct[of "\<Delta>@(X,Q)#\<Gamma>" "S" "T" _ "P"])
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   555
  apply(auto)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   556
  apply(rule S_Top)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   557
  defer
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   558
  defer
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   559
  defer
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   560
  apply(rule S_Var)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   561
  defer
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   562
  defer
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   563
  defer
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   564
  defer
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   565
  apply(rule S_Refl)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   566
  defer
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   567
  defer
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   568
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   569
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   570
apply (nominal_induct \<Delta> X Q \<Gamma> S T rule: subtype_of_rel_induct)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   571
*)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   572
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   573
section {* Weakening *}
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   574
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   575
constdefs 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   576
  extends :: "ty_context \<Rightarrow> ty_context \<Rightarrow> bool" ("_ extends _" [100,100] 100)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   577
  "\<Delta> extends \<Gamma> \<equiv> \<forall>X Q. (X,Q)\<in>set \<Gamma> \<longrightarrow> (X,Q)\<in>set \<Delta>"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   578
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   579
lemma extends_domain:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   580
  assumes a: "\<Delta> extends \<Gamma>"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   581
  shows "domain \<Gamma> \<subseteq> domain \<Delta>"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   582
  using a 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   583
  apply (auto simp add: extends_def)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   584
  apply (drule domain_existence)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   585
  apply (force simp add: domain_inclusion)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   586
  done
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   587
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   588
lemma extends_closed:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   589
  assumes a1: "T closed_in \<Gamma>"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   590
  and     a2: "\<Delta> extends \<Gamma>"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   591
  shows "T closed_in \<Delta>"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   592
  using a1 a2
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   593
  by (auto dest: extends_domain simp add: closed_in_def)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   594
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   595
lemma weakening:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   596
  assumes a1: "\<Gamma> \<turnstile> S <: T"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   597
  shows "\<turnstile> \<Delta> ok \<longrightarrow> \<Delta> extends \<Gamma> \<longrightarrow> \<Delta> \<turnstile> S <: T"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   598
  using a1 
18305
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   599
proof (nominal_induct \<Gamma> S T fresh: \<Delta> rule: subtype_of_rel_induct)
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   600
  case (S_Top \<Gamma> S) 
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   601
  have lh_drv_prem: "S closed_in \<Gamma>" by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   602
  show "\<turnstile> \<Delta> ok \<longrightarrow> \<Delta> extends \<Gamma> \<longrightarrow> \<Delta> \<turnstile> S <: Top"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   603
  proof (intro strip)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   604
    assume "\<turnstile> \<Delta> ok"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   605
    moreover
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   606
    assume "\<Delta> extends \<Gamma>"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   607
    hence "S closed_in \<Delta>" using lh_drv_prem by (rule_tac extends_closed)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   608
    ultimately show "\<Delta> \<turnstile> S <: Top" by force
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   609
  qed
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   610
next 
18305
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   611
  case (S_Var \<Gamma> X S T)
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   612
  have lh_drv_prem: "(X,S) \<in> set \<Gamma>" by fact
18305
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   613
  have ih: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<longrightarrow> \<Delta> extends \<Gamma> \<longrightarrow> \<Delta> \<turnstile> S <: T" by fact
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   614
  show "\<turnstile> \<Delta> ok \<longrightarrow> \<Delta> extends \<Gamma> \<longrightarrow> \<Delta> \<turnstile> TyVar X <: T"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   615
  proof (intro strip)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   616
    assume ok: "\<turnstile> \<Delta> ok"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   617
    and    extends: "\<Delta> extends \<Gamma>"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   618
    have "(X,S) \<in> set \<Delta>" using lh_drv_prem extends by (simp add: extends_def)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   619
    moreover
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   620
    have "\<Delta> \<turnstile> S <: T" using ok extends ih by simp
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   621
    ultimately show "\<Delta> \<turnstile> TyVar X <: T" using ok by force
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   622
  qed
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   623
next
18305
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   624
  case (S_Refl \<Gamma> X)
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   625
  have lh_drv_prem: "X \<in> domain \<Gamma>" by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   626
  show ?case
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   627
  proof (intro strip)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   628
    assume "\<turnstile> \<Delta> ok"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   629
    moreover
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   630
    assume "\<Delta> extends \<Gamma>"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   631
    hence "X \<in> domain \<Delta>" using lh_drv_prem by (force dest: extends_domain)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   632
    ultimately show "\<Delta> \<turnstile> TyVar X <: TyVar X" by force
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   633
  qed
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   634
next 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   635
  case S_Arrow thus ?case by force
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   636
next
18305
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   637
  case (S_All \<Gamma> X S1 S2 T1 T2)
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   638
  have fresh: "X\<sharp>\<Delta>" by fact
18305
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   639
  have ih1: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<longrightarrow> \<Delta> extends \<Gamma> \<longrightarrow> \<Delta> \<turnstile> T1 <: S1" by fact
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   640
  have ih2: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<longrightarrow> \<Delta> extends ((X,T1)#\<Gamma>) \<longrightarrow> \<Delta> \<turnstile> S2 <: T2" by fact
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   641
  have lh_drv_prem: "\<Gamma> \<turnstile> T1 <: S1" by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   642
  hence b5: "T1 closed_in \<Gamma>" by (simp add: subtype_implies_closed) 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   643
  show "\<turnstile> \<Delta> ok \<longrightarrow> \<Delta> extends \<Gamma> \<longrightarrow> \<Delta> \<turnstile> \<forall> [X<:S1].S2 <: \<forall> [X<:T1].T2"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   644
  proof (intro strip)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   645
    assume ok: "\<turnstile> \<Delta> ok"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   646
    and    extends: "\<Delta> extends \<Gamma>"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   647
    have "T1 closed_in \<Delta>" using extends b5 by (simp only: extends_closed)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   648
    hence "\<turnstile> ((X,T1)#\<Delta>) ok" using fresh ok by force   
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   649
    moreover 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   650
    have "((X,T1)#\<Delta>) extends ((X,T1)#\<Gamma>)" using extends by (force simp add: extends_def)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   651
    ultimately have "((X,T1)#\<Delta>) \<turnstile> S2 <: T2" using ih2 by simp
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   652
    moreover
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   653
    have "\<Delta> \<turnstile> T1 <: S1" using ok extends ih1 by simp 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   654
    ultimately show "\<Delta> \<turnstile> \<forall> [X<:S1].S2 <: \<forall> [X<:T1].T2" using ok by (force intro: S_All) 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   655
  qed
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   656
qed
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   657
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   658
text {* more automated proof *}
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   659
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   660
lemma weakening:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   661
  assumes a1: "\<Gamma> \<turnstile> S <: T"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   662
  shows "\<turnstile> \<Delta> ok \<longrightarrow> \<Delta> extends \<Gamma> \<longrightarrow> \<Delta> \<turnstile> S <: T"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   663
  using a1 
18305
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   664
proof (nominal_induct \<Gamma> S T fresh: \<Delta> rule: subtype_of_rel_induct)
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   665
  case (S_Top \<Gamma> S) thus ?case by (force intro: extends_closed)
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   666
next 
18305
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   667
  case (S_Var \<Gamma> X S T) thus ?case by (force simp add: extends_def)
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   668
next
18305
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   669
  case (S_Refl \<Gamma> X) thus ?case by (force dest: extends_domain)
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   670
next 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   671
  case S_Arrow thus ?case by force
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   672
next
18305
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   673
  case (S_All \<Gamma> X S1 S2 T1 T2)
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   674
  have fresh: "X\<sharp>\<Delta>" by fact
18305
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   675
  have ih1: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<longrightarrow> \<Delta> extends \<Gamma> \<longrightarrow> \<Delta> \<turnstile> T1 <: S1" by fact
a780f9c1538b changed everything until the interesting transitivity_narrowing
urbanc
parents: 18269
diff changeset
   676
  have ih2: "\<And> \<Delta>. \<turnstile> \<Delta> ok \<longrightarrow> \<Delta> extends ((X,T1)#\<Gamma>) \<longrightarrow> \<Delta> \<turnstile> S2 <: T2" by fact
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   677
  have lh_drv_prem: "\<Gamma> \<turnstile> T1 <: S1" by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   678
  hence b5: "T1 closed_in \<Gamma>" by (simp add: subtype_implies_closed) 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   679
  show "\<turnstile> \<Delta> ok \<longrightarrow> \<Delta> extends \<Gamma> \<longrightarrow> \<Delta> \<turnstile> \<forall> [X<:S1].S2 <: \<forall> [X<:T1].T2"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   680
  proof (intro strip)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   681
    assume ok: "\<turnstile> \<Delta> ok"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   682
    and    extends: "\<Delta> extends \<Gamma>"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   683
    have "T1 closed_in \<Delta>" using extends b5 by (simp only: extends_closed)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   684
    hence "\<turnstile> ((X,T1)#\<Delta>) ok" using fresh ok by force   
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   685
    moreover 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   686
    have "((X,T1)#\<Delta>) extends ((X,T1)#\<Gamma>)" using extends by (force simp add: extends_def)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   687
    ultimately have "((X,T1)#\<Delta>) \<turnstile> S2 <: T2" using ih2 by simp
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   688
    moreover
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   689
    have "\<Delta> \<turnstile> T1 <: S1" using ok extends ih1 by simp 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   690
    ultimately show "\<Delta> \<turnstile> \<forall> [X<:S1].S2 <: \<forall> [X<:T1].T2" using ok by (force intro: S_All) 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   691
  qed
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   692
qed 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   693
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   694
(* helper lemma to calculate the measure of the induction *)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   695
lemma measure_eq [simp]: "(x, y) \<in> measure f = (f x < f y)"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   696
  by (simp add: measure_def inv_image_def)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   697
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   698
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   699
lemma transitivity_and_narrowing:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   700
  "(\<forall>\<Gamma> S T. \<Gamma> \<turnstile> S <: Q \<longrightarrow> \<Gamma> \<turnstile> Q <: T \<longrightarrow> \<Gamma> \<turnstile> S <: T) \<and>
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   701
   (\<forall>\<Delta> \<Gamma> X P M N. (\<Delta>@(X,Q)#\<Gamma>) \<turnstile> M <: N \<longrightarrow> \<Gamma> \<turnstile> P <: Q \<longrightarrow> (\<Delta>@(X,P)#\<Gamma>) \<turnstile> M <: N)"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   702
proof (rule wf_induct [of "measure size_ty" _ Q, rule_format])
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   703
  show "wf (measure size_ty)" by simp
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   704
next
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   705
  case (goal2 Q)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   706
  note  IH1_outer = goal2[THEN conjunct1]
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   707
    and IH2_outer = goal2[THEN conjunct2, THEN spec, of _ "[]", simplified]
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   708
  (* CHECK: Not clear whether the condition size_ty Q = size_ty Q' is needed, or whether
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   709
     just doing it for Q'=Q is enough *)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   710
  have transitivity: 
18306
a28269045ff0 started to change the transitivity/narrowing case:
urbanc
parents: 18305
diff changeset
   711
    "\<And>\<Gamma> S T. \<Gamma> \<turnstile> S <: Q \<Longrightarrow> \<Gamma> \<turnstile> Q <: T \<longrightarrow> \<Gamma> \<turnstile> S <: T"
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   712
  proof - 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   713
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   714
    (* We first handle the case where T = Top once and for all; this saves some 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   715
       repeated argument below (just like on paper :-).  We establish a little lemma
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   716
       that is invoked where needed in each case of the induction. *) 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   717
    have top_case: 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   718
      "\<And>\<Gamma> S T' P. \<lbrakk>\<turnstile> \<Gamma> ok; S closed_in \<Gamma>; P \<longrightarrow> \<Gamma> \<turnstile> S <: T'; T'=Top \<or> P\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> S <: T'"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   719
    proof - 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   720
      fix \<Gamma> S T' P
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   721
      assume S_Top_prem1: "S closed_in \<Gamma>"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   722
      and    S_Top_prem2: "\<turnstile> \<Gamma> ok"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   723
      and    alternative: "P \<longrightarrow> \<Gamma> \<turnstile> S <: T'" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   724
      and    "T'=Top \<or> P" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   725
      moreover
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   726
      { assume "T'=Top"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   727
	hence "\<Gamma> \<turnstile> S <: T'" using S_Top_prem1 S_Top_prem2 by (simp add: S_Top)
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
      moreover 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   730
      { assume P 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   731
	with alternative have "\<Gamma> \<turnstile> S <: T'" by simp 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   732
      }
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   733
      ultimately show "\<Gamma> \<turnstile> S <: T'" by blast
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   734
    qed
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   735
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   736
    (* Now proceed by induction on the left-hand derivation *)
18306
a28269045ff0 started to change the transitivity/narrowing case:
urbanc
parents: 18305
diff changeset
   737
    fix \<Gamma> S T
a28269045ff0 started to change the transitivity/narrowing case:
urbanc
parents: 18305
diff changeset
   738
    assume a: "\<Gamma> \<turnstile> S <: Q"
a28269045ff0 started to change the transitivity/narrowing case:
urbanc
parents: 18305
diff changeset
   739
    from a show "\<Gamma> \<turnstile> Q <: T \<longrightarrow> \<Gamma> \<turnstile> S <: T"
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   740
     (* FIXME : nominal induct does not work here because Gamma S and Q are fixed variables *)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   741
     (* FIX: It would be better to use S',\<Gamma>',etc., instead of S1,\<Gamma>1, for consistency, in the cases
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   742
        where these do not refer to sub-phrases of S, etc. *)
18306
a28269045ff0 started to change the transitivity/narrowing case:
urbanc
parents: 18305
diff changeset
   743
    proof(nominal_induct \<Gamma> S Q\<equiv>Q fresh: \<Gamma> S T rule: subtype_of_rel_induct)
a28269045ff0 started to change the transitivity/narrowing case:
urbanc
parents: 18305
diff changeset
   744
    (* interesting case *)
a28269045ff0 started to change the transitivity/narrowing case:
urbanc
parents: 18305
diff changeset
   745
a28269045ff0 started to change the transitivity/narrowing case:
urbanc
parents: 18305
diff changeset
   746
      case (goal1 \<Gamma>1 S1)    --"lh drv.: \<Gamma> \<turnstile> S <: Q' \<equiv> \<Gamma>1 \<turnstile> S1 <: Top"
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   747
	--"automatic proof: thus ?case proof (auto intro!: S_Top dest: S_TopE)"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   748
      assume lh_drv_prem1: "\<turnstile> \<Gamma>1 ok"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   749
      and    lh_drv_prem2: "S1 closed_in \<Gamma>1"
18306
a28269045ff0 started to change the transitivity/narrowing case:
urbanc
parents: 18305
diff changeset
   750
      and    Q_eq: "Top=Q" 
a28269045ff0 started to change the transitivity/narrowing case:
urbanc
parents: 18305
diff changeset
   751
      show "\<Gamma>1 \<turnstile> Q <: T \<longrightarrow> \<Gamma>1 \<turnstile> S1 <: T" (* FIXME: why Ta? *)
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   752
      proof (intro strip)
18306
a28269045ff0 started to change the transitivity/narrowing case:
urbanc
parents: 18305
diff changeset
   753
        assume "\<Gamma>1 \<turnstile> Q <: T" --"rh drv." 
a28269045ff0 started to change the transitivity/narrowing case:
urbanc
parents: 18305
diff changeset
   754
        hence "T = Top" using Q_eq by (simp add: S_TopE)
a28269045ff0 started to change the transitivity/narrowing case:
urbanc
parents: 18305
diff changeset
   755
        thus "\<Gamma>1 \<turnstile> S1 <: T" using top_case[of "\<Gamma>1" "S1" "False" "T"] lh_drv_prem1 lh_drv_prem2
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   756
          by simp
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   757
      qed
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   758
    next
18306
a28269045ff0 started to change the transitivity/narrowing case:
urbanc
parents: 18305
diff changeset
   759
      case (goal2 \<Gamma>1 X1 S1 T1)     --"lh drv.: \<Gamma> \<turnstile> S <: Q' \<equiv> \<Gamma>1 \<turnstile> TVar(X1) <: S1"
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   760
        --"automatic proof: thus ?case proof (auto intro!: S_Var)"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   761
      have lh_drv_prem1: "\<turnstile> \<Gamma>1 ok" by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   762
      have lh_drv_prem2: "(X1,S1)\<in>set \<Gamma>1" by fact
18306
a28269045ff0 started to change the transitivity/narrowing case:
urbanc
parents: 18305
diff changeset
   763
      have IH_inner:    "\<And>T. \<Gamma>1 \<turnstile> T1 <: T \<longrightarrow> \<Gamma>1 \<turnstile> S1 <: T" by fact
a28269045ff0 started to change the transitivity/narrowing case:
urbanc
parents: 18305
diff changeset
   764
      show "\<Gamma>1 \<turnstile> T1 <: Ta \<longrightarrow> \<Gamma>1 \<turnstile> TyVar X1 <: Ta"
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   765
      proof (intro strip)
18306
a28269045ff0 started to change the transitivity/narrowing case:
urbanc
parents: 18305
diff changeset
   766
        assume "\<Gamma>1 \<turnstile> T1 <: Ta" --"right-hand drv."
a28269045ff0 started to change the transitivity/narrowing case:
urbanc
parents: 18305
diff changeset
   767
        with IH_inner have "\<Gamma>1 \<turnstile> S1 <: Ta" by simp
a28269045ff0 started to change the transitivity/narrowing case:
urbanc
parents: 18305
diff changeset
   768
        thus "\<Gamma>1 \<turnstile> TyVar X1 <: Ta" using lh_drv_prem1 lh_drv_prem2 by (simp add: S_Var)
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   769
      qed
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   770
    next
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   771
      case goal3 --"S_Refl case" show ?case by simp
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   772
    next
18306
a28269045ff0 started to change the transitivity/narrowing case:
urbanc
parents: 18305
diff changeset
   773
      case (goal4 \<Gamma>1 S1 S2 T1 T2) --"lh drv.: \<Gamma> \<turnstile> S <: Q' == \<Gamma>1 \<turnstile> S1 \<rightarrow> S2 <: T1 \<rightarrow> T2"
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   774
      have lh_drv_prem1: "\<Gamma>1 \<turnstile> T1 <: S1" by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   775
      have lh_drv_prem2: "\<Gamma>1 \<turnstile> S2 <: T2" by fact
18306
a28269045ff0 started to change the transitivity/narrowing case:
urbanc
parents: 18305
diff changeset
   776
      show "\<Gamma>1 \<turnstile> T1 \<rightarrow> T2 <: T \<longrightarrow> \<Gamma>1 \<turnstile> S1 \<rightarrow> S2 <: T"
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   777
      proof (intro strip)
18306
a28269045ff0 started to change the transitivity/narrowing case:
urbanc
parents: 18305
diff changeset
   778
        (*assume measure:  "size_ty Q = size_ty (T1 \<rightarrow> T2)" *)
a28269045ff0 started to change the transitivity/narrowing case:
urbanc
parents: 18305
diff changeset
   779
        assume  rh_deriv: "\<Gamma>1 \<turnstile> T1 \<rightarrow> T2 <: T'"
18246
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   780
	have "S1 closed_in \<Gamma>1" and "S2 closed_in \<Gamma>1" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   781
	  using lh_drv_prem1 lh_drv_prem2 by (simp_all add: subtype_implies_closed)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   782
        hence "(S1 \<rightarrow> S2) closed_in \<Gamma>1" by (simp add: closed_in_def ty.supp)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   783
        moreover
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   784
	have "\<turnstile> \<Gamma>1 ok" using rh_deriv by (rule subtype_implies_ok)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   785
        moreover
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   786
	have "T' = Top \<or> (\<exists>T3 T4.  T'= T3 \<rightarrow> T4 \<and> \<Gamma>1 \<turnstile> T3 <: T1 \<and> \<Gamma>1 \<turnstile> T2 <: T4)" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   787
	  using rh_deriv by (rule S_ArrowE_left)  
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   788
	moreover
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   789
	{ assume "\<exists>T3 T4. T'= T3 \<rightarrow> T4 \<and> \<Gamma>1 \<turnstile> T3 <: T1 \<and> \<Gamma>1 \<turnstile> T2 <: T4"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   790
	  then obtain T3 T4 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   791
	    where T'_inst: "T'= T3 \<rightarrow> T4" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   792
	    and   rh_drv_prem1: "\<Gamma>1 \<turnstile> T3 <: T1"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   793
	    and   rh_drv_prem2: "\<Gamma>1 \<turnstile> T2 <: T4" by force
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   794
          from IH1_outer[of "T1"] have "\<Gamma>1 \<turnstile> T3 <: S1" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   795
	    using measure rh_drv_prem1 lh_drv_prem1 by (force simp add: ty.inject)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   796
	  moreover
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   797
	  from IH1_outer[of "T2"] have "\<Gamma>1 \<turnstile> S2 <: T4" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   798
	    using measure rh_drv_prem2 lh_drv_prem2 by (force simp add: ty.inject)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   799
	  ultimately have "\<Gamma>1 \<turnstile> S1 \<rightarrow> S2 <: T'" using T'_inst by force
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   800
	}
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   801
	ultimately show "\<Gamma>1 \<turnstile> S1 \<rightarrow> S2 <: T'" using top_case[of "\<Gamma>1" "S1\<rightarrow>S2" _ "T'"] by blast
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   802
      qed
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   803
    next
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   804
      case (goal5 T' \<Gamma>1 X S1 S2 T1 T2) --"lh drv.: \<Gamma> \<turnstile> S <: Q' \<equiv> \<Gamma>1 \<turnstile> \<forall>[X:S1].S2 <: \<forall>[X:T1].T2" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   805
      have lh_drv_prem1: "\<Gamma>1 \<turnstile> T1 <: S1" by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   806
      have lh_drv_prem2: "((X,T1)#\<Gamma>1) \<turnstile> S2 <: T2" by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   807
      have fresh_cond: "X\<sharp>(\<Gamma>1,S1,T1)" by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   808
      show "size_ty Q = size_ty (\<forall>[X<:T1].T2) \<longrightarrow> \<Gamma>1 \<turnstile> \<forall>[X<:T1].T2 <: T' \<longrightarrow> \<Gamma>1 \<turnstile> \<forall>[X<:S1].S2 <: T'"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   809
      proof (intro strip)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   810
        assume measure: "size_ty Q = size_ty (\<forall>[X<:T1].T2)"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   811
        and    rh_deriv: "\<Gamma>1 \<turnstile> \<forall>[X<:T1].T2 <: T'"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   812
        have "S1 closed_in \<Gamma>1" and "S2 closed_in ((X,T1)#\<Gamma>1)" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   813
	  using lh_drv_prem1 lh_drv_prem2 by (simp_all add: subtype_implies_closed)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   814
	hence "(\<forall>[X<:S1].S2) closed_in \<Gamma>1" by (force simp add: closed_in_def ty.supp abs_supp)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   815
	moreover
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   816
	have "\<turnstile> \<Gamma>1 ok" using rh_deriv by (rule subtype_implies_ok)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   817
	moreover
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   818
        (* FIX: Maybe T3,T4 could be T1',T2' *)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   819
	have "T' = Top \<or> (\<exists>T3 T4. T'=\<forall>[X<:T3].T4 \<and> \<Gamma>1 \<turnstile> T3 <: T1 \<and> ((X,T3)#\<Gamma>1) \<turnstile> T2 <: T4)" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   820
	  using rh_deriv fresh_cond by (auto dest: S_AllE_left simp add: fresh_prod)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   821
        moreover
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   822
        { assume "\<exists>T3 T4. T'=\<forall>[X<:T3].T4 \<and> \<Gamma>1 \<turnstile> T3 <: T1 \<and> ((X,T3)#\<Gamma>1) \<turnstile> T2 <: T4"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   823
	  then obtain T3 T4 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   824
	    where T'_inst: "T'= \<forall>[X<:T3].T4" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   825
	    and   rh_drv_prem1: "\<Gamma>1 \<turnstile> T3 <: T1" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   826
	    and   rh_drv_prem2:"((X,T3)#\<Gamma>1) \<turnstile> T2 <: T4" by force
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   827
          from IH1_outer[of "T1"] have "\<Gamma>1 \<turnstile> T3 <: S1" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   828
	    using lh_drv_prem1 rh_drv_prem1 measure by (force simp add: ty.inject)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   829
          moreover
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   830
	  from IH2_outer[of "T1"] have "((X,T3)#\<Gamma>1) \<turnstile> S2 <: T2" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   831
	    using measure lh_drv_prem2 rh_drv_prem1 by force
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   832
	  with IH1_outer[of "T2"] have "((X,T3)#\<Gamma>1) \<turnstile> S2 <: T4" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   833
	    using measure rh_drv_prem2 by force
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   834
          moreover
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   835
	  ultimately have "\<Gamma>1 \<turnstile> \<forall>[X<:S1].S2 <: T'" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   836
	    using fresh_cond T'_inst by (simp add: fresh_prod S_All)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   837
        }
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   838
        ultimately show "\<Gamma>1 \<turnstile> \<forall>[X<:S1].S2 <: T'" using top_case[of "\<Gamma>1" "\<forall>[X<:S1].S2" _ "T'"] by blast
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   839
      qed
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   840
    qed
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   841
  qed
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   842
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   843
(* test
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   844
  have narrowing:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   845
    "\<And>\<Delta> \<Gamma> X M N. (\<Delta>@(X,Q)#\<Gamma>) \<turnstile> M <: N \<Longrightarrow> (\<forall>P. \<Gamma> \<turnstile> P<:Q \<longrightarrow> (\<Delta>@(X,P)#\<Gamma>) \<turnstile> M <: N)"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   846
  proof -
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   847
    fix \<Delta> \<Gamma> X M N
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   848
    assume a: "(\<Delta>@(X,Q)#\<Gamma>) \<turnstile> M <: N"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   849
    thus "\<forall>P. \<Gamma> \<turnstile> P <: Q \<longrightarrow> (\<Delta>@(X,P)#\<Gamma>) \<turnstile> M <: N"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   850
      thm subtype_of_rel_induct
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   851
      thm subtype_of_rel.induct
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   852
      using a proof (induct)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   853
      using a proof (induct rule: subtype_of_rel_induct[of "\<Delta>@(X,Q)#\<Gamma>" "M" "N" _ "()"])
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   854
*)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   855
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   856
  have narrowing:
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   857
    "\<And>\<Delta> \<Gamma> \<Gamma>' X M N. \<Gamma>' \<turnstile> M <: N \<Longrightarrow> \<Gamma>' = \<Delta>@(X,Q)#\<Gamma> \<longrightarrow> (\<forall>P. \<Gamma> \<turnstile> P<:Q \<longrightarrow> (\<Delta>@(X,P)#\<Gamma>) \<turnstile> M <: N)"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   858
  proof -
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   859
    fix \<Delta> \<Gamma> \<Gamma>' X M N
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   860
    assume "\<Gamma>' \<turnstile> M <: N"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   861
    thus "\<Gamma>' = \<Delta>@(X,Q)#\<Gamma> \<longrightarrow> (\<forall>P. \<Gamma> \<turnstile> P <: Q \<longrightarrow> (\<Delta>@(X,P)#\<Gamma>) \<turnstile> M <: N)"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   862
      (* FIXME : nominal induct does not work here because Gamma' M and N are fixed variables *)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   863
      (* FIX: Same comment about S1,\<Gamma>1 *)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   864
    proof (rule subtype_of_rel_induct[of "\<Gamma>'" "M" "N" "\<lambda>\<Gamma>' M N (\<Delta>,\<Gamma>,X). 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   865
	\<Gamma>' = \<Delta>@(X,Q)#\<Gamma> \<longrightarrow> (\<forall>P. \<Gamma> \<turnstile> P <: Q \<longrightarrow> (\<Delta>@(X,P)#\<Gamma>) \<turnstile> M <: N)" "(\<Delta>,\<Gamma>,X)", simplified], 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   866
	simp_all only: split_paired_all split_conv)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   867
      case (goal1 \<Delta>1 \<Gamma>1 X1 \<Gamma>2 S1)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   868
      have lh_drv_prem1: "\<turnstile> \<Gamma>2 ok" by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   869
      have lh_drv_prem2: "S1 closed_in \<Gamma>2" by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   870
      show "\<Gamma>2 = \<Delta>1@(X1,Q)#\<Gamma>1 \<longrightarrow> (\<forall>P. \<Gamma>1 \<turnstile> P <: Q \<longrightarrow> (\<Delta>1@(X1,P)#\<Gamma>1) \<turnstile> S1 <: Top)"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   871
      proof (intro strip)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   872
	fix P
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   873
	assume a1: "\<Gamma>2 = \<Delta>1@(X1,Q)#\<Gamma>1"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   874
	assume a2: "\<Gamma>1 \<turnstile> P <: Q"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   875
	from a2 have "P closed_in \<Gamma>1" by (simp add: subtype_implies_closed)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   876
	hence a3: "\<turnstile> (\<Delta>1@(X1,P)#\<Gamma>1) ok" using lh_drv_prem1 a1 by (rule_tac replace_type, simp_all) 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   877
	show "(\<Delta>1@(X1,P)#\<Gamma>1) \<turnstile> S1 <: Top"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   878
	  using a1 a3 lh_drv_prem2 by (force simp add: closed_in_def domain_append)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   879
      qed
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   880
    next
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   881
      case (goal2 \<Delta>1 \<Gamma>1 X1 \<Gamma>2 X2 S1 T1)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   882
      have lh_drv_prem1: "\<turnstile> \<Gamma>2 ok" by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   883
      have lh_drv_prem2: "(X2,S1)\<in>set \<Gamma>2" by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   884
      have lh_drv_prem3: "\<Gamma>2 \<turnstile> S1 <: T1" by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   885
      have IH_inner: 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   886
	"\<forall>\<Delta>1 \<Gamma>1 X1.  \<Gamma>2 = \<Delta>1@(X1,Q)#\<Gamma>1 \<longrightarrow> (\<forall>P. \<Gamma>1 \<turnstile> P <: Q \<longrightarrow> (\<Delta>1@(X1,P)#\<Gamma>1) \<turnstile> S1 <: T1)" by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   887
      show "\<Gamma>2 = (\<Delta>1@(X1,Q)#\<Gamma>1) \<longrightarrow> (\<forall>P. \<Gamma>1 \<turnstile> P <: Q \<longrightarrow> (\<Delta>1@(X1,P)#\<Gamma>1) \<turnstile> TyVar X2 <: T1)"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   888
      proof (intro strip)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   889
	fix P
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   890
        assume a1: "\<Gamma>2 = \<Delta>1@(X1,Q)#\<Gamma>1"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   891
	and    a2: "\<Gamma>1 \<turnstile> P <: Q"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   892
	from a2 have "P closed_in \<Gamma>1" by (simp add: subtype_implies_closed)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   893
	hence a3: "\<turnstile> (\<Delta>1@(X1,P)#\<Gamma>1) ok" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   894
	  using lh_drv_prem1 a1 by (rule_tac replace_type, simp_all)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   895
	show "(\<Delta>1@(X1,P)#\<Gamma>1) \<turnstile> TyVar X2 <: T1"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   896
 	proof (cases "X1=X2")
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   897
	  case False
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   898
	  have b0: "X1\<noteq>X2" by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   899
	  from lh_drv_prem3 a1 a2 IH_inner 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   900
	  have b1: "(\<Delta>1@(X1,P)#\<Gamma>1) \<turnstile> S1 <: T1" by simp
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   901
	  from lh_drv_prem2 a1 b0 have b2: "(X2,S1)\<in>set (\<Delta>1@(X1,P)#\<Gamma>1)" by simp
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   902
	  show "(\<Delta>1@(X1,P)#\<Gamma>1) \<turnstile> TyVar X2 <: T1" using a3 b2 b1 by (rule S_Var)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   903
	next 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   904
	  case True
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   905
	  have b0: "X1=X2" by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   906
	  have a4: "S1=Q"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   907
	  proof -
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   908
	    have c0: "(X2,Q)\<in>set \<Gamma>2" using b0 a1 by simp
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   909
	    with lh_drv_prem1 lh_drv_prem2 show "S1=Q" by (simp add: uniqueness_of_ctxt)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   910
	  qed
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   911
	  have a5: "(\<Delta>1@(X1,P)#\<Gamma>1) extends \<Gamma>1" by (force simp add: extends_def)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   912
	  have a6: "(\<Delta>1@(X2,P)#\<Gamma>1) \<turnstile> P <: Q" using b0 a5 a2 a3 by (simp add: weakening)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   913
	  have a7: "(\<Delta>1@(X2,P)#\<Gamma>1) \<turnstile> Q <: T1" using b0 IH_inner a1 lh_drv_prem3 a2 a4 by simp
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   914
	  have a8: "(\<Delta>1@(X2,P)#\<Gamma>1) \<turnstile> P <: T1" using a6 a7 transitivity by blast
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   915
	  have a9: "(X2,P)\<in>set (\<Delta>1@(X1,P)#\<Gamma>1)" using b0 by simp
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   916
	  show "(\<Delta>1@(X1,P)#\<Gamma>1) \<turnstile> TyVar X2 <: T1" using a3 b0 a8 a9 by (force simp add: S_Var)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   917
	qed
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   918
      qed
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   919
    next
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   920
      case (goal3 \<Delta>1 \<Gamma>1 X1 \<Gamma>2 X2)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   921
      have lh_drv_prem1: "\<turnstile> \<Gamma>2 ok" by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   922
      have lh_drv_prem2: "X2 \<in> domain \<Gamma>2" by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   923
      show "\<Gamma>2 = (\<Delta>1@(X1,Q)#\<Gamma>1) \<longrightarrow> (\<forall>P. \<Gamma>1 \<turnstile> P <: Q \<longrightarrow> (\<Delta>1@(X1,P)#\<Gamma>1) \<turnstile> TyVar X2 <: TyVar X2)"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   924
      proof (intro strip)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   925
	fix P
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   926
	assume a1: "\<Gamma>2 = \<Delta>1@(X1,Q)#\<Gamma>1"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   927
	and    a2: "\<Gamma>1 \<turnstile> P <: Q"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   928
	from a2 have "P closed_in \<Gamma>1" by (simp add: subtype_implies_closed)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   929
	hence a3: "\<turnstile> (\<Delta>1@(X1,P)#\<Gamma>1) ok" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   930
	  using lh_drv_prem1 a1 by (rule_tac replace_type, simp_all)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   931
	have b0: "X2 \<in> domain (\<Delta>1@(X1,P)#\<Gamma>1)" using lh_drv_prem2 a1 by (simp add: domain_append)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   932
	show "(\<Delta>1@(X1,P)#\<Gamma>1) \<turnstile> TyVar X2 <: TyVar X2" using a3 b0 by (rule S_Refl)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   933
      qed
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   934
    next
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   935
      case goal4 thus ?case by blast
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   936
    next
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   937
      case (goal5 \<Delta>1 \<Gamma>1 X1 \<Gamma>2 X2 S1 S2 T1 T2)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   938
      have IH_inner1: 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   939
	"\<forall>\<Delta>1 \<Gamma>1 X1.  \<Gamma>2 = \<Delta>1@(X1,Q)#\<Gamma>1 \<longrightarrow> (\<forall>P. \<Gamma>1 \<turnstile> P <: Q \<longrightarrow> (\<Delta>1@(X1,P)#\<Gamma>1) \<turnstile> T1 <: S1)" by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   940
      have IH_inner2: 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   941
	"\<forall>\<Delta>1 \<Gamma>1 X1. (X2,T1)#\<Gamma>2 = \<Delta>1@(X1,Q)#\<Gamma>1 \<longrightarrow> (\<forall>P. \<Gamma>1 \<turnstile> P <: Q \<longrightarrow> (\<Delta>1@(X1,P)#\<Gamma>1) \<turnstile> S2 <: T2)" 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   942
	by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   943
      have lh_drv_prem1: "\<Gamma>2 \<turnstile> T1 <: S1" by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   944
      have lh_drv_prem2: "X2 \<sharp> (\<Gamma>2,S1, T1)" by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   945
      have lh_drv_prem3: "((X2,T1) # \<Gamma>2) \<turnstile> S2 <: T2" by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   946
      have freshness: "X2\<sharp>(\<Delta>1, \<Gamma>1, X1)" by fact
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   947
      show "\<Gamma>2 = (\<Delta>1@(X1,Q)#\<Gamma>1) \<longrightarrow> 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   948
            (\<forall>P. \<Gamma>1 \<turnstile> P <: Q \<longrightarrow> (\<Delta>1@(X1,P)#\<Gamma>1) \<turnstile> \<forall> [X2<:S1].S2 <: \<forall> [X2<:T1].T2)"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   949
      proof (intro strip)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   950
	fix P
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   951
	assume a1: "\<Gamma>2 = \<Delta>1@(X1,Q)#\<Gamma>1"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   952
	and    a2: "\<Gamma>1 \<turnstile> P <: Q"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   953
	have b0: "(\<Delta>1@(X1,P)#\<Gamma>1) \<turnstile> T1 <: S1" using a1 a2 lh_drv_prem1 IH_inner1 by simp
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   954
	have b1: "(((X2,T1)#\<Delta>1)@(X1,P)#\<Gamma>1) \<turnstile> S2 <: T2" using a1 a2 lh_drv_prem3 IH_inner2
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   955
	  apply auto
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   956
	  apply (drule_tac x="(X2,T1)#\<Delta>1" in spec)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   957
	  apply(simp)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   958
	  done
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   959
	have b3: "X2\<sharp>(\<Delta>1@(X1,P)#\<Gamma>1)" using lh_drv_prem2 freshness a2
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   960
	  by (auto simp add: fresh_prod fresh_list_append fresh_list_cons dest: subtype_implies_fresh)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   961
	show "(\<Delta>1@(X1,P)#\<Gamma>1) \<turnstile> \<forall> [X2<:S1].S2 <: \<forall> [X2<:T1].T2" using b0 b3 b1 by force 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   962
      qed
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   963
    qed
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   964
  qed
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   965
  from transitivity narrowing show ?case by blast 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   966
qed
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   967
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   968
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   969
 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   970
section {* Terms *}
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   971
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   972
nominal_datatype trm = Var   "vrs"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   973
                     | Lam   "\<guillemotleft>vrs\<guillemotright>trm" "ty"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   974
                     | TAbs  "\<guillemotleft>tyvrs\<guillemotright>trm" "ty"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   975
                     | App   "trm" "trm"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   976
                     | TApp  "trm" "ty"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   977
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   978
consts
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   979
  Lam_syn   :: "vrs \<Rightarrow> ty \<Rightarrow> trm \<Rightarrow> trm"   ("Lam [_:_]._" [100,100,100] 100)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   980
  TAbs_syn  :: "tyvrs \<Rightarrow> ty \<Rightarrow> trm \<Rightarrow> trm" ("TAbs [_<:_]._" [100,100,100] 100)
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   981
translations 
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   982
  "Lam  [a:tys].t" \<rightleftharpoons> "trm.Lam a t tys"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   983
  "TAbs [a<:tys].t" \<rightleftharpoons> "trm.TAbs a t tys"
676d2e625d98 added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
diff changeset
   984