src/HOL/Nominal/Examples/Lam_Funs.thy
author wenzelm
Tue, 10 Jul 2007 23:29:43 +0200
changeset 23719 ccd9cb15c062
parent 22832 6a16085eaaa1
child 25751 a4e69ce247e0
permissions -rw-r--r--
more markup for inner and outer syntax; added enclose;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21107
diff changeset
     1
(* $Id$ *)
21107
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
     2
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
     3
theory Lam_Funs
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
     4
imports "../Nominal"
21107
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
     5
begin
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
     6
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
     7
atom_decl name
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
     8
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22540
diff changeset
     9
nominal_datatype lam = 
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22540
diff changeset
    10
    Var "name"
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22540
diff changeset
    11
  | App "lam" "lam"
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22540
diff changeset
    12
  | Lam "\<guillemotleft>name\<guillemotright>lam" ("Lam [_]._" [100,100] 100)
21107
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
    13
21557
3c8e29a6e4f0 added the function for free variables of a lambda-term, which is a
urbanc
parents: 21555
diff changeset
    14
text {* depth of a lambda-term *}
3c8e29a6e4f0 added the function for free variables of a lambda-term, which is a
urbanc
parents: 21555
diff changeset
    15
21555
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21404
diff changeset
    16
consts 
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21404
diff changeset
    17
  depth :: "lam \<Rightarrow> nat"
21107
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
    18
21555
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21404
diff changeset
    19
nominal_primrec
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
    20
  "depth (Var x) = (1::nat)"
21555
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21404
diff changeset
    21
  "depth (App t1 t2) = (max (depth t1) (depth t2)) + 1"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
    22
  "depth (Lam [a].t) = (depth t) + (1::nat)"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
    23
  apply(finite_guess)+
21555
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21404
diff changeset
    24
  apply(rule TrueI)+
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21404
diff changeset
    25
  apply(simp add: fresh_nat)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
    26
  apply(fresh_guess)+
21555
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21404
diff changeset
    27
  done
21107
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
    28
21557
3c8e29a6e4f0 added the function for free variables of a lambda-term, which is a
urbanc
parents: 21555
diff changeset
    29
text {* free variables of a lambda-term *}
3c8e29a6e4f0 added the function for free variables of a lambda-term, which is a
urbanc
parents: 21555
diff changeset
    30
3c8e29a6e4f0 added the function for free variables of a lambda-term, which is a
urbanc
parents: 21555
diff changeset
    31
consts 
3c8e29a6e4f0 added the function for free variables of a lambda-term, which is a
urbanc
parents: 21555
diff changeset
    32
  frees :: "lam \<Rightarrow> name set"
3c8e29a6e4f0 added the function for free variables of a lambda-term, which is a
urbanc
parents: 21555
diff changeset
    33
3c8e29a6e4f0 added the function for free variables of a lambda-term, which is a
urbanc
parents: 21555
diff changeset
    34
nominal_primrec (invariant: "\<lambda>s::name set. finite s")
3c8e29a6e4f0 added the function for free variables of a lambda-term, which is a
urbanc
parents: 21555
diff changeset
    35
  "frees (Var a) = {a}"
3c8e29a6e4f0 added the function for free variables of a lambda-term, which is a
urbanc
parents: 21555
diff changeset
    36
  "frees (App t1 t2) = (frees t1) \<union> (frees t2)"
3c8e29a6e4f0 added the function for free variables of a lambda-term, which is a
urbanc
parents: 21555
diff changeset
    37
  "frees (Lam [a].t) = (frees t) - {a}"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
    38
apply(finite_guess)+
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
    39
apply(simp)+ 
21557
3c8e29a6e4f0 added the function for free variables of a lambda-term, which is a
urbanc
parents: 21555
diff changeset
    40
apply(simp add: fresh_def)
3c8e29a6e4f0 added the function for free variables of a lambda-term, which is a
urbanc
parents: 21555
diff changeset
    41
apply(simp add: supp_of_fin_sets[OF pt_name_inst, OF at_name_inst, OF fs_at_inst[OF at_name_inst]])
3c8e29a6e4f0 added the function for free variables of a lambda-term, which is a
urbanc
parents: 21555
diff changeset
    42
apply(simp add: supp_atm)
3c8e29a6e4f0 added the function for free variables of a lambda-term, which is a
urbanc
parents: 21555
diff changeset
    43
apply(force)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
    44
apply(fresh_guess)+
21557
3c8e29a6e4f0 added the function for free variables of a lambda-term, which is a
urbanc
parents: 21555
diff changeset
    45
done
3c8e29a6e4f0 added the function for free variables of a lambda-term, which is a
urbanc
parents: 21555
diff changeset
    46
3c8e29a6e4f0 added the function for free variables of a lambda-term, which is a
urbanc
parents: 21555
diff changeset
    47
lemma frees_equals_support:
3c8e29a6e4f0 added the function for free variables of a lambda-term, which is a
urbanc
parents: 21555
diff changeset
    48
  shows "frees t = supp t"
3c8e29a6e4f0 added the function for free variables of a lambda-term, which is a
urbanc
parents: 21555
diff changeset
    49
by (nominal_induct t rule: lam.induct)
3c8e29a6e4f0 added the function for free variables of a lambda-term, which is a
urbanc
parents: 21555
diff changeset
    50
   (simp_all add: lam.supp supp_atm abs_supp)
3c8e29a6e4f0 added the function for free variables of a lambda-term, which is a
urbanc
parents: 21555
diff changeset
    51
21107
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
    52
text {* capture-avoiding substitution *}
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
    53
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
    54
21555
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21404
diff changeset
    55
consts
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21404
diff changeset
    56
  subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam"  ("_[_::=_]" [100,100,100] 100)
21107
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
    57
21555
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21404
diff changeset
    58
nominal_primrec
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21404
diff changeset
    59
  "(Var x)[y::=t'] = (if x=y then t' else (Var x))"
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21404
diff changeset
    60
  "(App t1 t2)[y::=t'] = App (t1[y::=t']) (t2[y::=t'])"
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21404
diff changeset
    61
  "x\<sharp>(y,t') \<Longrightarrow> (Lam [x].t)[y::=t'] = Lam [x].(t[y::=t'])"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
    62
apply(finite_guess)+
21555
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21404
diff changeset
    63
apply(rule TrueI)+
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21404
diff changeset
    64
apply(simp add: abs_fresh)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
    65
apply(fresh_guess)+
21107
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
    66
done
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
    67
22540
e4817fa0f6a1 adapted to new nominal_inductive
urbanc
parents: 22492
diff changeset
    68
lemma subst_eqvt[eqvt]:
21107
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
    69
  fixes pi:: "name prm"
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
    70
  shows "pi\<bullet>(t1[b::=t2]) = (pi\<bullet>t1)[(pi\<bullet>b)::=(pi\<bullet>t2)]"
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
    71
apply(nominal_induct t1 avoiding: b t2 rule: lam.induct)
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
    72
apply(auto simp add: perm_bij fresh_prod fresh_atm fresh_bij)
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
    73
done
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
    74
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
    75
lemma subst_supp: 
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
    76
  shows "supp(t1[a::=t2]) \<subseteq> (((supp(t1)-{a})\<union>supp(t2))::name set)"
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
    77
apply(nominal_induct t1 avoiding: a t2 rule: lam.induct)
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
    78
apply(auto simp add: lam.supp supp_atm fresh_prod abs_supp)
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
    79
apply(blast)+
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
    80
done
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
    81
21555
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21404
diff changeset
    82
text{* parallel substitution *}
21107
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
    83
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
    84
consts
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
    85
 psubst :: "(name\<times>lam) list \<Rightarrow> lam \<Rightarrow> lam" ("_<_>" [100,100] 900)
21107
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
    86
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
    87
fun
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
    88
  lookup :: "(name\<times>lam) list \<Rightarrow> name \<Rightarrow> lam"   
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
    89
where
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
    90
  "lookup [] x        = Var x"
22492
43545e640877 Unified function syntax
krauss
parents: 22418
diff changeset
    91
| "lookup ((y,T)#\<theta>) x = (if x=y then T else lookup \<theta> x)"
21107
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
    92
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
    93
lemma lookup_eqvt[eqvt]:
21107
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
    94
  fixes pi::"name prm"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
    95
  shows "pi\<bullet>(lookup \<theta> x) = lookup (pi\<bullet>\<theta>) (pi\<bullet>x)"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
    96
by (induct \<theta>) (auto simp add: perm_bij)
21107
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
    97
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
    98
lemma lookup_fresh:
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
    99
  fixes z::"name"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
   100
  assumes "z\<sharp>\<theta>" "z\<sharp>x"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
   101
  shows "z\<sharp> lookup \<theta> x"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
   102
using assms 
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
   103
by (induct rule: lookup.induct) (auto simp add: fresh_list_cons)
21107
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
   104
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
   105
lemma lookup_fresh':
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
   106
  assumes "z\<sharp>\<theta>"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
   107
  shows "lookup \<theta> z = Var z"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
   108
using assms 
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
   109
by (induct rule: lookup.induct)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
   110
   (auto simp add: fresh_list_cons fresh_prod fresh_atm)
21107
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
   111
21555
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21404
diff changeset
   112
nominal_primrec
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
   113
  "\<theta><(Var x)> = (lookup \<theta> x)"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
   114
  "\<theta><(App t1 t2)> = App (\<theta><t1>) (\<theta><t2>)"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
   115
  "x\<sharp>\<theta>\<Longrightarrow>\<theta><(Lam [x].t)> = Lam [x].(\<theta><t>)"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
   116
apply(finite_guess)+
21555
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21404
diff changeset
   117
apply(rule TrueI)+
229c0158de92 adapted function definitions to new syntax
urbanc
parents: 21404
diff changeset
   118
apply(simp add: abs_fresh)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 21557
diff changeset
   119
apply(fresh_guess)+
21107
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
   120
done
e69c0e82955a new file for defining functions in the lambda-calculus
urbanc
parents:
diff changeset
   121
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21107
diff changeset
   122
end