src/HOL/Nominal/Examples/W.thy
author haftmann
Fri, 27 Jun 2025 08:09:26 +0200
changeset 82775 61c39a9e5415
parent 82664 e9f3b94eb6a0
permissions -rw-r--r--
typo
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
26195
8292f8723e99 added new example
urbanc
parents:
diff changeset
     1
theory W
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 63167
diff changeset
     2
imports "HOL-Nominal.Nominal"
26195
8292f8723e99 added new example
urbanc
parents:
diff changeset
     3
begin
8292f8723e99 added new example
urbanc
parents:
diff changeset
     4
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58219
diff changeset
     5
text \<open>Example for strong induction rules avoiding sets of atoms.\<close>
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
     6
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
     7
atom_decl tvar var 
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
     8
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
     9
abbreviation
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80149
diff changeset
    10
  "difference_list" :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (\<open>_ - _\<close> [60,60] 60) 
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
    11
where
68386
98cf1c823c48 Keep filter input syntax
nipkow
parents: 68249
diff changeset
    12
  "xs - ys \<equiv> [x \<leftarrow> xs. x\<notin>set ys]"
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
    13
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
    14
lemma difference_eqvt_tvar[eqvt]:
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
    15
  fixes pi::"tvar prm"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
    16
    and   Xs Ys::"tvar list"
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
    17
  shows "pi\<bullet>(Xs - Ys) = (pi\<bullet>Xs) - (pi\<bullet>Ys)"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
    18
  by (induct Xs) (simp_all add: eqvts)
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
    19
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
    20
lemma difference_fresh [simp]:
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
    21
  fixes X::"tvar"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
    22
    and   Xs Ys::"tvar list"
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
    23
  shows "X\<sharp>(Xs - Ys) \<longleftrightarrow> X\<sharp>Xs \<or> X\<in>set Ys"
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
    24
  by (induct Xs) (auto simp add: fresh_list_nil fresh_list_cons fresh_atm)
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
    25
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    26
lemma difference_supset:
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    27
  fixes xs::"'a list"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    28
  and   ys::"'a list"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    29
  and   zs::"'a list"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    30
  assumes asm: "set xs \<subseteq> set ys"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    31
  shows "xs - ys = []"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    32
using asm
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    33
by (induct xs) (auto)
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    34
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
    35
nominal_datatype ty = 
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
    36
    TVar "tvar"
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80149
diff changeset
    37
  | Fun "ty" "ty" (\<open>_\<rightarrow>_\<close> [100,100] 100)
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
    38
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
    39
nominal_datatype tyS = 
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
    40
    Ty  "ty"
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80149
diff changeset
    41
  | ALL "\<guillemotleft>tvar\<guillemotright>tyS" (\<open>\<forall>[_]._\<close> [100,100] 100)
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
    42
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
    43
nominal_datatype trm = 
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
    44
    Var "var"
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
    45
  | App "trm" "trm" 
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80149
diff changeset
    46
  | Lam "\<guillemotleft>var\<guillemotright>trm" (\<open>Lam [_]._\<close> [100,100] 100)
58219
61b852f90161 improved 'datatype_compat' further for recursion through functions
blanchet
parents: 55417
diff changeset
    47
  | Let "\<guillemotleft>var\<guillemotright>trm" "trm"
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
    48
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
    49
abbreviation
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80149
diff changeset
    50
  LetBe :: "var \<Rightarrow> trm \<Rightarrow> trm \<Rightarrow> trm" (\<open>Let _ be _ in _\<close> [100,100,100] 100)
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
    51
where
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
    52
 "Let x be t1 in t2 \<equiv> trm.Let x t2 t1"
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
    53
41798
c3aa3c87ef21 modernized specifications;
wenzelm
parents: 39246
diff changeset
    54
type_synonym 
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
    55
  Ctxt  = "(var\<times>tyS) list" 
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
    56
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58219
diff changeset
    57
text \<open>free type variables\<close>
29098
6b23a58cc67c Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents: 28655
diff changeset
    58
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    59
consts ftv :: "'a \<Rightarrow> tvar list"
29098
6b23a58cc67c Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents: 28655
diff changeset
    60
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    61
overloading 
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    62
  ftv_prod \<equiv> "ftv :: ('a \<times> 'b) \<Rightarrow> tvar list"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    63
  ftv_tvar \<equiv> "ftv :: tvar \<Rightarrow> tvar list"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    64
  ftv_var  \<equiv> "ftv :: var \<Rightarrow> tvar list"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    65
  ftv_list \<equiv> "ftv :: 'a list \<Rightarrow> tvar list"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    66
  ftv_ty   \<equiv> "ftv :: ty \<Rightarrow> tvar list"
29098
6b23a58cc67c Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents: 28655
diff changeset
    67
begin
6b23a58cc67c Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents: 28655
diff changeset
    68
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    69
primrec 
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    70
  ftv_prod
29098
6b23a58cc67c Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents: 28655
diff changeset
    71
where
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    72
 "ftv_prod (x, y) = (ftv x) @ (ftv y)"
29098
6b23a58cc67c Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents: 28655
diff changeset
    73
6b23a58cc67c Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents: 28655
diff changeset
    74
definition
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    75
  ftv_tvar :: "tvar \<Rightarrow> tvar list"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    76
where
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    77
[simp]: "ftv_tvar X \<equiv> [(X::tvar)]"
29098
6b23a58cc67c Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents: 28655
diff changeset
    78
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    79
definition
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    80
  ftv_var :: "var \<Rightarrow> tvar list"
29098
6b23a58cc67c Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents: 28655
diff changeset
    81
where
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    82
[simp]: "ftv_var x \<equiv> []"
29098
6b23a58cc67c Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents: 28655
diff changeset
    83
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    84
primrec 
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    85
  ftv_list
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    86
where
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    87
  "ftv_list [] = []"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    88
| "ftv_list (x#xs) = (ftv x)@(ftv_list xs)"
29098
6b23a58cc67c Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents: 28655
diff changeset
    89
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    90
nominal_primrec 
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    91
  ftv_ty :: "ty \<Rightarrow> tvar list"
29098
6b23a58cc67c Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents: 28655
diff changeset
    92
where
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    93
  "ftv_ty (TVar X) = [X]"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
    94
| "ftv_ty (T1 \<rightarrow>T2) = (ftv_ty T1) @ (ftv_ty T2)"
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
    95
by (rule TrueI)+
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
    96
29098
6b23a58cc67c Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents: 28655
diff changeset
    97
end
6b23a58cc67c Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents: 28655
diff changeset
    98
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
    99
lemma ftv_ty_eqvt[eqvt]:
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   100
  fixes pi::"tvar prm"
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   101
  and   T::"ty"
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   102
  shows "pi\<bullet>(ftv T) = ftv (pi\<bullet>T)"
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   103
by (nominal_induct T rule: ty.strong_induct)
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   104
   (perm_simp add: append_eqvt)+
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   105
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   106
overloading 
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   107
  ftv_tyS  \<equiv> "ftv :: tyS \<Rightarrow> tvar list"
29098
6b23a58cc67c Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents: 28655
diff changeset
   108
begin
6b23a58cc67c Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents: 28655
diff changeset
   109
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   110
nominal_primrec 
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   111
  ftv_tyS :: "tyS \<Rightarrow> tvar list"
29098
6b23a58cc67c Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents: 28655
diff changeset
   112
where
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   113
  "ftv_tyS (Ty T)    = ((ftv (T::ty))::tvar list)"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   114
| "ftv_tyS (\<forall>[X].S) = (ftv_tyS S) - [X]"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   115
apply(finite_guess add: ftv_ty_eqvt fs_tvar1 | rule TrueI)+
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   116
  apply (metis difference_fresh list.set_intros(1))
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   117
apply(fresh_guess add: ftv_ty_eqvt fs_tvar1)+
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   118
done
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   119
29098
6b23a58cc67c Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents: 28655
diff changeset
   120
end
6b23a58cc67c Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents: 28655
diff changeset
   121
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   122
lemma ftv_tyS_eqvt[eqvt]:
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   123
  fixes pi::"tvar prm"
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   124
  and   S::"tyS"
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   125
  shows "pi\<bullet>(ftv S) = ftv (pi\<bullet>S)"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   126
proof (nominal_induct S rule: tyS.strong_induct)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   127
  case (Ty ty)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   128
  then show ?case
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   129
    by (simp add: ftv_ty_eqvt)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   130
next
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   131
  case (ALL tvar tyS)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   132
  then show ?case 
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   133
    by (metis difference_eqvt_tvar ftv_ty.simps(1) ftv_tyS.simps(2) ftv_ty_eqvt ty.perm(3) tyS.perm(4))
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   134
qed
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   135
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   136
lemma ftv_Ctxt_eqvt[eqvt]:
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   137
  fixes pi::"tvar prm"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   138
    and   \<Gamma>::"Ctxt"
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   139
  shows "pi\<bullet>(ftv \<Gamma>) = ftv (pi\<bullet>\<Gamma>)"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   140
  by (induct \<Gamma>) (auto simp add: eqvts)
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   141
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58219
diff changeset
   142
text \<open>Valid\<close>
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   143
inductive
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   144
  valid :: "Ctxt \<Rightarrow> bool"
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   145
where
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   146
  V_Nil[intro]:  "valid []"
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   147
| V_Cons[intro]: "\<lbrakk>valid \<Gamma>;x\<sharp>\<Gamma>\<rbrakk>\<Longrightarrow> valid ((x,S)#\<Gamma>)"
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   148
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   149
equivariance valid
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   150
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58219
diff changeset
   151
text \<open>General\<close>
39246
9e58f0499f57 modernized primrec
haftmann
parents: 32137
diff changeset
   152
primrec gen :: "ty \<Rightarrow> tvar list \<Rightarrow> tyS" where
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   153
  "gen T [] = Ty T"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 32137
diff changeset
   154
| "gen T (X#Xs) = \<forall>[X].(gen T Xs)"
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   155
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   156
lemma gen_eqvt[eqvt]:
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   157
  fixes pi::"tvar prm"
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   158
  shows "pi\<bullet>(gen T Xs) = gen (pi\<bullet>T) (pi\<bullet>Xs)"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   159
  by (induct Xs) (simp_all add: eqvts)
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   160
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   161
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   162
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   163
abbreviation 
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   164
  close :: "Ctxt \<Rightarrow> ty \<Rightarrow> tyS"
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   165
where 
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   166
  "close \<Gamma> T \<equiv> gen T ((ftv T) - (ftv \<Gamma>))"
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   167
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   168
lemma close_eqvt[eqvt]:
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   169
  fixes pi::"tvar prm"
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   170
  shows "pi\<bullet>(close \<Gamma> T) = close (pi\<bullet>\<Gamma>) (pi\<bullet>T)"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   171
  by (simp_all only: eqvts)
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   172
  
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58219
diff changeset
   173
text \<open>Substitution\<close>
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   174
41798
c3aa3c87ef21 modernized specifications;
wenzelm
parents: 39246
diff changeset
   175
type_synonym Subst = "(tvar\<times>ty) list"
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   176
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   177
consts
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80149
diff changeset
   178
  psubst :: "Subst \<Rightarrow> 'a \<Rightarrow> 'a"       (\<open>_<_>\<close> [100,60] 120)
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   179
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   180
abbreviation 
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80149
diff changeset
   181
  subst :: "'a \<Rightarrow> tvar \<Rightarrow> ty \<Rightarrow> 'a"  (\<open>_[_::=_]\<close> [100,100,100] 100)
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   182
where
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   183
  "smth[X::=T] \<equiv> ([(X,T)])<smth>" 
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   184
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   185
fun lookup :: "Subst \<Rightarrow> tvar \<Rightarrow> ty"   
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   186
where
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   187
  "lookup [] X        = TVar X"
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   188
| "lookup ((Y,T)#\<theta>) X = (if X=Y then T else lookup \<theta> X)"
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   189
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   190
lemma lookup_eqvt[eqvt]:
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   191
  fixes pi::"tvar prm"
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   192
  shows "pi\<bullet>(lookup \<theta> X) = lookup (pi\<bullet>\<theta>) (pi\<bullet>X)"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   193
  by (induct \<theta>) (auto simp add: eqvts)
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   194
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   195
lemma lookup_fresh:
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   196
  fixes X::"tvar"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   197
  assumes a: "X\<sharp>\<theta>"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   198
  shows "lookup \<theta> X = TVar X"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   199
  using a
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   200
  by (induct \<theta>)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   201
    (auto simp add: fresh_list_cons fresh_prod fresh_atm)
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   202
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   203
overloading 
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   204
  psubst_ty \<equiv> "psubst :: Subst \<Rightarrow> ty \<Rightarrow> ty"
29098
6b23a58cc67c Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents: 28655
diff changeset
   205
begin
6b23a58cc67c Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents: 28655
diff changeset
   206
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   207
nominal_primrec 
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   208
  psubst_ty
29098
6b23a58cc67c Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents: 28655
diff changeset
   209
where
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   210
  "\<theta><TVar X>   = lookup \<theta> X"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 41893
diff changeset
   211
| "\<theta><T\<^sub>1 \<rightarrow> T\<^sub>2> = (\<theta><T\<^sub>1>) \<rightarrow> (\<theta><T\<^sub>2>)"
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   212
by (rule TrueI)+
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   213
29098
6b23a58cc67c Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents: 28655
diff changeset
   214
end
6b23a58cc67c Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents: 28655
diff changeset
   215
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   216
lemma psubst_ty_eqvt[eqvt]:
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   217
  fixes pi::"tvar prm"
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   218
  and   \<theta>::"Subst"
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   219
  and   T::"ty"
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   220
  shows "pi\<bullet>(\<theta><T>) = (pi\<bullet>\<theta>)<(pi\<bullet>T)>"
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   221
by (induct T rule: ty.induct) (simp_all add: eqvts)
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   222
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   223
overloading 
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   224
  psubst_tyS \<equiv> "psubst :: Subst \<Rightarrow> tyS \<Rightarrow> tyS"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   225
begin
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   226
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   227
nominal_primrec 
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   228
  psubst_tyS :: "Subst \<Rightarrow> tyS \<Rightarrow> tyS"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   229
where 
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   230
  "\<theta><(Ty T)> = Ty (\<theta><T>)"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   231
| "X\<sharp>\<theta> \<Longrightarrow> \<theta><(\<forall>[X].S)> = \<forall>[X].(\<theta><S>)"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   232
apply(finite_guess add: psubst_ty_eqvt fs_tvar1 | simp add: abs_fresh)+
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   233
apply(fresh_guess add: psubst_ty_eqvt fs_tvar1)+
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   234
done
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   235
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   236
end
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   237
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   238
overloading 
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   239
  psubst_Ctxt \<equiv> "psubst :: Subst \<Rightarrow> Ctxt \<Rightarrow> Ctxt"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   240
begin
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   241
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   242
fun
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   243
  psubst_Ctxt :: "Subst \<Rightarrow> Ctxt \<Rightarrow> Ctxt"
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   244
where
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   245
  "psubst_Ctxt \<theta> [] = []"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   246
| "psubst_Ctxt \<theta> ((x,S)#\<Gamma>) = (x,\<theta><S>)#(psubst_Ctxt \<theta> \<Gamma>)"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   247
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   248
end
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   249
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   250
lemma fresh_lookup:
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   251
  fixes X::"tvar"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   252
  and   \<theta>::"Subst"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   253
  and   Y::"tvar"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   254
  assumes asms: "X\<sharp>Y" "X\<sharp>\<theta>"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   255
  shows "X\<sharp>(lookup \<theta> Y)"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   256
  using asms
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   257
  by (induct \<theta>)
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   258
     (auto simp add: fresh_list_cons fresh_prod fresh_atm)
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   259
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   260
lemma fresh_psubst_ty:
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   261
  fixes X::"tvar"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   262
  and   \<theta>::"Subst"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   263
  and   T::"ty"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   264
  assumes asms: "X\<sharp>\<theta>" "X\<sharp>T"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   265
  shows "X\<sharp>\<theta><T>"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   266
  using asms
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   267
  by (nominal_induct T rule: ty.strong_induct)
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   268
     (auto simp add: fresh_list_append fresh_list_cons fresh_prod fresh_lookup)
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   269
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   270
lemma fresh_psubst_tyS:
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   271
  fixes X::"tvar"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   272
  and   \<theta>::"Subst"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   273
  and   S::"tyS"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   274
  assumes asms: "X\<sharp>\<theta>" "X\<sharp>S"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   275
  shows "X\<sharp>\<theta><S>"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   276
  using asms
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   277
  by (nominal_induct S avoiding: \<theta>  X rule: tyS.strong_induct)
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   278
     (auto simp add: fresh_psubst_ty abs_fresh)
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   279
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   280
lemma fresh_psubst_Ctxt:
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   281
  fixes X::"tvar"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   282
  and   \<theta>::"Subst"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   283
  and   \<Gamma>::"Ctxt"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   284
  assumes asms: "X\<sharp>\<theta>" "X\<sharp>\<Gamma>"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   285
  shows "X\<sharp>\<theta><\<Gamma>>"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   286
using asms
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   287
by (induct \<Gamma>)
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   288
   (auto simp add: fresh_psubst_tyS fresh_list_cons)
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   289
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   290
lemma subst_freshfact2_ty: 
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   291
  fixes   X::"tvar"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   292
  and     Y::"tvar"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   293
  and     T::"ty"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   294
  assumes asms: "X\<sharp>S"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   295
  shows "X\<sharp>T[X::=S]"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   296
  using asms
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   297
by (nominal_induct T rule: ty.strong_induct)
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   298
   (auto simp add: fresh_atm)
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   299
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58219
diff changeset
   300
text \<open>instance of a type scheme\<close>
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   301
inductive
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80149
diff changeset
   302
  inst :: "ty \<Rightarrow> tyS \<Rightarrow> bool"(\<open>_ \<prec> _\<close> [50,51] 50)  
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   303
where
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   304
  I_Ty[intro]:  "T \<prec> (Ty T)" 
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   305
| I_All[intro]: "\<lbrakk>X\<sharp>T'; T \<prec> S\<rbrakk> \<Longrightarrow> T[X::=T'] \<prec> \<forall>[X].S"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   306
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   307
equivariance inst[tvar] 
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   308
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   309
nominal_inductive inst
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   310
  by (simp_all add: abs_fresh subst_freshfact2_ty)
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   311
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   312
lemma subst_forget_ty:
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   313
  fixes T::"ty"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   314
  and   X::"tvar"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   315
  assumes a: "X\<sharp>T"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   316
  shows "T[X::=S] = T"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   317
  using a
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   318
  by  (nominal_induct T rule: ty.strong_induct)
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   319
      (auto simp add: fresh_atm)
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   320
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   321
lemma psubst_ty_lemma:
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   322
  fixes \<theta>::"Subst"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   323
  and   X::"tvar"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   324
  and   T'::"ty"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   325
  and   T::"ty"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   326
  assumes a: "X\<sharp>\<theta>" 
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   327
  shows "\<theta><T[X::=T']> = (\<theta><T>)[X::=\<theta><T'>]"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   328
using a
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   329
proof (nominal_induct T avoiding: \<theta> X T' rule: ty.strong_induct)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   330
  case (TVar tvar)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   331
  then show ?case
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   332
    by (simp add: fresh_atm(1) fresh_lookup lookup_fresh subst_forget_ty)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   333
qed auto
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   334
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   335
lemma general_preserved:
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   336
  fixes \<theta>::"Subst"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   337
  assumes a: "T \<prec> S"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   338
  shows "\<theta><T> \<prec> \<theta><S>"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   339
  using a
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   340
proof (nominal_induct T S avoiding: \<theta> rule: inst.strong_induct)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   341
  case (I_Ty T)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   342
  then show ?case
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   343
    by (simp add: inst.I_Ty)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   344
next
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   345
  case (I_All X T' T S)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   346
  then show ?case
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   347
    by (simp add: fresh_psubst_ty inst.I_All psubst_ty_lemma)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   348
qed
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   349
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   350
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 58219
diff changeset
   351
text\<open>typing judgements\<close>
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   352
inductive
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80149
diff changeset
   353
  typing :: "Ctxt \<Rightarrow> trm \<Rightarrow> ty \<Rightarrow> bool" (\<open> _ \<turnstile> _ : _ \<close> [60,60,60] 60) 
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   354
where
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   355
  T_VAR[intro]: "\<lbrakk>valid \<Gamma>; (x,S)\<in>set \<Gamma>; T \<prec> S\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> Var x : T"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 41893
diff changeset
   356
| T_APP[intro]: "\<lbrakk>\<Gamma> \<turnstile> t\<^sub>1 : T\<^sub>1\<rightarrow>T\<^sub>2; \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>1\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> App t\<^sub>1 t\<^sub>2 : T\<^sub>2" 
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 41893
diff changeset
   357
| T_LAM[intro]: "\<lbrakk>x\<sharp>\<Gamma>;((x,Ty T\<^sub>1)#\<Gamma>) \<turnstile> t : T\<^sub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [x].t : T\<^sub>1\<rightarrow>T\<^sub>2"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 41893
diff changeset
   358
| T_LET[intro]: "\<lbrakk>x\<sharp>\<Gamma>; \<Gamma> \<turnstile> t\<^sub>1 : T\<^sub>1; ((x,close \<Gamma> T\<^sub>1)#\<Gamma>) \<turnstile> t\<^sub>2 : T\<^sub>2; set (ftv T\<^sub>1 - ftv \<Gamma>) \<sharp>* T\<^sub>2\<rbrakk> 
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 41893
diff changeset
   359
                 \<Longrightarrow> \<Gamma> \<turnstile> Let x be t\<^sub>1 in t\<^sub>2 : T\<^sub>2"
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   360
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   361
equivariance typing[tvar]
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   362
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   363
lemma fresh_tvar_trm: 
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   364
  fixes X::"tvar"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   365
  and   t::"trm"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   366
  shows "X\<sharp>t"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   367
by (nominal_induct t rule: trm.strong_induct) 
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   368
   (simp_all add: fresh_atm abs_fresh)
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   369
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   370
lemma ftv_ty: 
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   371
  fixes T::"ty"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   372
  shows "supp T = set (ftv T)"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   373
by (nominal_induct T rule: ty.strong_induct) 
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   374
   (simp_all add: ty.supp supp_atm)
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   375
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   376
lemma ftv_tyS: 
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   377
  fixes S::"tyS"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   378
  shows "supp S = set (ftv S)"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   379
by (nominal_induct S rule: tyS.strong_induct) 
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   380
   (auto simp add: tyS.supp abs_supp ftv_ty)
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   381
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   382
lemma ftv_Ctxt: 
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   383
  fixes \<Gamma>::"Ctxt"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   384
  shows "supp \<Gamma> = set (ftv \<Gamma>)"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   385
proof (induct \<Gamma>)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   386
  case Nil
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   387
  then show ?case
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   388
    by (simp add: supp_list_nil)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   389
next
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   390
  case (Cons c \<Gamma>)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   391
  show ?case
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   392
  proof (cases c)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   393
    case (Pair a b)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   394
    with Cons show ?thesis
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   395
      by (simp add: ftv_tyS supp_atm(3) supp_list_cons supp_prod)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   396
  qed
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   397
qed
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   398
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   399
lemma ftv_tvars:
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   400
  fixes Tvs::"tvar list"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   401
  shows "supp Tvs = set Tvs"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   402
  by (induct Tvs) (simp_all add: supp_list_nil supp_list_cons supp_atm)
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   403
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   404
lemma difference_supp: 
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   405
  fixes xs ys::"tvar list"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   406
  shows "((supp (xs - ys))::tvar set) = supp xs - supp ys"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   407
  by (induct xs) (auto simp add: supp_list_nil supp_list_cons ftv_tvars)
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   408
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   409
lemma set_supp_eq: 
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   410
  fixes xs::"tvar list"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   411
  shows "set xs = supp xs"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   412
  by (induct xs) (simp_all add: supp_list_nil supp_list_cons supp_atm)
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   413
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   414
nominal_inductive2 typing
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 41893
diff changeset
   415
  avoids T_LET: "set (ftv T\<^sub>1 - ftv \<Gamma>)"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   416
     apply (simp_all add: fresh_star_def fresh_def ftv_Ctxt fresh_tvar_trm)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   417
  by (meson fresh_def fresh_tvar_trm)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   418
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   419
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   420
lemma perm_fresh_fresh_aux:
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   421
  "\<forall>(x,y)\<in>set (pi::tvar prm). x \<sharp> z \<and> y \<sharp> z \<Longrightarrow> pi \<bullet> (z::'a::pt_tvar) = z"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   422
proof (induct pi rule: rev_induct)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   423
  case Nil
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   424
  then show ?case
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   425
    by simp
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   426
next
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   427
  case (snoc x xs)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   428
  then show ?case
80149
40a3fc07a587 More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 80140
diff changeset
   429
    unfolding split_paired_all pt_tvar2
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   430
    using perm_fresh_fresh(1) by fastforce
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   431
qed
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   432
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   433
lemma freshs_mem:
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   434
  fixes S::"tvar set"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   435
  assumes "x \<in> S"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   436
    and     "S \<sharp>* z"
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   437
  shows  "x \<sharp> z"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   438
  using assms by (simp add: fresh_star_def)
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   439
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   440
lemma fresh_gen_set:
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   441
  fixes X::"tvar"
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   442
  and   Xs::"tvar list"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   443
  assumes "X\<in>set Xs" 
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   444
  shows "X\<sharp>gen T Xs"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   445
  using assms by (induct Xs) (auto simp: abs_fresh)
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   446
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   447
lemma close_fresh:
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   448
  fixes \<Gamma>::"Ctxt"
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   449
  shows "\<forall>(X::tvar)\<in>set ((ftv T) - (ftv \<Gamma>)). X\<sharp>(close \<Gamma> T)"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   450
  by (meson fresh_gen_set)
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   451
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   452
lemma gen_supp: 
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   453
  shows "(supp (gen T Xs)::tvar set) = supp T - supp Xs"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   454
by (induct Xs) 
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   455
   (auto simp add: supp_list_nil supp_list_cons tyS.supp abs_supp supp_atm)
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   456
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   457
lemma minus_Int_eq: 
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   458
  shows "T - (T - U) = T \<inter> U"
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   459
  by blast
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   460
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   461
lemma close_supp: 
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   462
  shows "supp (close \<Gamma> T) = set (ftv T) \<inter> set (ftv \<Gamma>)"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   463
  using difference_supp ftv_ty gen_supp set_supp_eq by auto
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   464
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   465
lemma better_T_LET:
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   466
  assumes x: "x\<sharp>\<Gamma>"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 41893
diff changeset
   467
  and t1: "\<Gamma> \<turnstile> t\<^sub>1 : T\<^sub>1"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 41893
diff changeset
   468
  and t2: "((x,close \<Gamma> T\<^sub>1)#\<Gamma>) \<turnstile> t\<^sub>2 : T\<^sub>2"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 41893
diff changeset
   469
  shows "\<Gamma> \<turnstile> Let x be t\<^sub>1 in t\<^sub>2 : T\<^sub>2"
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   470
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 41893
diff changeset
   471
  have fin: "finite (set (ftv T\<^sub>1 - ftv \<Gamma>))" by simp
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 41893
diff changeset
   472
  obtain pi where pi1: "(pi \<bullet> set (ftv T\<^sub>1 - ftv \<Gamma>)) \<sharp>* (T\<^sub>2, \<Gamma>)"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 41893
diff changeset
   473
    and pi2: "set pi \<subseteq> set (ftv T\<^sub>1 - ftv \<Gamma>) \<times> (pi \<bullet> set (ftv T\<^sub>1 - ftv \<Gamma>))"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 41893
diff changeset
   474
    by (rule at_set_avoiding [OF at_tvar_inst fin fs_tvar1, of "(T\<^sub>2, \<Gamma>)"])
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 41893
diff changeset
   475
  from pi1 have pi1': "(pi \<bullet> set (ftv T\<^sub>1 - ftv \<Gamma>)) \<sharp>* \<Gamma>"
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   476
    by (simp add: fresh_star_prod)
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   477
  have Gamma_fresh: "\<forall>(x,y)\<in>set pi. x \<sharp> \<Gamma> \<and> y \<sharp> \<Gamma>"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   478
    using freshs_mem [OF _ pi1'] subsetD [OF pi2] ftv_Ctxt fresh_def by fastforce
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   479
  have "\<And>x y. \<lbrakk>x \<in> set (ftv T\<^sub>1 - ftv \<Gamma>); y \<in> pi \<bullet> set (ftv T\<^sub>1 - ftv \<Gamma>)\<rbrakk>
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   480
              \<Longrightarrow> x \<sharp> close \<Gamma> T\<^sub>1 \<and> y \<sharp> close \<Gamma> T\<^sub>1"
82664
e9f3b94eb6a0 more modern qualification of auxiliary operations
haftmann
parents: 80914
diff changeset
   481
    using pi1' fresh_def fresh_gen_set freshs_mem ftv_Ctxt ftv_ty gen_supp
e9f3b94eb6a0 more modern qualification of auxiliary operations
haftmann
parents: 80914
diff changeset
   482
    by (metis (lifting) DiffE mem_Collect_eq set_filter)
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   483
  then have close_fresh': "\<forall>(x, y)\<in>set pi. x \<sharp> close \<Gamma> T\<^sub>1 \<and> y \<sharp> close \<Gamma> T\<^sub>1"
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   484
    using pi2 by auto
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   485
  note x
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   486
  moreover from Gamma_fresh perm_boolI [OF t1, of pi]
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 41893
diff changeset
   487
  have "\<Gamma> \<turnstile> t\<^sub>1 : pi \<bullet> T\<^sub>1"
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   488
    by (simp add: perm_fresh_fresh_aux eqvts fresh_tvar_trm)
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   489
  moreover from t2 close_fresh'
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 41893
diff changeset
   490
  have "(x,(pi \<bullet> close \<Gamma> T\<^sub>1))#\<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2"
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   491
    by (simp add: perm_fresh_fresh_aux)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 41893
diff changeset
   492
  with Gamma_fresh have "(x,close \<Gamma> (pi \<bullet> T\<^sub>1))#\<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2"
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   493
    by (simp add: close_eqvt perm_fresh_fresh_aux)
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   494
  moreover from pi1 Gamma_fresh
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 41893
diff changeset
   495
  have "set (ftv (pi \<bullet> T\<^sub>1) - ftv \<Gamma>) \<sharp>* T\<^sub>2"
28655
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   496
    by (simp only: eqvts fresh_star_prod perm_fresh_fresh_aux)
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   497
  ultimately show ?thesis by (rule T_LET)
2822c56dd1cf Example for using the generalized version of nominal_inductive.
berghofe
parents: 26195
diff changeset
   498
qed
26195
8292f8723e99 added new example
urbanc
parents:
diff changeset
   499
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   500
lemma ftv_ty_subst:
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   501
  fixes T::"ty"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   502
  and   \<theta>::"Subst"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   503
  and   X Y ::"tvar"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   504
  assumes "X \<in> set (ftv T)"
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   505
  and     "Y \<in> set (ftv (lookup \<theta> X))"
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   506
  shows "Y \<in> set (ftv (\<theta><T>))"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   507
  using assms
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   508
  by (nominal_induct T rule: ty.strong_induct) (auto)
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   509
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   510
lemma ftv_tyS_subst:
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   511
  fixes S::"tyS"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   512
  and   \<theta>::"Subst"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   513
  and   X Y::"tvar"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   514
  assumes "X \<in> set (ftv S)"
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   515
  and     "Y \<in> set (ftv (lookup \<theta> X))"
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   516
  shows "Y \<in> set (ftv (\<theta><S>))"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   517
  using assms
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   518
by (nominal_induct S avoiding: \<theta> Y rule: tyS.strong_induct) 
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   519
   (auto simp add: ftv_ty_subst fresh_atm)
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   520
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   521
lemma ftv_Ctxt_subst:
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   522
  fixes \<Gamma>::"Ctxt"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   523
  and   \<theta>::"Subst"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   524
assumes "X \<in> set (ftv \<Gamma>)"
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   525
  and   "Y \<in> set (ftv (lookup \<theta> X))"
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   526
  shows "Y \<in> set (ftv (\<theta><\<Gamma>>))"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   527
  using assms by (induct \<Gamma>) (auto simp add: ftv_tyS_subst)
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   528
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   529
lemma gen_preserved1:
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   530
  assumes "Xs \<sharp>* \<theta>"
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   531
  shows "\<theta><gen T Xs> = gen (\<theta><T>) Xs"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   532
  using assms by (induct Xs) (auto simp add: fresh_star_def)
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   533
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   534
lemma gen_preserved2:
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   535
  fixes T::"ty"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   536
  and   \<Gamma>::"Ctxt"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   537
  assumes "((ftv T) - (ftv \<Gamma>)) \<sharp>* \<theta>"
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   538
  shows "((ftv (\<theta><T>)) - (ftv (\<theta><\<Gamma>>))) = ((ftv T) - (ftv \<Gamma>))"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   539
  using assms
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   540
proof(nominal_induct T rule: ty.strong_induct)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   541
  case (TVar tvar)
80149
40a3fc07a587 More tidying of proofs
paulson <lp15@cam.ac.uk>
parents: 80140
diff changeset
   542
  then show ?case 
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   543
    apply(auto simp add: fresh_star_def ftv_Ctxt_subst)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   544
    by (metis filter.simps fresh_def fresh_psubst_Ctxt ftv_Ctxt ftv_ty.simps(1) lookup_fresh)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   545
next
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   546
  case (Fun ty1 ty2)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   547
  then show ?case
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   548
    by (simp add: fresh_star_list) 
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   549
qed
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   550
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   551
lemma close_preserved:
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   552
  fixes \<Gamma>::"Ctxt"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   553
  assumes "((ftv T) - (ftv \<Gamma>)) \<sharp>* \<theta>"
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   554
  shows "\<theta><close \<Gamma> T> = close (\<theta><\<Gamma>>) (\<theta><T>)"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   555
  using assms by (simp add: gen_preserved1 gen_preserved2)
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   556
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   557
lemma var_fresh_for_ty:
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   558
  fixes x::"var"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   559
    and   T::"ty"
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   560
  shows "x\<sharp>T"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   561
  by (nominal_induct T rule: ty.strong_induct) (simp_all add: fresh_atm)
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   562
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   563
lemma var_fresh_for_tyS:
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   564
  fixes x::"var" and S::"tyS"
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   565
  shows "x\<sharp>S"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   566
  by (nominal_induct S rule: tyS.strong_induct) (simp_all add: abs_fresh var_fresh_for_ty)
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   567
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   568
lemma psubst_fresh_Ctxt:
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   569
  fixes x::"var" and \<Gamma>::"Ctxt" and \<theta>::"Subst"
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   570
  shows "x\<sharp>\<theta><\<Gamma>> = x\<sharp>\<Gamma>"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   571
  by (induct \<Gamma>) (auto simp add: fresh_list_cons fresh_list_nil fresh_prod var_fresh_for_tyS)
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   572
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   573
lemma psubst_valid:
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   574
  fixes \<theta>::Subst
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   575
    and   \<Gamma>::Ctxt
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   576
  assumes "valid \<Gamma>"
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   577
  shows "valid (\<theta><\<Gamma>>)"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   578
  using assms by (induct) (auto simp add: psubst_fresh_Ctxt)
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   579
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   580
lemma psubst_in:
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   581
  fixes \<Gamma>::"Ctxt"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   582
  and   \<theta>::"Subst"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   583
  and   pi::"tvar prm"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   584
  and   S::"tyS"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   585
  assumes a: "(x,S)\<in>set \<Gamma>"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   586
  shows "(x,\<theta><S>)\<in>set (\<theta><\<Gamma>>)"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   587
  using a by (induct \<Gamma>) (auto simp add: calc_atm)
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   588
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   589
lemma typing_preserved:
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   590
  fixes \<theta>::"Subst"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   591
    and   pi::"tvar prm"
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   592
  assumes "\<Gamma> \<turnstile> t : T"
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   593
  shows "(\<theta><\<Gamma>>) \<turnstile> t : (\<theta><T>)"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   594
  using assms
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   595
proof (nominal_induct \<Gamma> t T avoiding: \<theta> rule: typing.strong_induct)
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   596
  case (T_VAR \<Gamma> x S T)
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   597
  have a1: "valid \<Gamma>" by fact
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   598
  have a2: "(x, S) \<in> set \<Gamma>" by fact
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   599
  have a3: "T \<prec> S" by fact
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   600
  have  "valid (\<theta><\<Gamma>>)" using a1 by (simp add: psubst_valid)
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   601
  moreover
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   602
  have  "(x,\<theta><S>)\<in>set (\<theta><\<Gamma>>)" using a2 by (simp add: psubst_in)
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   603
  moreover
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   604
  have "\<theta><T> \<prec> \<theta><S>" using a3 by (simp add: general_preserved)
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   605
  ultimately show "(\<theta><\<Gamma>>) \<turnstile> Var x : (\<theta><T>)" by (simp add: typing.T_VAR)
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   606
next
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   607
  case (T_APP \<Gamma> t1 T1 T2 t2)
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   608
  have "\<theta><\<Gamma>> \<turnstile> t1 : \<theta><T1 \<rightarrow> T2>" by fact
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   609
  then have "\<theta><\<Gamma>> \<turnstile> t1 : (\<theta><T1>) \<rightarrow> (\<theta><T2>)" by simp
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   610
  moreover
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   611
  have "\<theta><\<Gamma>> \<turnstile> t2 : \<theta><T1>" by fact
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   612
  ultimately show "\<theta><\<Gamma>> \<turnstile> App t1 t2 : \<theta><T2>" by (simp add: typing.T_APP)
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   613
next
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   614
  case (T_LAM x \<Gamma> T1 t T2)
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   615
  fix pi::"tvar prm" and \<theta>::"Subst"
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   616
  have "x\<sharp>\<Gamma>" by fact
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   617
  then have "x\<sharp>\<theta><\<Gamma>>" by (simp add: psubst_fresh_Ctxt)
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   618
  moreover
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   619
  have "\<theta><((x, Ty T1)#\<Gamma>)> \<turnstile> t : \<theta><T2>" by fact
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   620
  then have "((x, Ty (\<theta><T1>))#(\<theta><\<Gamma>>)) \<turnstile> t : \<theta><T2>" by (simp add: calc_atm)
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   621
  ultimately show "\<theta><\<Gamma>> \<turnstile> Lam [x].t : \<theta><T1 \<rightarrow> T2>" by (simp add: typing.T_LAM)
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   622
next
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   623
  case (T_LET x \<Gamma> t1 T1 t2 T2)
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   624
  have vc: "((ftv T1) - (ftv \<Gamma>)) \<sharp>* \<theta>" by fact
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   625
  have "x\<sharp>\<Gamma>" by fact
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   626
   then have a1: "x\<sharp>\<theta><\<Gamma>>" by (simp add: calc_atm psubst_fresh_Ctxt)
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   627
  have a2: "\<theta><\<Gamma>> \<turnstile> t1 : \<theta><T1>" by fact 
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   628
  have a3: "\<theta><((x, close \<Gamma> T1)#\<Gamma>)> \<turnstile> t2 : \<theta><T2>" by fact
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   629
  from a2 a3 show "\<theta><\<Gamma>> \<turnstile> Let x be t1 in t2 : \<theta><T2>"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 68386
diff changeset
   630
    by (simp add: a1 better_T_LET close_preserved vc)
32137
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   631
qed
3b260527fc11 tuned proofs and added some lemmas
Christian Urban <urbanc@in.tum.de>
parents: 29608
diff changeset
   632
26195
8292f8723e99 added new example
urbanc
parents:
diff changeset
   633
end