Nominal-General/Atoms.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Tue, 04 May 2010 12:34:33 +0200
changeset 2042 495b6feb76a8
parent 1774 c34347ec7ab3
child 2470 bdb1eab47161
permissions -rw-r--r--
Move ExLeroy to New Parser
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
(*  Title:      Atoms
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
    Authors:    Brian Huffman, Christian Urban
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
    Instantiations of concrete atoms 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
*)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
theory Atoms
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
imports Nominal2_Atoms
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
begin
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
section {* Manual instantiation of class @{text at}. *}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
typedef (open) name = "{a. sort_of a = Sort ''name'' []}"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
by (rule exists_eq_simple_sort)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
instantiation name :: at
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
begin
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
definition
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
  "p \<bullet> a = Abs_name (p \<bullet> Rep_name a)"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
definition
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
  "atom a = Rep_name a"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
instance
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
apply (rule at_class)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
apply (rule type_definition_name)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
apply (rule atom_name_def)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
apply (rule permute_name_def)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
done
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
end
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
lemma sort_of_atom_name: 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
  shows "sort_of (atom (a::name)) = Sort ''name'' []"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
  unfolding atom_name_def using Rep_name by simp
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
text {* Custom syntax for concrete atoms of type at *}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
term "a:::name"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
text {* 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
  a:::name stands for (atom a) with a being of concrete atom 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
  type name. The above lemma can therefore also be stated as
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
  "sort_of (a:::name) = Sort ''name'' []"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
  This does not work for multi-sorted atoms.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
*}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
section {* Automatic instantiation of class @{text at}. *}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
atom_decl name2
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
lemma sort_of_atom_name2:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
  "sort_of (atom (a::name2)) = Sort ''Atoms.name2'' []"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
unfolding atom_name2_def 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
using Rep_name2 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
by simp
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
text {* example swappings *}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
lemma
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
  fixes a b::"atom"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
  assumes "sort_of a = sort_of b"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
  shows "(a \<rightleftharpoons> b) \<bullet> (a, b) = (b, a)"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
using assms
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
by simp
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    69
lemma
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
  fixes a b::"name2"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    71
  shows "(a \<leftrightarrow> b) \<bullet> (a, b) = (b, a)"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    72
by simp
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
section {* An example for multiple-sort atoms *}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    76
datatype ty =
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    77
  TVar string
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
| Fun ty ty ("_ \<rightarrow> _")
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80
primrec
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
  sort_of_ty::"ty \<Rightarrow> atom_sort"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
where
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
  "sort_of_ty (TVar s) = Sort ''TVar'' [Sort s []]"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    84
| "sort_of_ty (Fun ty1 ty2) = Sort ''Fun'' [sort_of_ty ty1, sort_of_ty ty2]"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
lemma sort_of_ty_eq_iff:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
  shows "sort_of_ty x = sort_of_ty y \<longleftrightarrow> x = y"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
apply(induct x arbitrary: y)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
apply(case_tac [!] y)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
apply(simp_all)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
done
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
declare sort_of_ty.simps [simp del]
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    95
typedef (open) var = "{a. sort_of a \<in> range sort_of_ty}"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    96
  by (rule_tac x="Atom (sort_of_ty x) y" in exI, simp)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
instantiation var :: at_base
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
begin
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   101
definition
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
  "p \<bullet> a = Abs_var (p \<bullet> Rep_var a)"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
definition
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
  "atom a = Rep_var a"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   107
instance
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
apply (rule at_base_class)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   109
apply (rule type_definition_var)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
apply (rule atom_var_def)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
apply (rule permute_var_def)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
done
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   114
end
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
text {* Constructor for variables. *}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
definition
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
  Var :: "nat \<Rightarrow> ty \<Rightarrow> var"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
where
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
  "Var x t = Abs_var (Atom (sort_of_ty t) x)"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
lemma Var_eq_iff [simp]:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
  shows "Var x s = Var y t \<longleftrightarrow> x = y \<and> s = t"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
  unfolding Var_def
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
  by (auto simp add: Abs_var_inject sort_of_ty_eq_iff)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
lemma sort_of_atom_var [simp]:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
  "sort_of (atom (Var n ty)) = sort_of_ty ty"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
  unfolding atom_var_def Var_def
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
  by (simp add: Abs_var_inverse)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
lemma 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   134
  assumes "\<alpha> \<noteq> \<beta>" 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   135
  shows "(Var x \<alpha> \<leftrightarrow> Var y \<alpha>) \<bullet> (Var x \<alpha>, Var x \<beta>) = (Var y \<alpha>, Var x \<beta>)"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   136
  using assms by simp
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   137
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   138
text {* Projecting out the type component of a variable. *}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   139
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   140
definition
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   141
  ty_of :: "var \<Rightarrow> ty"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   142
where
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   143
  "ty_of x = inv sort_of_ty (sort_of (atom x))"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   144
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   145
text {*
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   146
  Functions @{term Var}/@{term ty_of} satisfy many of the same
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   147
  properties as @{term Atom}/@{term sort_of}.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   148
*}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   150
lemma ty_of_Var [simp]:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   151
  shows "ty_of (Var x t) = t"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   152
  unfolding ty_of_def
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   153
  unfolding sort_of_atom_var
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
  apply (rule inv_f_f)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
  apply (simp add: inj_on_def sort_of_ty_eq_iff)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   156
  done
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
lemma ty_of_permute [simp]:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
  shows "ty_of (p \<bullet> x) = ty_of x"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   160
  unfolding ty_of_def
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   161
  unfolding atom_eqvt [symmetric]
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   162
  by simp
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   163
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   164
end