src/HOL/Nominal/Examples/Pattern.thy
author wenzelm
Mon, 23 Sep 2024 22:33:37 +0200
changeset 80935 b5bdcdbf5ec1
parent 80914 d97fdabd9e2b
child 81125 ec121999a9cb
permissions -rw-r--r--
proper 'no_syntax' (amending 8e72f55295fd);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
     1
section \<open>Simply-typed lambda-calculus with let and tuple patterns\<close>
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
     2
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
     3
theory Pattern
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 63167
diff changeset
     4
imports "HOL-Nominal.Nominal"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
     5
begin
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
     6
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
     7
no_syntax
80935
b5bdcdbf5ec1 proper 'no_syntax' (amending 8e72f55295fd);
wenzelm
parents: 80914
diff changeset
     8
  "_Map" :: "maplets => 'a \<rightharpoonup> 'b"  (\<open>(\<open>indent=1 notation=\<open>mixfix maplets\<close>\<close>[_])\<close>)
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
     9
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    10
atom_decl name
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    11
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    12
nominal_datatype ty =
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    13
    Atom nat
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80142
diff changeset
    14
  | Arrow ty ty  (infixr \<open>\<rightarrow>\<close> 200)
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80142
diff changeset
    15
  | TupleT ty ty  (infixr \<open>\<otimes>\<close> 210)
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    16
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    17
lemma fresh_type [simp]: "(a::name) \<sharp> (T::ty)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    18
  by (induct T rule: ty.induct) (simp_all add: fresh_nat)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    19
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    20
lemma supp_type [simp]: "supp (T::ty) = ({} :: name set)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    21
  by (induct T rule: ty.induct) (simp_all add: ty.supp supp_nat)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    22
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    23
lemma perm_type: "(pi::name prm) \<bullet> (T::ty) = T"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    24
  by (induct T rule: ty.induct) (simp_all add: perm_nat_def)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    25
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    26
nominal_datatype trm =
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    27
    Var name
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80142
diff changeset
    28
  | Tuple trm trm  (\<open>(1'\<langle>_,/ _'\<rangle>)\<close>)
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    29
  | Abs ty "\<guillemotleft>name\<guillemotright>trm"
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80142
diff changeset
    30
  | App trm trm  (infixl \<open>\<cdot>\<close> 200)
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    31
  | Let ty trm btrm
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    32
and btrm =
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    33
    Base trm
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    34
  | Bind ty "\<guillemotleft>name\<guillemotright>btrm"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    35
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    36
abbreviation
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80142
diff changeset
    37
  Abs_syn :: "name \<Rightarrow> ty \<Rightarrow> trm \<Rightarrow> trm"  (\<open>(3\<lambda>_:_./ _)\<close> [0, 0, 10] 10) 
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    38
where
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    39
  "\<lambda>x:T. t \<equiv> Abs T x t"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    40
58310
91ea607a34d8 updated news
blanchet
parents: 58249
diff changeset
    41
datatype pat =
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    42
    PVar name ty
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80142
diff changeset
    43
  | PTuple pat pat  (\<open>(1'\<langle>\<langle>_,/ _'\<rangle>\<rangle>)\<close>)
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    44
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    45
(* FIXME: The following should be done automatically by the nominal package *)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    46
overloading pat_perm \<equiv> "perm :: name prm \<Rightarrow> pat \<Rightarrow> pat" (unchecked)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    47
begin
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    48
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    49
primrec pat_perm
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    50
where
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    51
  "pat_perm pi (PVar x ty) = PVar (pi \<bullet> x) (pi \<bullet> ty)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    52
| "pat_perm pi \<langle>\<langle>p, q\<rangle>\<rangle> = \<langle>\<langle>pat_perm pi p, pat_perm pi q\<rangle>\<rangle>"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    53
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    54
end
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    55
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    56
declare pat_perm.simps [eqvt]
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    57
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    58
lemma supp_PVar [simp]: "((supp (PVar x T))::name set) = supp x"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    59
  by (simp add: supp_def perm_fresh_fresh)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    60
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    61
lemma supp_PTuple [simp]: "((supp \<langle>\<langle>p, q\<rangle>\<rangle>)::name set) = supp p \<union> supp q"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    62
  by (simp add: supp_def Collect_disj_eq del: disj_not1)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    63
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    64
instance pat :: pt_name
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
    65
proof
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
    66
  fix x :: pat
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
    67
  show "([]::(name \<times> _) list) \<bullet> x = x"
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
    68
    by (induct x) simp_all
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
    69
  fix pi1 pi2 :: "(name \<times> name) list"
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
    70
  show "(pi1 @ pi2) \<bullet> x = pi1 \<bullet> pi2 \<bullet> x"
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
    71
    by (induct x) (simp_all add: pt_name2)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
    72
  assume "pi1 \<triangleq> pi2"
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
    73
  then show "pi1 \<bullet> x = pi2 \<bullet> x"
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
    74
    by (induct x) (simp_all add: pt_name3)
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    75
qed
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    76
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    77
instance pat :: fs_name
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
    78
proof
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
    79
  fix x :: pat
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
    80
  show "finite (supp x::name set)"
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
    81
    by (induct x) (simp_all add: fin_supp)
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    82
qed
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    83
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    84
(* the following function cannot be defined using nominal_primrec, *)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    85
(* since variable parameters are currently not allowed.            *)
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80142
diff changeset
    86
primrec abs_pat :: "pat \<Rightarrow> btrm \<Rightarrow> btrm" (\<open>(3\<lambda>[_]./ _)\<close> [0, 10] 10)
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    87
where
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    88
  "(\<lambda>[PVar x T]. t) = Bind T x t"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    89
| "(\<lambda>[\<langle>\<langle>p, q\<rangle>\<rangle>]. t) = (\<lambda>[p]. \<lambda>[q]. t)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    90
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    91
lemma abs_pat_eqvt [eqvt]:
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    92
  "(pi :: name prm) \<bullet> (\<lambda>[p]. t) = (\<lambda>[pi \<bullet> p]. (pi \<bullet> t))"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    93
  by (induct p arbitrary: t) simp_all
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    94
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    95
lemma abs_pat_fresh [simp]:
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    96
  "(x::name) \<sharp> (\<lambda>[p]. t) = (x \<in> supp p \<or> x \<sharp> t)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    97
  by (induct p arbitrary: t) (simp_all add: abs_fresh supp_atm)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    98
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
    99
lemma abs_pat_alpha:
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   100
  assumes fresh: "((pi::name prm) \<bullet> supp p::name set) \<sharp>* t"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   101
  and pi: "set pi \<subseteq> supp p \<times> pi \<bullet> supp p"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   102
  shows "(\<lambda>[p]. t) = (\<lambda>[pi \<bullet> p]. pi \<bullet> t)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   103
proof -
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   104
  note pt_name_inst at_name_inst pi
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   105
  moreover have "(supp p::name set) \<sharp>* (\<lambda>[p]. t)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   106
    by (simp add: fresh_star_def)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   107
  moreover from fresh
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   108
  have "(pi \<bullet> supp p::name set) \<sharp>* (\<lambda>[p]. t)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   109
    by (simp add: fresh_star_def)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   110
  ultimately have "pi \<bullet> (\<lambda>[p]. t) = (\<lambda>[p]. t)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   111
    by (rule pt_freshs_freshs)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   112
  then show ?thesis by (simp add: eqvts)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   113
qed
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   114
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   115
primrec pat_vars :: "pat \<Rightarrow> name list"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   116
where
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   117
  "pat_vars (PVar x T) = [x]"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   118
| "pat_vars \<langle>\<langle>p, q\<rangle>\<rangle> = pat_vars q @ pat_vars p"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   119
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   120
lemma pat_vars_eqvt [eqvt]:
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   121
  "(pi :: name prm) \<bullet> (pat_vars p) = pat_vars (pi \<bullet> p)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   122
  by (induct p rule: pat.induct) (simp_all add: eqvts)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   123
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   124
lemma set_pat_vars_supp: "set (pat_vars p) = supp p"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   125
  by (induct p) (auto simp add: supp_atm)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   126
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   127
lemma distinct_eqvt [eqvt]:
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   128
  "(pi :: name prm) \<bullet> (distinct (xs::name list)) = distinct (pi \<bullet> xs)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   129
  by (induct xs) (simp_all add: eqvts)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   130
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   131
primrec pat_type :: "pat \<Rightarrow> ty"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   132
where
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   133
  "pat_type (PVar x T) = T"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   134
| "pat_type \<langle>\<langle>p, q\<rangle>\<rangle> = pat_type p \<otimes> pat_type q"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   135
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   136
lemma pat_type_eqvt [eqvt]:
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   137
  "(pi :: name prm) \<bullet> (pat_type p) = pat_type (pi \<bullet> p)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   138
  by (induct p) simp_all
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   139
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   140
lemma pat_type_perm_eq: "pat_type ((pi :: name prm) \<bullet> p) = pat_type p"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   141
  by (induct p) (simp_all add: perm_type)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   142
41798
c3aa3c87ef21 modernized specifications;
wenzelm
parents: 34990
diff changeset
   143
type_synonym ctx = "(name \<times> ty) list"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   144
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   145
inductive
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80142
diff changeset
   146
  ptyping :: "pat \<Rightarrow> ty \<Rightarrow> ctx \<Rightarrow> bool"  (\<open>\<turnstile> _ : _ \<Rightarrow> _\<close> [60, 60, 60] 60)
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   147
where
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   148
  PVar: "\<turnstile> PVar x T : T \<Rightarrow> [(x, T)]"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   149
| PTuple: "\<turnstile> p : T \<Rightarrow> \<Delta>\<^sub>1 \<Longrightarrow> \<turnstile> q : U \<Rightarrow> \<Delta>\<^sub>2 \<Longrightarrow> \<turnstile> \<langle>\<langle>p, q\<rangle>\<rangle> : T \<otimes> U \<Rightarrow> \<Delta>\<^sub>2 @ \<Delta>\<^sub>1"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   150
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   151
lemma pat_vars_ptyping:
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   152
  assumes "\<turnstile> p : T \<Rightarrow> \<Delta>"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   153
  shows "pat_vars p = map fst \<Delta>" using assms
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   154
  by induct simp_all
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   155
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   156
inductive
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   157
  valid :: "ctx \<Rightarrow> bool"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   158
where
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   159
  Nil [intro!]: "valid []"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   160
| Cons [intro!]: "valid \<Gamma> \<Longrightarrow> x \<sharp> \<Gamma> \<Longrightarrow> valid ((x, T) # \<Gamma>)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   161
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   162
inductive_cases validE[elim!]: "valid ((x, T) # \<Gamma>)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   163
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   164
lemma fresh_ctxt_set_eq: "((x::name) \<sharp> (\<Gamma>::ctx)) = (x \<notin> fst ` set \<Gamma>)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   165
  by (induct \<Gamma>) (auto simp add: fresh_list_nil fresh_list_cons fresh_prod fresh_atm)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   166
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   167
lemma valid_distinct: "valid \<Gamma> = distinct (map fst \<Gamma>)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   168
  by (induct \<Gamma>) (auto simp add: fresh_ctxt_set_eq [symmetric])
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   169
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   170
abbreviation
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80142
diff changeset
   171
  "sub_ctx" :: "ctx \<Rightarrow> ctx \<Rightarrow> bool" (\<open>_ \<sqsubseteq> _\<close>) 
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   172
where
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   173
  "\<Gamma>\<^sub>1 \<sqsubseteq> \<Gamma>\<^sub>2 \<equiv> \<forall>x. x \<in> set \<Gamma>\<^sub>1 \<longrightarrow> x \<in> set \<Gamma>\<^sub>2"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   174
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   175
abbreviation
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80142
diff changeset
   176
  Let_syn :: "pat \<Rightarrow> trm \<Rightarrow> trm \<Rightarrow> trm"  (\<open>(LET (_ =/ _)/ IN (_))\<close> 10)
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   177
where
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   178
  "LET p = t IN u \<equiv> Let (pat_type p) t (\<lambda>[p]. Base u)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   179
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80142
diff changeset
   180
inductive typing :: "ctx \<Rightarrow> trm \<Rightarrow> ty \<Rightarrow> bool" (\<open>_ \<turnstile> _ : _\<close> [60, 60, 60] 60)
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   181
where
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   182
  Var [intro]: "valid \<Gamma> \<Longrightarrow> (x, T) \<in> set \<Gamma> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   183
| Tuple [intro]: "\<Gamma> \<turnstile> t : T \<Longrightarrow> \<Gamma> \<turnstile> u : U \<Longrightarrow> \<Gamma> \<turnstile> \<langle>t, u\<rangle> : T \<otimes> U"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   184
| Abs [intro]: "(x, T) # \<Gamma> \<turnstile> t : U \<Longrightarrow> \<Gamma> \<turnstile> (\<lambda>x:T. t) : T \<rightarrow> U"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   185
| App [intro]: "\<Gamma> \<turnstile> t : T \<rightarrow> U \<Longrightarrow> \<Gamma> \<turnstile> u : T \<Longrightarrow> \<Gamma> \<turnstile> t \<cdot> u : U"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   186
| Let: "((supp p)::name set) \<sharp>* t \<Longrightarrow>
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   187
    \<Gamma> \<turnstile> t : T \<Longrightarrow> \<turnstile> p : T \<Rightarrow> \<Delta> \<Longrightarrow> \<Delta> @ \<Gamma> \<turnstile> u : U \<Longrightarrow>
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   188
    \<Gamma> \<turnstile> (LET p = t IN u) : U"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   190
equivariance ptyping
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   191
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   192
equivariance valid
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   193
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   194
equivariance typing
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   195
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   196
lemma valid_typing:
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   197
  assumes "\<Gamma> \<turnstile> t : T"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   198
  shows "valid \<Gamma>" using assms
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   199
  by induct auto
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   200
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   201
lemma pat_var:
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   202
  assumes "\<turnstile> p : T \<Rightarrow> \<Delta>"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   203
  shows "(supp p::name set) = supp \<Delta>" using assms
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   204
  by induct (auto simp add: supp_list_nil supp_list_cons supp_prod supp_list_append)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   205
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   206
lemma valid_app_fresh:
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   207
  assumes "valid (\<Delta> @ \<Gamma>)" and "(x::name) \<in> supp \<Delta>"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   208
  shows "x \<sharp> \<Gamma>" using assms
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   209
  by (induct \<Delta>)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   210
    (auto simp add: supp_list_nil supp_list_cons supp_prod supp_atm fresh_list_append)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   211
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   212
lemma pat_freshs:
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   213
  assumes "\<turnstile> p : T \<Rightarrow> \<Delta>"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   214
  shows "(supp p::name set) \<sharp>* c = (supp \<Delta>::name set) \<sharp>* c" using assms
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   215
  by (auto simp add: fresh_star_def pat_var)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   216
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   217
lemma valid_app_mono:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   218
  assumes "valid (\<Delta> @ \<Gamma>\<^sub>1)" and "(supp \<Delta>::name set) \<sharp>* \<Gamma>\<^sub>2" and "valid \<Gamma>\<^sub>2" and "\<Gamma>\<^sub>1 \<sqsubseteq> \<Gamma>\<^sub>2"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   219
  shows "valid (\<Delta> @ \<Gamma>\<^sub>2)" using assms
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   220
  by (induct \<Delta>)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   221
    (auto simp add: supp_list_cons fresh_star_Un_elim supp_prod
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   222
       fresh_list_append supp_atm fresh_star_insert_elim fresh_star_empty_elim)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   223
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   224
nominal_inductive2 typing
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   225
avoids
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   226
  Abs: "{x}"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   227
| Let: "(supp p)::name set"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   228
  by (auto simp add: fresh_star_def abs_fresh fin_supp pat_var
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   229
    dest!: valid_typing valid_app_fresh)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   230
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   231
lemma better_T_Let [intro]:
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   232
  assumes t: "\<Gamma> \<turnstile> t : T" and p: "\<turnstile> p : T \<Rightarrow> \<Delta>" and u: "\<Delta> @ \<Gamma> \<turnstile> u : U"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   233
  shows "\<Gamma> \<turnstile> (LET p = t IN u) : U"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   234
proof -
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   235
  obtain pi::"name prm" where pi: "(pi \<bullet> (supp p::name set)) \<sharp>* (t, Base u, \<Gamma>)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   236
    and pi': "set pi \<subseteq> supp p \<times> (pi \<bullet> supp p)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   237
    by (rule at_set_avoiding [OF at_name_inst fin_supp fin_supp])
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   238
  from p u have p_fresh: "(supp p::name set) \<sharp>* \<Gamma>"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   239
    by (auto simp add: fresh_star_def pat_var dest!: valid_typing valid_app_fresh)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   240
  from pi have p_fresh': "(pi \<bullet> (supp p::name set)) \<sharp>* \<Gamma>"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   241
    by (simp add: fresh_star_prod_elim)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   242
  from pi have p_fresh'': "(pi \<bullet> (supp p::name set)) \<sharp>* Base u"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   243
    by (simp add: fresh_star_prod_elim)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   244
  from pi have "(supp (pi \<bullet> p)::name set) \<sharp>* t"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   245
    by (simp add: fresh_star_prod_elim eqvts)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   246
  moreover note t
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   247
  moreover from p have "pi \<bullet> (\<turnstile> p : T \<Rightarrow> \<Delta>)" by (rule perm_boolI)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   248
  then have "\<turnstile> (pi \<bullet> p) : T \<Rightarrow> (pi \<bullet> \<Delta>)" by (simp add: eqvts perm_type)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   249
  moreover from u have "pi \<bullet> (\<Delta> @ \<Gamma> \<turnstile> u : U)" by (rule perm_boolI)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   250
  with pt_freshs_freshs [OF pt_name_inst at_name_inst pi' p_fresh p_fresh']
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   251
  have "(pi \<bullet> \<Delta>) @ \<Gamma> \<turnstile> (pi \<bullet> u) : U" by (simp add: eqvts perm_type)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   252
  ultimately have "\<Gamma> \<turnstile> (LET (pi \<bullet> p) = t IN (pi \<bullet> u)) : U"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   253
    by (rule Let)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   254
  then show ?thesis by (simp add: abs_pat_alpha [OF p_fresh'' pi'] pat_type_perm_eq)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   255
qed
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   256
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   257
lemma weakening: 
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   258
  assumes "\<Gamma>\<^sub>1 \<turnstile> t : T" and "valid \<Gamma>\<^sub>2" and "\<Gamma>\<^sub>1 \<sqsubseteq> \<Gamma>\<^sub>2"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   259
  shows "\<Gamma>\<^sub>2 \<turnstile> t : T" using assms
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   260
proof (nominal_induct \<Gamma>\<^sub>1 t T avoiding: \<Gamma>\<^sub>2 rule: typing.strong_induct)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   261
  case (Abs x T \<Gamma> t U)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   262
  then show ?case
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   263
    by (simp add: typing.Abs valid.Cons)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   264
next
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   265
  case (Let p t \<Gamma> T \<Delta> u U)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   266
  then show ?case
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   267
    by (smt (verit, ccfv_threshold) Un_iff pat_freshs set_append typing.simps valid_app_mono valid_typing) 
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   268
qed auto
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   269
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   270
inductive
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80142
diff changeset
   271
  match :: "pat \<Rightarrow> trm \<Rightarrow> (name \<times> trm) list \<Rightarrow> bool"  (\<open>\<turnstile> _ \<rhd> _ \<Rightarrow> _\<close> [50, 50, 50] 50)
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   272
where
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   273
  PVar: "\<turnstile> PVar x T \<rhd> t \<Rightarrow> [(x, t)]"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   274
| PProd: "\<turnstile> p \<rhd> t \<Rightarrow> \<theta> \<Longrightarrow> \<turnstile> q \<rhd> u \<Rightarrow> \<theta>' \<Longrightarrow> \<turnstile> \<langle>\<langle>p, q\<rangle>\<rangle> \<rhd> \<langle>t, u\<rangle> \<Rightarrow> \<theta> @ \<theta>'"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   275
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   276
fun
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   277
  lookup :: "(name \<times> trm) list \<Rightarrow> name \<Rightarrow> trm"   
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   278
where
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   279
  "lookup [] x = Var x"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   280
| "lookup ((y, e) # \<theta>) x = (if x = y then e else lookup \<theta> x)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   281
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   282
lemma lookup_eqvt[eqvt]:
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   283
  fixes pi :: "name prm"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   284
  and   \<theta> :: "(name \<times> trm) list"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   285
  and   X :: "name"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   286
  shows "pi \<bullet> (lookup \<theta> X) = lookup (pi \<bullet> \<theta>) (pi \<bullet> X)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   287
  by (induct \<theta>) (auto simp add: eqvts)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   288
 
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   289
nominal_primrec
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80142
diff changeset
   290
  psubst :: "(name \<times> trm) list \<Rightarrow> trm \<Rightarrow> trm"  (\<open>_\<lparr>_\<rparr>\<close> [95,0] 210)
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80142
diff changeset
   291
  and psubstb :: "(name \<times> trm) list \<Rightarrow> btrm \<Rightarrow> btrm"  (\<open>_\<lparr>_\<rparr>\<^sub>b\<close> [95,0] 210)
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   292
where
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   293
  "\<theta>\<lparr>Var x\<rparr> = (lookup \<theta> x)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   294
| "\<theta>\<lparr>t \<cdot> u\<rparr> = \<theta>\<lparr>t\<rparr> \<cdot> \<theta>\<lparr>u\<rparr>"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   295
| "\<theta>\<lparr>\<langle>t, u\<rangle>\<rparr> = \<langle>\<theta>\<lparr>t\<rparr>, \<theta>\<lparr>u\<rparr>\<rangle>"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   296
| "\<theta>\<lparr>Let T t u\<rparr> = Let T (\<theta>\<lparr>t\<rparr>) (\<theta>\<lparr>u\<rparr>\<^sub>b)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   297
| "x \<sharp> \<theta> \<Longrightarrow> \<theta>\<lparr>\<lambda>x:T. t\<rparr> = (\<lambda>x:T. \<theta>\<lparr>t\<rparr>)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   298
| "\<theta>\<lparr>Base t\<rparr>\<^sub>b = Base (\<theta>\<lparr>t\<rparr>)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   299
| "x \<sharp> \<theta> \<Longrightarrow> \<theta>\<lparr>Bind T x t\<rparr>\<^sub>b = Bind T x (\<theta>\<lparr>t\<rparr>\<^sub>b)"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   300
  by (finite_guess | simp add: abs_fresh | fresh_guess)+
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   301
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   302
lemma lookup_fresh:
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   303
  "x = y \<longrightarrow> x \<in> set (map fst \<theta>) \<Longrightarrow> \<forall>(y, t)\<in>set \<theta>. x \<sharp> t \<Longrightarrow> x \<sharp> lookup \<theta> y"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   304
  by (induct \<theta>) (use fresh_atm in force)+
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   305
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   306
lemma psubst_fresh:
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   307
  assumes "x \<in> set (map fst \<theta>)" and "\<forall>(y, t)\<in>set \<theta>. x \<sharp> t"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   308
  shows "x \<sharp> \<theta>\<lparr>t\<rparr>" and "x \<sharp> \<theta>\<lparr>t'\<rparr>\<^sub>b" using assms
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   309
proof (nominal_induct t and t' avoiding: \<theta> rule: trm_btrm.strong_inducts)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   310
  case (Var name)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   311
  then show ?case
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   312
    by (metis lookup_fresh simps(1))
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   313
qed (auto simp: abs_fresh)
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   314
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   315
lemma psubst_eqvt[eqvt]:
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   316
  fixes pi :: "name prm" 
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   317
  shows "pi \<bullet> (\<theta>\<lparr>t\<rparr>) = (pi \<bullet> \<theta>)\<lparr>pi \<bullet> t\<rparr>"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   318
  and "pi \<bullet> (\<theta>\<lparr>t'\<rparr>\<^sub>b) = (pi \<bullet> \<theta>)\<lparr>pi \<bullet> t'\<rparr>\<^sub>b"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   319
  by (nominal_induct t and t' avoiding: \<theta> rule: trm_btrm.strong_inducts)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   320
    (simp_all add: eqvts fresh_bij)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   321
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   322
abbreviation 
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80142
diff changeset
   323
  subst :: "trm \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> trm" (\<open>_[_\<mapsto>_]\<close> [100,0,0] 100)
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   324
where 
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   325
  "t[x\<mapsto>t'] \<equiv> [(x,t')]\<lparr>t\<rparr>"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   326
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   327
abbreviation 
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80142
diff changeset
   328
  substb :: "btrm \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> btrm" (\<open>_[_\<mapsto>_]\<^sub>b\<close> [100,0,0] 100)
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   329
where 
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   330
  "t[x\<mapsto>t']\<^sub>b \<equiv> [(x,t')]\<lparr>t\<rparr>\<^sub>b"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   331
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   332
lemma lookup_forget:
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   333
  "(supp (map fst \<theta>)::name set) \<sharp>* x \<Longrightarrow> lookup \<theta> x = Var x"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   334
  by (induct \<theta>) (auto simp add: split_paired_all fresh_star_def fresh_atm supp_list_cons supp_atm)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   335
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   336
lemma supp_fst: "(x::name) \<in> supp (map fst (\<theta>::(name \<times> trm) list)) \<Longrightarrow> x \<in> supp \<theta>"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   337
  by (induct \<theta>) (auto simp add: supp_list_nil supp_list_cons supp_prod)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   338
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   339
lemma psubst_forget:
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   340
  "(supp (map fst \<theta>)::name set) \<sharp>* t \<Longrightarrow> \<theta>\<lparr>t\<rparr> = t"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   341
  "(supp (map fst \<theta>)::name set) \<sharp>* t' \<Longrightarrow> \<theta>\<lparr>t'\<rparr>\<^sub>b = t'"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   342
proof (nominal_induct t and t' avoiding: \<theta> rule: trm_btrm.strong_inducts)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   343
  case (Var name)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   344
  then show ?case
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   345
    by (simp add: fresh_star_set lookup_forget)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   346
next
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   347
  case (Abs ty name trm)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   348
  then show ?case
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   349
    apply (simp add: fresh_def)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   350
    by (metis abs_fresh(1) fresh_star_set supp_fst trm.fresh(3))
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   351
next
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   352
  case (Bind ty name btrm)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   353
  then show ?case
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   354
    apply (simp add: fresh_def)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   355
    by (metis abs_fresh(1) btrm.fresh(2) fresh_star_set supp_fst)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   356
qed (auto simp: fresh_star_set)
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   357
80142
34e0ddfc6dcc More tidying of Nominal proofs
paulson <lp15@cam.ac.uk>
parents: 80140
diff changeset
   358
lemma psubst_nil[simp]: "[]\<lparr>t\<rparr> = t" "[]\<lparr>t'\<rparr>\<^sub>b = t'"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   359
  by (induct t and t' rule: trm_btrm.inducts) (simp_all add: fresh_list_nil)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   360
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   361
lemma psubst_cons:
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   362
  assumes "(supp (map fst \<theta>)::name set) \<sharp>* u"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   363
  shows "((x, u) # \<theta>)\<lparr>t\<rparr> = \<theta>\<lparr>t[x\<mapsto>u]\<rparr>" and "((x, u) # \<theta>)\<lparr>t'\<rparr>\<^sub>b = \<theta>\<lparr>t'[x\<mapsto>u]\<^sub>b\<rparr>\<^sub>b"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   364
  using assms
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   365
  by (nominal_induct t and t' avoiding: x u \<theta> rule: trm_btrm.strong_inducts)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   366
    (simp_all add: fresh_list_nil fresh_list_cons psubst_forget)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   367
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   368
lemma psubst_append:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   369
  "(supp (map fst (\<theta>\<^sub>1 @ \<theta>\<^sub>2))::name set) \<sharp>* map snd (\<theta>\<^sub>1 @ \<theta>\<^sub>2) \<Longrightarrow> (\<theta>\<^sub>1 @ \<theta>\<^sub>2)\<lparr>t\<rparr> = \<theta>\<^sub>2\<lparr>\<theta>\<^sub>1\<lparr>t\<rparr>\<rparr>"
80142
34e0ddfc6dcc More tidying of Nominal proofs
paulson <lp15@cam.ac.uk>
parents: 80140
diff changeset
   370
proof (induct \<theta>\<^sub>1 arbitrary: t)
34e0ddfc6dcc More tidying of Nominal proofs
paulson <lp15@cam.ac.uk>
parents: 80140
diff changeset
   371
  case Nil
34e0ddfc6dcc More tidying of Nominal proofs
paulson <lp15@cam.ac.uk>
parents: 80140
diff changeset
   372
  then show ?case
34e0ddfc6dcc More tidying of Nominal proofs
paulson <lp15@cam.ac.uk>
parents: 80140
diff changeset
   373
    by (auto simp: psubst_nil)
34e0ddfc6dcc More tidying of Nominal proofs
paulson <lp15@cam.ac.uk>
parents: 80140
diff changeset
   374
next
34e0ddfc6dcc More tidying of Nominal proofs
paulson <lp15@cam.ac.uk>
parents: 80140
diff changeset
   375
  case (Cons a \<theta>\<^sub>1)
34e0ddfc6dcc More tidying of Nominal proofs
paulson <lp15@cam.ac.uk>
parents: 80140
diff changeset
   376
  then show ?case
34e0ddfc6dcc More tidying of Nominal proofs
paulson <lp15@cam.ac.uk>
parents: 80140
diff changeset
   377
  proof (cases a)
34e0ddfc6dcc More tidying of Nominal proofs
paulson <lp15@cam.ac.uk>
parents: 80140
diff changeset
   378
    case (Pair a b)
34e0ddfc6dcc More tidying of Nominal proofs
paulson <lp15@cam.ac.uk>
parents: 80140
diff changeset
   379
    with Cons show ?thesis
34e0ddfc6dcc More tidying of Nominal proofs
paulson <lp15@cam.ac.uk>
parents: 80140
diff changeset
   380
      apply (simp add: supp_list_cons fresh_star_set fresh_list_cons)
34e0ddfc6dcc More tidying of Nominal proofs
paulson <lp15@cam.ac.uk>
parents: 80140
diff changeset
   381
      by (metis Un_iff fresh_star_set map_append psubst_cons(1) supp_list_append)
34e0ddfc6dcc More tidying of Nominal proofs
paulson <lp15@cam.ac.uk>
parents: 80140
diff changeset
   382
  qed
34e0ddfc6dcc More tidying of Nominal proofs
paulson <lp15@cam.ac.uk>
parents: 80140
diff changeset
   383
qed
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   384
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   385
lemma abs_pat_psubst [simp]:
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   386
  "(supp p::name set) \<sharp>* \<theta> \<Longrightarrow> \<theta>\<lparr>\<lambda>[p]. t\<rparr>\<^sub>b = (\<lambda>[p]. \<theta>\<lparr>t\<rparr>\<^sub>b)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   387
  by (induct p arbitrary: t) (auto simp add: fresh_star_def supp_atm)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   388
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   389
lemma valid_insert:
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   390
  assumes "valid (\<Delta> @ [(x, T)] @ \<Gamma>)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   391
  shows "valid (\<Delta> @ \<Gamma>)" using assms
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   392
  by (induct \<Delta>)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   393
    (auto simp add: fresh_list_append fresh_list_cons)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   394
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   395
lemma fresh_set: 
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   396
  shows "y \<sharp> xs = (\<forall>x\<in>set xs. y \<sharp> x)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   397
  by (induct xs) (simp_all add: fresh_list_nil fresh_list_cons)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   398
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   399
lemma context_unique:
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   400
  assumes "valid \<Gamma>"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   401
  and "(x, T) \<in> set \<Gamma>"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   402
  and "(x, U) \<in> set \<Gamma>"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   403
  shows "T = U" using assms
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   404
  by induct (auto simp add: fresh_set fresh_prod fresh_atm)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   405
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   406
lemma subst_type_aux:
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   407
  assumes a: "\<Delta> @ [(x, U)] @ \<Gamma> \<turnstile> t : T"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   408
  and b: "\<Gamma> \<turnstile> u : U"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   409
  shows "\<Delta> @ \<Gamma> \<turnstile> t[x\<mapsto>u] : T" using a b
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   410
proof (nominal_induct \<Gamma>'\<equiv>"\<Delta> @ [(x, U)] @ \<Gamma>" t T avoiding: x u \<Delta> rule: typing.strong_induct)
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 33268
diff changeset
   411
  case (Var y T x u \<Delta>)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   412
  from \<open>valid (\<Delta> @ [(x, U)] @ \<Gamma>)\<close>
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 33268
diff changeset
   413
  have valid: "valid (\<Delta> @ \<Gamma>)" by (rule valid_insert)
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   414
  show "\<Delta> @ \<Gamma> \<turnstile> Var y[x\<mapsto>u] : T"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   415
  proof cases
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   416
    assume eq: "x = y"
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 33268
diff changeset
   417
    from Var eq have "T = U" by (auto intro: context_unique)
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 33268
diff changeset
   418
    with Var eq valid show "\<Delta> @ \<Gamma> \<turnstile> Var y[x\<mapsto>u] : T" by (auto intro: weakening)
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   419
  next
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   420
    assume ineq: "x \<noteq> y"
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 33268
diff changeset
   421
    from Var ineq have "(y, T) \<in> set (\<Delta> @ \<Gamma>)" by simp
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 33268
diff changeset
   422
    then show "\<Delta> @ \<Gamma> \<turnstile> Var y[x\<mapsto>u] : T" using ineq valid by auto
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   423
  qed
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   424
next
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   425
  case (Tuple t\<^sub>1 T\<^sub>1 t\<^sub>2 T\<^sub>2)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   426
  from refl \<open>\<Gamma> \<turnstile> u : U\<close>
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   427
  have "\<Delta> @ \<Gamma> \<turnstile> t\<^sub>1[x\<mapsto>u] : T\<^sub>1" by (rule Tuple)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   428
  moreover from refl \<open>\<Gamma> \<turnstile> u : U\<close>
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   429
  have "\<Delta> @ \<Gamma> \<turnstile> t\<^sub>2[x\<mapsto>u] : T\<^sub>2" by (rule Tuple)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   430
  ultimately have "\<Delta> @ \<Gamma> \<turnstile> \<langle>t\<^sub>1[x\<mapsto>u], t\<^sub>2[x\<mapsto>u]\<rangle> : T\<^sub>1 \<otimes> T\<^sub>2" ..
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   431
  then show ?case by simp
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   432
next
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 33268
diff changeset
   433
  case (Let p t T \<Delta>' s S)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   434
  from refl \<open>\<Gamma> \<turnstile> u : U\<close>
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   435
  have "\<Delta> @ \<Gamma> \<turnstile> t[x\<mapsto>u] : T" by (rule Let)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   436
  moreover note \<open>\<turnstile> p : T \<Rightarrow> \<Delta>'\<close>
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 33268
diff changeset
   437
  moreover have "\<Delta>' @ (\<Delta> @ [(x, U)] @ \<Gamma>) = (\<Delta>' @ \<Delta>) @ [(x, U)] @ \<Gamma>" by simp
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   438
  then have "(\<Delta>' @ \<Delta>) @ \<Gamma> \<turnstile> s[x\<mapsto>u] : S" using \<open>\<Gamma> \<turnstile> u : U\<close> by (rule Let)
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   439
  then have "\<Delta>' @ \<Delta> @ \<Gamma> \<turnstile> s[x\<mapsto>u] : S" by simp
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   440
  ultimately have "\<Delta> @ \<Gamma> \<turnstile> (LET p = t[x\<mapsto>u] IN s[x\<mapsto>u]) : S"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   441
    by (rule better_T_Let)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   442
  moreover from Let have "(supp p::name set) \<sharp>* [(x, u)]"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   443
    by (simp add: fresh_star_def fresh_list_nil fresh_list_cons)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   444
  ultimately show ?case by simp
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   445
next
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 33268
diff changeset
   446
  case (Abs y T t S)
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 33268
diff changeset
   447
  have "(y, T) # \<Delta> @ [(x, U)] @ \<Gamma> = ((y, T) # \<Delta>) @ [(x, U)] @ \<Gamma>"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   448
    by simp
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   449
  then have "((y, T) # \<Delta>) @ \<Gamma> \<turnstile> t[x\<mapsto>u] : S" using \<open>\<Gamma> \<turnstile> u : U\<close> by (rule Abs)
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   450
  then have "(y, T) # \<Delta> @ \<Gamma> \<turnstile> t[x\<mapsto>u] : S" by simp
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   451
  then have "\<Delta> @ \<Gamma> \<turnstile> (\<lambda>y:T. t[x\<mapsto>u]) : T \<rightarrow> S"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   452
    by (rule typing.Abs)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   453
  moreover from Abs have "y \<sharp> [(x, u)]"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   454
    by (simp add: fresh_list_nil fresh_list_cons)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   455
  ultimately show ?case by simp
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   456
next
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   457
  case (App t\<^sub>1 T S t\<^sub>2)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   458
  from refl \<open>\<Gamma> \<turnstile> u : U\<close>
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   459
  have "\<Delta> @ \<Gamma> \<turnstile> t\<^sub>1[x\<mapsto>u] : T \<rightarrow> S" by (rule App)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   460
  moreover from refl \<open>\<Gamma> \<turnstile> u : U\<close>
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   461
  have "\<Delta> @ \<Gamma> \<turnstile> t\<^sub>2[x\<mapsto>u] : T" by (rule App)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   462
  ultimately have "\<Delta> @ \<Gamma> \<turnstile> (t\<^sub>1[x\<mapsto>u]) \<cdot> (t\<^sub>2[x\<mapsto>u]) : S"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   463
    by (rule typing.App)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   464
  then show ?case by simp
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   465
qed
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   466
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   467
lemmas subst_type = subst_type_aux [of "[]", simplified]
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   468
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   469
lemma match_supp_fst:
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   470
  assumes "\<turnstile> p \<rhd> u \<Rightarrow> \<theta>" shows "(supp (map fst \<theta>)::name set) = supp p" using assms
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   471
  by induct (simp_all add: supp_list_nil supp_list_cons supp_list_append)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   472
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   473
lemma match_supp_snd:
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   474
  assumes "\<turnstile> p \<rhd> u \<Rightarrow> \<theta>" shows "(supp (map snd \<theta>)::name set) = supp u" using assms
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   475
  by induct (simp_all add: supp_list_nil supp_list_cons supp_list_append trm.supp)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   476
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   477
lemma match_fresh: "\<turnstile> p \<rhd> u \<Rightarrow> \<theta> \<Longrightarrow> (supp p::name set) \<sharp>* u \<Longrightarrow>
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   478
  (supp (map fst \<theta>)::name set) \<sharp>* map snd \<theta>"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   479
  by (simp add: fresh_star_def fresh_def match_supp_fst match_supp_snd)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   480
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   481
lemma match_type_aux:
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   482
  assumes "\<turnstile> p : U \<Rightarrow> \<Delta>"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   483
  and "\<Gamma>\<^sub>2 \<turnstile> u : U"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   484
  and "\<Gamma>\<^sub>1 @ \<Delta> @ \<Gamma>\<^sub>2 \<turnstile> t : T"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   485
  and "\<turnstile> p \<rhd> u \<Rightarrow> \<theta>"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   486
  and "(supp p::name set) \<sharp>* u"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   487
  shows "\<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile> \<theta>\<lparr>t\<rparr> : T" using assms
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   488
proof (induct arbitrary: \<Gamma>\<^sub>1 \<Gamma>\<^sub>2 t u T \<theta>)
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   489
  case (PVar x U)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   490
  from \<open>\<Gamma>\<^sub>1 @ [(x, U)] @ \<Gamma>\<^sub>2 \<turnstile> t : T\<close> \<open>\<Gamma>\<^sub>2 \<turnstile> u : U\<close>
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   491
  have "\<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile> t[x\<mapsto>u] : T" by (rule subst_type_aux)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   492
  moreover from \<open>\<turnstile> PVar x U \<rhd> u \<Rightarrow> \<theta>\<close> have "\<theta> = [(x, u)]"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   493
    by cases simp_all
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   494
  ultimately show ?case by simp
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   495
next
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   496
  case (PTuple p S \<Delta>\<^sub>1 q U \<Delta>\<^sub>2)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   497
  from \<open>\<turnstile> \<langle>\<langle>p, q\<rangle>\<rangle> \<rhd> u \<Rightarrow> \<theta>\<close> obtain u\<^sub>1 u\<^sub>2 \<theta>\<^sub>1 \<theta>\<^sub>2
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   498
    where u: "u = \<langle>u\<^sub>1, u\<^sub>2\<rangle>" and \<theta>: "\<theta> = \<theta>\<^sub>1 @ \<theta>\<^sub>2"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   499
    and p: "\<turnstile> p \<rhd> u\<^sub>1 \<Rightarrow> \<theta>\<^sub>1" and q: "\<turnstile> q \<rhd> u\<^sub>2 \<Rightarrow> \<theta>\<^sub>2"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   500
    by cases simp_all
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   501
  with PTuple have "\<Gamma>\<^sub>2 \<turnstile> \<langle>u\<^sub>1, u\<^sub>2\<rangle> : S \<otimes> U" by simp
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   502
  then obtain u\<^sub>1: "\<Gamma>\<^sub>2 \<turnstile> u\<^sub>1 : S" and u\<^sub>2: "\<Gamma>\<^sub>2 \<turnstile> u\<^sub>2 : U"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   503
    by cases (simp_all add: ty.inject trm.inject)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   504
  note u\<^sub>1
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   505
  moreover from \<open>\<Gamma>\<^sub>1 @ (\<Delta>\<^sub>2 @ \<Delta>\<^sub>1) @ \<Gamma>\<^sub>2 \<turnstile> t : T\<close>
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   506
  have "(\<Gamma>\<^sub>1 @ \<Delta>\<^sub>2) @ \<Delta>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile> t : T" by simp
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   507
  moreover note p
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   508
  moreover from \<open>supp \<langle>\<langle>p, q\<rangle>\<rangle> \<sharp>* u\<close> and u
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   509
  have "(supp p::name set) \<sharp>* u\<^sub>1" by (simp add: fresh_star_def)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   510
  ultimately have \<theta>\<^sub>1: "(\<Gamma>\<^sub>1 @ \<Delta>\<^sub>2) @ \<Gamma>\<^sub>2 \<turnstile> \<theta>\<^sub>1\<lparr>t\<rparr> : T"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   511
    by (rule PTuple)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   512
  note u\<^sub>2
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   513
  moreover from \<theta>\<^sub>1
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   514
  have "\<Gamma>\<^sub>1 @ \<Delta>\<^sub>2 @ \<Gamma>\<^sub>2 \<turnstile> \<theta>\<^sub>1\<lparr>t\<rparr> : T" by simp
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   515
  moreover note q
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   516
  moreover from \<open>supp \<langle>\<langle>p, q\<rangle>\<rangle> \<sharp>* u\<close> and u
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   517
  have "(supp q::name set) \<sharp>* u\<^sub>2" by (simp add: fresh_star_def)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   518
  ultimately have "\<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile> \<theta>\<^sub>2\<lparr>\<theta>\<^sub>1\<lparr>t\<rparr>\<rparr> : T"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   519
    by (rule PTuple)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   520
  moreover from \<open>\<turnstile> \<langle>\<langle>p, q\<rangle>\<rangle> \<rhd> u \<Rightarrow> \<theta>\<close> \<open>supp \<langle>\<langle>p, q\<rangle>\<rangle> \<sharp>* u\<close>
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   521
  have "(supp (map fst \<theta>)::name set) \<sharp>* map snd \<theta>"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   522
    by (rule match_fresh)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   523
  ultimately show ?case using \<theta> by (simp add: psubst_append)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   524
qed
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   525
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   526
lemmas match_type = match_type_aux [where \<Gamma>\<^sub>1="[]", simplified]
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   527
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 80142
diff changeset
   528
inductive eval :: "trm \<Rightarrow> trm \<Rightarrow> bool" (\<open>_ \<longmapsto> _\<close> [60,60] 60)
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   529
where
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   530
  TupleL: "t \<longmapsto> t' \<Longrightarrow> \<langle>t, u\<rangle> \<longmapsto> \<langle>t', u\<rangle>"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   531
| TupleR: "u \<longmapsto> u' \<Longrightarrow> \<langle>t, u\<rangle> \<longmapsto> \<langle>t, u'\<rangle>"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   532
| Abs: "t \<longmapsto> t' \<Longrightarrow> (\<lambda>x:T. t) \<longmapsto> (\<lambda>x:T. t')"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   533
| AppL: "t \<longmapsto> t' \<Longrightarrow> t \<cdot> u \<longmapsto> t' \<cdot> u"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   534
| AppR: "u \<longmapsto> u' \<Longrightarrow> t \<cdot> u \<longmapsto> t \<cdot> u'"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   535
| Beta: "x \<sharp> u \<Longrightarrow> (\<lambda>x:T. t) \<cdot> u \<longmapsto> t[x\<mapsto>u]"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   536
| Let: "((supp p)::name set) \<sharp>* t \<Longrightarrow> distinct (pat_vars p) \<Longrightarrow>
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   537
    \<turnstile> p \<rhd> t \<Rightarrow> \<theta> \<Longrightarrow> (LET p = t IN u) \<longmapsto> \<theta>\<lparr>u\<rparr>"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   538
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   539
equivariance match
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   540
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   541
equivariance eval
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   542
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   543
lemma match_vars:
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   544
  assumes "\<turnstile> p \<rhd> t \<Rightarrow> \<theta>" and "x \<in> supp p"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   545
  shows "x \<in> set (map fst \<theta>)" using assms
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   546
  by induct (auto simp add: supp_atm)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   547
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   548
lemma match_fresh_mono:
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   549
  assumes "\<turnstile> p \<rhd> t \<Rightarrow> \<theta>" and "(x::name) \<sharp> t"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   550
  shows "\<forall>(y, t)\<in>set \<theta>. x \<sharp> t" using assms
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   551
  by induct auto
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   552
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   553
nominal_inductive2 eval
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   554
avoids
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   555
  Abs: "{x}"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   556
| Beta: "{x}"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   557
| Let: "(supp p)::name set"
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   558
proof (simp_all add: fresh_star_def abs_fresh fin_supp)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   559
  show "x \<sharp> t[x\<mapsto>u]" if "x \<sharp> u" for x t u
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   560
    by (simp add: \<open>x \<sharp> u\<close> psubst_fresh(1))
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   561
next
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   562
  show "\<forall>x\<in>supp p. (x::name) \<sharp> \<theta>\<lparr>u\<rparr>" 
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   563
    if "\<forall>x\<in>supp p. (x::name) \<sharp> t" and "\<turnstile> p \<rhd> t \<Rightarrow> \<theta>"
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   564
    for p t \<theta> u
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   565
    by (meson that match_fresh_mono match_vars psubst_fresh(1))
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   566
qed
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   567
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   568
lemma typing_case_Abs:
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   569
  assumes ty: "\<Gamma> \<turnstile> (\<lambda>x:T. t) : S"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   570
  and fresh: "x \<sharp> \<Gamma>"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   571
  and R: "\<And>U. S = T \<rightarrow> U \<Longrightarrow> (x, T) # \<Gamma> \<turnstile> t : U \<Longrightarrow> P"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   572
  shows P using ty
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   573
proof cases
34990
81e8fdfeb849 Adapted to changes in cases method.
berghofe
parents: 34915
diff changeset
   574
  case (Abs x' T' t' U)
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   575
  obtain y::name where y: "y \<sharp> (x, \<Gamma>, \<lambda>x':T'. t')"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   576
    by (rule exists_fresh) (auto intro: fin_supp)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   577
  from \<open>(\<lambda>x:T. t) = (\<lambda>x':T'. t')\<close> [symmetric]
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   578
  have x: "x \<sharp> (\<lambda>x':T'. t')" by (simp add: abs_fresh)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   579
  have x': "x' \<sharp> (\<lambda>x':T'. t')" by (simp add: abs_fresh)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   580
  from \<open>(x', T') # \<Gamma> \<turnstile> t' : U\<close> have x'': "x' \<sharp> \<Gamma>"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   581
    by (auto dest: valid_typing)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   582
  have "(\<lambda>x:T. t) = (\<lambda>x':T'. t')" by fact
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   583
  also from x x' y have "\<dots> = [(x, y)] \<bullet> [(x', y)] \<bullet> (\<lambda>x':T'. t')"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   584
    by (simp only: perm_fresh_fresh fresh_prod)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   585
  also have "\<dots> = (\<lambda>x:T'. [(x, y)] \<bullet> [(x', y)] \<bullet> t')"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   586
    by (simp add: swap_simps perm_fresh_fresh)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   587
  finally have "(\<lambda>x:T. t) = (\<lambda>x:T'. [(x, y)] \<bullet> [(x', y)] \<bullet> t')" .
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   588
  then have T: "T = T'" and t: "[(x, y)] \<bullet> [(x', y)] \<bullet> t' = t"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   589
    by (simp_all add: trm.inject alpha)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   590
  from Abs T have "S = T \<rightarrow> U" by simp
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   591
  moreover from \<open>(x', T') # \<Gamma> \<turnstile> t' : U\<close>
34990
81e8fdfeb849 Adapted to changes in cases method.
berghofe
parents: 34915
diff changeset
   592
  have "[(x, y)] \<bullet> [(x', y)] \<bullet> ((x', T') # \<Gamma> \<turnstile> t' : U)"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   593
    by (simp add: perm_bool)
34990
81e8fdfeb849 Adapted to changes in cases method.
berghofe
parents: 34915
diff changeset
   594
  with T t y x'' fresh have "(x, T) # \<Gamma> \<turnstile> t : U"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   595
    by (simp add: eqvts swap_simps perm_fresh_fresh fresh_prod)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   596
  ultimately show ?thesis by (rule R)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   597
qed simp_all
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   598
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   599
nominal_primrec ty_size :: "ty \<Rightarrow> nat"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   600
where
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   601
  "ty_size (Atom n) = 0"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   602
| "ty_size (T \<rightarrow> U) = ty_size T + ty_size U + 1"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   603
| "ty_size (T \<otimes> U) = ty_size T + ty_size U + 1"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   604
  by (rule TrueI)+
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   605
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   606
lemma bind_tuple_ineq:
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   607
  "ty_size (pat_type p) < ty_size U \<Longrightarrow> Bind U x t \<noteq> (\<lambda>[p]. u)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   608
  by (induct p arbitrary: U x t u) (auto simp add: btrm.inject)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   609
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   610
lemma valid_appD: assumes "valid (\<Gamma> @ \<Delta>)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   611
  shows "valid \<Gamma>" "valid \<Delta>" using assms
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   612
  by (induct \<Gamma>'\<equiv>"\<Gamma> @ \<Delta>" arbitrary: \<Gamma> \<Delta>)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   613
    (auto simp add: Cons_eq_append_conv fresh_list_append)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   614
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   615
lemma valid_app_freshs: assumes "valid (\<Gamma> @ \<Delta>)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   616
  shows "(supp \<Gamma>::name set) \<sharp>* \<Delta>" "(supp \<Delta>::name set) \<sharp>* \<Gamma>" using assms
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   617
  by (induct \<Gamma>'\<equiv>"\<Gamma> @ \<Delta>" arbitrary: \<Gamma> \<Delta>)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   618
    (auto simp add: Cons_eq_append_conv fresh_star_def
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   619
     fresh_list_nil fresh_list_cons supp_list_nil supp_list_cons fresh_list_append
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   620
     supp_prod fresh_prod supp_atm fresh_atm
44687
20deab90494e tuned proof
haftmann
parents: 41798
diff changeset
   621
     dest: notE [OF iffD1 [OF fresh_def]])
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   622
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   623
lemma perm_mem_left: "(x::name) \<in> ((pi::name prm) \<bullet> A) \<Longrightarrow> (rev pi \<bullet> x) \<in> A"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   624
  by (drule perm_boolI [of _ "rev pi"]) (simp add: eqvts perm_pi_simp)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   625
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   626
lemma perm_mem_right: "(rev (pi::name prm) \<bullet> (x::name)) \<in> A \<Longrightarrow> x \<in> (pi \<bullet> A)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   627
  by (drule perm_boolI [of _ pi]) (simp add: eqvts perm_pi_simp)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   628
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   629
lemma perm_cases:
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   630
  assumes pi: "set pi \<subseteq> A \<times> A"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   631
  shows "((pi::name prm) \<bullet> B) \<subseteq> A \<union> B"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   632
proof
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   633
  fix x assume "x \<in> pi \<bullet> B"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   634
  then show "x \<in> A \<union> B" using pi
80140
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   635
  proof (induct pi arbitrary: x B rule: rev_induct)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   636
    case Nil
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   637
    then show ?case
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   638
      by simp
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   639
  next
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   640
    case (snoc y xs)
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   641
    then show ?case
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   642
      apply simp
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   643
      by (metis SigmaE perm_mem_left perm_pi_simp(2) pt_name2 swap_simps(3))
6d56e85f674e More proof tidying for Nominal
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   644
  qed
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   645
qed
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   646
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   647
lemma abs_pat_alpha':
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   648
  assumes eq: "(\<lambda>[p]. t) = (\<lambda>[q]. u)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   649
  and ty: "pat_type p = pat_type q"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   650
  and pv: "distinct (pat_vars p)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   651
  and qv: "distinct (pat_vars q)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   652
  shows "\<exists>pi::name prm. p = pi \<bullet> q \<and> t = pi \<bullet> u \<and>
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   653
    set pi \<subseteq> (supp p \<union> supp q) \<times> (supp p \<union> supp q)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   654
  using assms
45129
1fce03e3e8ad tuned proofs -- eliminated vacuous "induct arbitrary: ..." situations;
wenzelm
parents: 44687
diff changeset
   655
proof (induct p arbitrary: q t u)
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   656
  case (PVar x T)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   657
  note PVar' = this
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   658
  show ?case
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   659
  proof (cases q)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   660
    case (PVar x' T')
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   661
    with \<open>(\<lambda>[PVar x T]. t) = (\<lambda>[q]. u)\<close>
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   662
    have "x = x' \<and> t = u \<or> x \<noteq> x' \<and> t = [(x, x')] \<bullet> u \<and> x \<sharp> u"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   663
      by (simp add: btrm.inject alpha)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   664
    then show ?thesis
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   665
    proof
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   666
      assume "x = x' \<and> t = u"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   667
      with PVar PVar' have "PVar x T = ([]::name prm) \<bullet> q \<and>
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33189
diff changeset
   668
        t = ([]::name prm) \<bullet> u \<and>
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33189
diff changeset
   669
        set ([]::name prm) \<subseteq> (supp (PVar x T) \<union> supp q) \<times>
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   670
          (supp (PVar x T) \<union> supp q)" by simp
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   671
      then show ?thesis ..
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   672
    next
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   673
      assume "x \<noteq> x' \<and> t = [(x, x')] \<bullet> u \<and> x \<sharp> u"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   674
      with PVar PVar' have "PVar x T = [(x, x')] \<bullet> q \<and>
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33189
diff changeset
   675
        t = [(x, x')] \<bullet> u \<and>
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33189
diff changeset
   676
        set [(x, x')] \<subseteq> (supp (PVar x T) \<union> supp q) \<times>
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   677
          (supp (PVar x T) \<union> supp q)"
33268
02de0317f66f eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 33189
diff changeset
   678
        by (simp add: perm_swap swap_simps supp_atm perm_type)
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   679
      then show ?thesis ..
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   680
    qed
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   681
  next
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   682
    case (PTuple p\<^sub>1 p\<^sub>2)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   683
    with PVar have "ty_size (pat_type p\<^sub>1) < ty_size T" by simp
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   684
    then have "Bind T x t \<noteq> (\<lambda>[p\<^sub>1]. \<lambda>[p\<^sub>2]. u)"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   685
      by (rule bind_tuple_ineq)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   686
    moreover from PTuple PVar
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   687
    have "Bind T x t = (\<lambda>[p\<^sub>1]. \<lambda>[p\<^sub>2]. u)" by simp
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   688
    ultimately show ?thesis ..
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   689
  qed
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   690
next
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   691
  case (PTuple p\<^sub>1 p\<^sub>2)
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   692
  note PTuple' = this
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   693
  show ?case
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   694
  proof (cases q)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   695
    case (PVar x T)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   696
    with PTuple have "ty_size (pat_type p\<^sub>1) < ty_size T" by auto
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   697
    then have "Bind T x u \<noteq> (\<lambda>[p\<^sub>1]. \<lambda>[p\<^sub>2]. t)"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   698
      by (rule bind_tuple_ineq)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   699
    moreover from PTuple PVar
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   700
    have "Bind T x u = (\<lambda>[p\<^sub>1]. \<lambda>[p\<^sub>2]. t)" by simp
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   701
    ultimately show ?thesis ..
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   702
  next
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   703
    case (PTuple p\<^sub>1' p\<^sub>2')
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   704
    with PTuple' have "(\<lambda>[p\<^sub>1]. \<lambda>[p\<^sub>2]. t) = (\<lambda>[p\<^sub>1']. \<lambda>[p\<^sub>2']. u)" by simp
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   705
    moreover from PTuple PTuple' have "pat_type p\<^sub>1 = pat_type p\<^sub>1'"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   706
      by (simp add: ty.inject)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   707
    moreover from PTuple' have "distinct (pat_vars p\<^sub>1)" by simp
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   708
    moreover from PTuple PTuple' have "distinct (pat_vars p\<^sub>1')" by simp
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   709
    ultimately have "\<exists>pi::name prm. p\<^sub>1 = pi \<bullet> p\<^sub>1' \<and>
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   710
      (\<lambda>[p\<^sub>2]. t) = pi \<bullet> (\<lambda>[p\<^sub>2']. u) \<and>
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   711
      set pi \<subseteq> (supp p\<^sub>1 \<union> supp p\<^sub>1') \<times> (supp p\<^sub>1 \<union> supp p\<^sub>1')"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   712
      by (rule PTuple')
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   713
    then obtain pi::"name prm" where
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   714
      "p\<^sub>1 = pi \<bullet> p\<^sub>1'" "(\<lambda>[p\<^sub>2]. t) = pi \<bullet> (\<lambda>[p\<^sub>2']. u)" and
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   715
      pi: "set pi \<subseteq> (supp p\<^sub>1 \<union> supp p\<^sub>1') \<times> (supp p\<^sub>1 \<union> supp p\<^sub>1')" by auto
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   716
    from \<open>(\<lambda>[p\<^sub>2]. t) = pi \<bullet> (\<lambda>[p\<^sub>2']. u)\<close>
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   717
    have "(\<lambda>[p\<^sub>2]. t) = (\<lambda>[pi \<bullet> p\<^sub>2']. pi \<bullet> u)"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   718
      by (simp add: eqvts)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   719
    moreover from PTuple PTuple' have "pat_type p\<^sub>2 = pat_type (pi \<bullet> p\<^sub>2')"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   720
      by (simp add: ty.inject pat_type_perm_eq)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   721
    moreover from PTuple' have "distinct (pat_vars p\<^sub>2)" by simp
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   722
    moreover from PTuple PTuple' have "distinct (pat_vars (pi \<bullet> p\<^sub>2'))"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   723
      by (simp add: pat_vars_eqvt [symmetric] distinct_eqvt [symmetric])
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   724
    ultimately have "\<exists>pi'::name prm. p\<^sub>2 = pi' \<bullet> pi \<bullet> p\<^sub>2' \<and>
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   725
      t = pi' \<bullet> pi \<bullet> u \<and>
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   726
      set pi' \<subseteq> (supp p\<^sub>2 \<union> supp (pi \<bullet> p\<^sub>2')) \<times> (supp p\<^sub>2 \<union> supp (pi \<bullet> p\<^sub>2'))"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   727
      by (rule PTuple')
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   728
    then obtain pi'::"name prm" where
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   729
      "p\<^sub>2 = pi' \<bullet> pi \<bullet> p\<^sub>2'" "t = pi' \<bullet> pi \<bullet> u" and
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   730
      pi': "set pi' \<subseteq> (supp p\<^sub>2 \<union> supp (pi \<bullet> p\<^sub>2')) \<times>
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   731
        (supp p\<^sub>2 \<union> supp (pi \<bullet> p\<^sub>2'))" by auto
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   732
    from PTuple PTuple' have "pi \<bullet> distinct (pat_vars \<langle>\<langle>p\<^sub>1', p\<^sub>2'\<rangle>\<rangle>)" by simp
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   733
    then have "distinct (pat_vars \<langle>\<langle>pi \<bullet> p\<^sub>1', pi \<bullet> p\<^sub>2'\<rangle>\<rangle>)" by (simp only: eqvts)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   734
    with \<open>p\<^sub>1 = pi \<bullet> p\<^sub>1'\<close> PTuple'
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   735
    have fresh: "(supp p\<^sub>2 \<union> supp (pi \<bullet> p\<^sub>2') :: name set) \<sharp>* p\<^sub>1"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   736
      by (auto simp add: set_pat_vars_supp fresh_star_def fresh_def eqvts)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   737
    from \<open>p\<^sub>1 = pi \<bullet> p\<^sub>1'\<close> have "pi' \<bullet> (p\<^sub>1 = pi \<bullet> p\<^sub>1')" by (rule perm_boolI)
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   738
    with pt_freshs_freshs [OF pt_name_inst at_name_inst pi' fresh fresh]
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   739
    have "p\<^sub>1 = pi' \<bullet> pi \<bullet> p\<^sub>1'" by (simp add: eqvts)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   740
    with \<open>p\<^sub>2 = pi' \<bullet> pi \<bullet> p\<^sub>2'\<close> have "\<langle>\<langle>p\<^sub>1, p\<^sub>2\<rangle>\<rangle> = (pi' @ pi) \<bullet> \<langle>\<langle>p\<^sub>1', p\<^sub>2'\<rangle>\<rangle>"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   741
      by (simp add: pt_name2)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   742
    moreover
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   743
    have "((supp p\<^sub>2 \<union> (pi \<bullet> supp p\<^sub>2')) \<times> (supp p\<^sub>2 \<union> (pi \<bullet> supp p\<^sub>2'))::(name \<times> name) set) \<subseteq>
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   744
      (supp p\<^sub>2 \<union> (supp p\<^sub>1 \<union> supp p\<^sub>1' \<union> supp p\<^sub>2')) \<times> (supp p\<^sub>2 \<union> (supp p\<^sub>1 \<union> supp p\<^sub>1' \<union> supp p\<^sub>2'))"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   745
      by (rule subset_refl Sigma_mono Un_mono perm_cases [OF pi])+
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   746
    with pi' have "set pi' \<subseteq> \<dots>" by (simp add: supp_eqvt [symmetric])
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   747
    with pi have "set (pi' @ pi) \<subseteq> (supp \<langle>\<langle>p\<^sub>1, p\<^sub>2\<rangle>\<rangle> \<union> supp \<langle>\<langle>p\<^sub>1', p\<^sub>2'\<rangle>\<rangle>) \<times>
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   748
      (supp \<langle>\<langle>p\<^sub>1, p\<^sub>2\<rangle>\<rangle> \<union> supp \<langle>\<langle>p\<^sub>1', p\<^sub>2'\<rangle>\<rangle>)"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   749
      by (simp add: Sigma_Un_distrib1 Sigma_Un_distrib2 Un_ac) blast
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   750
    moreover note \<open>t = pi' \<bullet> pi \<bullet> u\<close>
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   751
    ultimately have "\<langle>\<langle>p\<^sub>1, p\<^sub>2\<rangle>\<rangle> = (pi' @ pi) \<bullet> q \<and> t = (pi' @ pi) \<bullet> u \<and>
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   752
      set (pi' @ pi) \<subseteq> (supp \<langle>\<langle>p\<^sub>1, p\<^sub>2\<rangle>\<rangle> \<union> supp q) \<times>
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   753
        (supp \<langle>\<langle>p\<^sub>1, p\<^sub>2\<rangle>\<rangle> \<union> supp q)" using PTuple
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   754
      by (simp add: pt_name2)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   755
    then show ?thesis ..
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   756
  qed
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   757
qed
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   758
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   759
lemma typing_case_Let:
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   760
  assumes ty: "\<Gamma> \<turnstile> (LET p = t IN u) : U"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   761
  and fresh: "(supp p::name set) \<sharp>* \<Gamma>"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   762
  and distinct: "distinct (pat_vars p)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   763
  and R: "\<And>T \<Delta>. \<Gamma> \<turnstile> t : T \<Longrightarrow> \<turnstile> p : T \<Rightarrow> \<Delta> \<Longrightarrow> \<Delta> @ \<Gamma> \<turnstile> u : U \<Longrightarrow> P"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   764
  shows P using ty
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   765
proof cases
34990
81e8fdfeb849 Adapted to changes in cases method.
berghofe
parents: 34915
diff changeset
   766
  case (Let p' t' T \<Delta> u')
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   767
  then have "(supp \<Delta>::name set) \<sharp>* \<Gamma>"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   768
    by (auto intro: valid_typing valid_app_freshs)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   769
  with Let have "(supp p'::name set) \<sharp>* \<Gamma>"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   770
    by (simp add: pat_var)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   771
  with fresh have fresh': "(supp p \<union> supp p' :: name set) \<sharp>* \<Gamma>"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   772
    by (auto simp add: fresh_star_def)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   773
  from Let have "(\<lambda>[p]. Base u) = (\<lambda>[p']. Base u')"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   774
    by (simp add: trm.inject)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   775
  moreover from Let have "pat_type p = pat_type p'"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   776
    by (simp add: trm.inject)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   777
  moreover note distinct
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   778
  moreover from \<open>\<Delta> @ \<Gamma> \<turnstile> u' : U\<close> have "valid (\<Delta> @ \<Gamma>)"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   779
    by (rule valid_typing)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   780
  then have "valid \<Delta>" by (rule valid_appD)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   781
  with \<open>\<turnstile> p' : T \<Rightarrow> \<Delta>\<close> have "distinct (pat_vars p')"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   782
    by (simp add: valid_distinct pat_vars_ptyping)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   783
  ultimately have "\<exists>pi::name prm. p = pi \<bullet> p' \<and> Base u = pi \<bullet> Base u' \<and>
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   784
    set pi \<subseteq> (supp p \<union> supp p') \<times> (supp p \<union> supp p')"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   785
    by (rule abs_pat_alpha')
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   786
  then obtain pi::"name prm" where pi: "p = pi \<bullet> p'" "u = pi \<bullet> u'"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   787
    and pi': "set pi \<subseteq> (supp p \<union> supp p') \<times> (supp p \<union> supp p')"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   788
    by (auto simp add: btrm.inject)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   789
  from Let have "\<Gamma> \<turnstile> t : T" by (simp add: trm.inject)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   790
  moreover from \<open>\<turnstile> p' : T \<Rightarrow> \<Delta>\<close> have "\<turnstile> (pi \<bullet> p') : (pi \<bullet> T) \<Rightarrow> (pi \<bullet> \<Delta>)"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   791
    by (simp add: ptyping.eqvt)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   792
  with pi have "\<turnstile> p : T \<Rightarrow> (pi \<bullet> \<Delta>)" by (simp add: perm_type)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   793
  moreover from Let
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   794
  have "(pi \<bullet> \<Delta>) @ (pi \<bullet> \<Gamma>) \<turnstile> (pi \<bullet> u') : (pi \<bullet> U)"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   795
    by (simp add: append_eqvt [symmetric] typing.eqvt)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   796
  with pi have "(pi \<bullet> \<Delta>) @ \<Gamma> \<turnstile> u : U"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   797
    by (simp add: perm_type pt_freshs_freshs
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   798
      [OF pt_name_inst at_name_inst pi' fresh' fresh'])
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   799
  ultimately show ?thesis by (rule R)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   800
qed simp_all
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   801
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   802
lemma preservation:
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   803
  assumes "t \<longmapsto> t'" and "\<Gamma> \<turnstile> t : T"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   804
  shows "\<Gamma> \<turnstile> t' : T" using assms
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   805
proof (nominal_induct avoiding: \<Gamma> T rule: eval.strong_induct)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   806
  case (TupleL t t' u)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   807
  from \<open>\<Gamma> \<turnstile> \<langle>t, u\<rangle> : T\<close> obtain T\<^sub>1 T\<^sub>2
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   808
    where "T = T\<^sub>1 \<otimes> T\<^sub>2" "\<Gamma> \<turnstile> t : T\<^sub>1" "\<Gamma> \<turnstile> u : T\<^sub>2"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   809
    by cases (simp_all add: trm.inject)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   810
  from \<open>\<Gamma> \<turnstile> t : T\<^sub>1\<close> have "\<Gamma> \<turnstile> t' : T\<^sub>1" by (rule TupleL)
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   811
  then have "\<Gamma> \<turnstile> \<langle>t', u\<rangle> : T\<^sub>1 \<otimes> T\<^sub>2" using \<open>\<Gamma> \<turnstile> u : T\<^sub>2\<close>
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   812
    by (rule Tuple)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   813
  with \<open>T = T\<^sub>1 \<otimes> T\<^sub>2\<close> show ?case by simp
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   814
next
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   815
  case (TupleR u u' t)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   816
  from \<open>\<Gamma> \<turnstile> \<langle>t, u\<rangle> : T\<close> obtain T\<^sub>1 T\<^sub>2
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 45129
diff changeset
   817
    where "T = T\<^sub>1 \<otimes> T\<^sub>2" "\<Gamma> \<turnstile> t : T\<^sub>1" "\<Gamma> \<turnstile> u : T\<^sub>2"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   818
    by cases (simp_all add: trm.inject)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   819
  from \<open>\<Gamma> \<turnstile> u : T\<^sub>2\<close> have "\<Gamma> \<turnstile> u' : T\<^sub>2" by (rule TupleR)
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   820
  with \<open>\<Gamma> \<turnstile> t : T\<^sub>1\<close> have "\<Gamma> \<turnstile> \<langle>t, u'\<rangle> : T\<^sub>1 \<otimes> T\<^sub>2"
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   821
    by (rule Tuple)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   822
  with \<open>T = T\<^sub>1 \<otimes> T\<^sub>2\<close> show ?case by simp
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   823
next
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   824
  case (Abs t t' x S)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   825
  from \<open>\<Gamma> \<turnstile> (\<lambda>x:S. t) : T\<close> \<open>x \<sharp> \<Gamma>\<close> obtain U where
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   826
    T: "T = S \<rightarrow> U" and U: "(x, S) # \<Gamma> \<turnstile> t : U"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   827
    by (rule typing_case_Abs)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   828
  from U have "(x, S) # \<Gamma> \<turnstile> t' : U" by (rule Abs)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   829
  then have "\<Gamma> \<turnstile> (\<lambda>x:S. t') : S \<rightarrow> U"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   830
    by (rule typing.Abs)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   831
  with T show ?case by simp
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   832
next
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   833
  case (Beta x u S t)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   834
  from \<open>\<Gamma> \<turnstile> (\<lambda>x:S. t) \<cdot> u : T\<close> \<open>x \<sharp> \<Gamma>\<close>
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   835
  obtain "(x, S) # \<Gamma> \<turnstile> t : T" and "\<Gamma> \<turnstile> u : S"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   836
    by cases (auto simp add: trm.inject ty.inject elim: typing_case_Abs)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   837
  then show ?case by (rule subst_type)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   838
next
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   839
  case (Let p t \<theta> u)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   840
  from \<open>\<Gamma> \<turnstile> (LET p = t IN u) : T\<close> \<open>supp p \<sharp>* \<Gamma>\<close> \<open>distinct (pat_vars p)\<close>
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   841
  obtain U \<Delta> where "\<turnstile> p : U \<Rightarrow> \<Delta>" "\<Gamma> \<turnstile> t : U" "\<Delta> @ \<Gamma> \<turnstile> u : T"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   842
    by (rule typing_case_Let)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   843
  then show ?case using \<open>\<turnstile> p \<rhd> t \<Rightarrow> \<theta>\<close> \<open>supp p \<sharp>* t\<close>
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   844
    by (rule match_type)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   845
next
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   846
  case (AppL t t' u)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   847
  from \<open>\<Gamma> \<turnstile> t \<cdot> u : T\<close> obtain U where
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   848
    t: "\<Gamma> \<turnstile> t : U \<rightarrow> T" and u: "\<Gamma> \<turnstile> u : U"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   849
    by cases (auto simp add: trm.inject)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   850
  from t have "\<Gamma> \<turnstile> t' : U \<rightarrow> T" by (rule AppL)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   851
  then show ?case using u by (rule typing.App)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   852
next
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   853
  case (AppR u u' t)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   854
  from \<open>\<Gamma> \<turnstile> t \<cdot> u : T\<close> obtain U where
33189
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   855
    t: "\<Gamma> \<turnstile> t : U \<rightarrow> T" and u: "\<Gamma> \<turnstile> u : U"
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   856
    by cases (auto simp add: trm.inject)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   857
  from u have "\<Gamma> \<turnstile> u' : U" by (rule AppR)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   858
  with t show ?case by (rule typing.App)
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   859
qed
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   860
82a40677c1f8 Added Pattern.thy to Nominal/Examples.
berghofe
parents:
diff changeset
   861
end