Nominal/Atoms.thy
author Christian Urban <urbanc@in.tum.de>
Wed, 01 Jun 2011 22:55:14 +0100
changeset 2803 04f7c4ce8588
parent 2782 2cb34b1e7e19
child 2891 304dfe6cc83a
permissions -rw-r--r--
hopefully final fix for ho-functions
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
*)
2782
2cb34b1e7e19 made the subtyping work again
Christian Urban <urbanc@in.tum.de>
parents: 2781
diff changeset
     6
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
theory Atoms
2470
bdb1eab47161 moved everything out of Nominal_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
     8
imports Nominal2_Base
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
begin
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
2470
bdb1eab47161 moved everything out of Nominal_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    11
bdb1eab47161 moved everything out of Nominal_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    12
bdb1eab47161 moved everything out of Nominal_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    13
section {* @{const nat_of} is an example of a function 
bdb1eab47161 moved everything out of Nominal_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    14
  without finite support *}
bdb1eab47161 moved everything out of Nominal_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    15
bdb1eab47161 moved everything out of Nominal_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    16
bdb1eab47161 moved everything out of Nominal_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    17
lemma not_fresh_nat_of:
bdb1eab47161 moved everything out of Nominal_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    18
  shows "\<not> a \<sharp> nat_of"
bdb1eab47161 moved everything out of Nominal_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    19
unfolding fresh_def supp_def
bdb1eab47161 moved everything out of Nominal_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    20
proof (clarsimp)
bdb1eab47161 moved everything out of Nominal_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    21
  assume "finite {b. (a \<rightleftharpoons> b) \<bullet> nat_of \<noteq> nat_of}"
bdb1eab47161 moved everything out of Nominal_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    22
  hence "finite ({a} \<union> {b. (a \<rightleftharpoons> b) \<bullet> nat_of \<noteq> nat_of})"
bdb1eab47161 moved everything out of Nominal_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    23
    by simp
bdb1eab47161 moved everything out of Nominal_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    24
  then obtain b where
bdb1eab47161 moved everything out of Nominal_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    25
    b1: "b \<noteq> a" and
bdb1eab47161 moved everything out of Nominal_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    26
    b2: "sort_of b = sort_of a" and
bdb1eab47161 moved everything out of Nominal_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    27
    b3: "(a \<rightleftharpoons> b) \<bullet> nat_of = nat_of"
bdb1eab47161 moved everything out of Nominal_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    28
    by (rule obtain_atom) auto
bdb1eab47161 moved everything out of Nominal_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    29
  have "nat_of a = (a \<rightleftharpoons> b) \<bullet> (nat_of a)" by (simp add: permute_nat_def)
bdb1eab47161 moved everything out of Nominal_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    30
  also have "\<dots> = ((a \<rightleftharpoons> b) \<bullet> nat_of) ((a \<rightleftharpoons> b) \<bullet> a)" by (simp add: permute_fun_app_eq)
bdb1eab47161 moved everything out of Nominal_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    31
  also have "\<dots> = nat_of ((a \<rightleftharpoons> b) \<bullet> a)" using b3 by simp
bdb1eab47161 moved everything out of Nominal_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    32
  also have "\<dots> = nat_of b" using b2 by simp
bdb1eab47161 moved everything out of Nominal_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    33
  finally have "nat_of a = nat_of b" by simp
bdb1eab47161 moved everything out of Nominal_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    34
  with b2 have "a = b" by (simp add: atom_components_eq_iff)
bdb1eab47161 moved everything out of Nominal_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    35
  with b1 show "False" by simp
bdb1eab47161 moved everything out of Nominal_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    36
qed
bdb1eab47161 moved everything out of Nominal_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    37
bdb1eab47161 moved everything out of Nominal_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    38
lemma supp_nat_of:
bdb1eab47161 moved everything out of Nominal_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    39
  shows "supp nat_of = UNIV"
bdb1eab47161 moved everything out of Nominal_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    40
  using not_fresh_nat_of [unfolded fresh_def] by auto
bdb1eab47161 moved everything out of Nominal_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    41
bdb1eab47161 moved everything out of Nominal_Supp
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    42
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
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
    44
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
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
    46
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
    47
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
instantiation name :: at
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
begin
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
definition
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
  "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
    53
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
definition
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
  "atom a = Rep_name a"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
instance
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
apply (rule at_class)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
apply (rule type_definition_name)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
apply (rule atom_name_def)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
apply (rule permute_name_def)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
done
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
end
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
lemma sort_of_atom_name: 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
  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
    68
  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
    69
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
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
    71
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    72
term "a:::name"
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
text {* 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
  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
    76
  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
    77
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
  "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
    79
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80
  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
    81
*}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
2744
56b8d977d1c0 more on the pearl paper
Christian Urban <urbanc@in.tum.de>
parents: 2597
diff changeset
    83
(* fixme 
56b8d977d1c0 more on the pearl paper
Christian Urban <urbanc@in.tum.de>
parents: 2597
diff changeset
    84
lemma supp_of_finite_name_set:
56b8d977d1c0 more on the pearl paper
Christian Urban <urbanc@in.tum.de>
parents: 2597
diff changeset
    85
  fixes S::"name set"
56b8d977d1c0 more on the pearl paper
Christian Urban <urbanc@in.tum.de>
parents: 2597
diff changeset
    86
  assumes "infinte S" "infinite (UNIV - S)"
56b8d977d1c0 more on the pearl paper
Christian Urban <urbanc@in.tum.de>
parents: 2597
diff changeset
    87
  shows "supp S = UNIV"
56b8d977d1c0 more on the pearl paper
Christian Urban <urbanc@in.tum.de>
parents: 2597
diff changeset
    88
proof -
56b8d977d1c0 more on the pearl paper
Christian Urban <urbanc@in.tum.de>
parents: 2597
diff changeset
    89
  { fix a    
56b8d977d1c0 more on the pearl paper
Christian Urban <urbanc@in.tum.de>
parents: 2597
diff changeset
    90
    have "a \<in> supp S"
56b8d977d1c0 more on the pearl paper
Christian Urban <urbanc@in.tum.de>
parents: 2597
diff changeset
    91
    proof (cases "a \<in> atom ` S")
56b8d977d1c0 more on the pearl paper
Christian Urban <urbanc@in.tum.de>
parents: 2597
diff changeset
    92
      assume "a \<in> atom ` S"
56b8d977d1c0 more on the pearl paper
Christian Urban <urbanc@in.tum.de>
parents: 2597
diff changeset
    93
      then have "\<forall>b \<in> UNIV - S. (a \<rightleftharpoons> atom b) \<bullet> S \<noteq> S"
56b8d977d1c0 more on the pearl paper
Christian Urban <urbanc@in.tum.de>
parents: 2597
diff changeset
    94
	apply(clarify)
56b8d977d1c0 more on the pearl paper
Christian Urban <urbanc@in.tum.de>
parents: 2597
diff changeset
    95
	
56b8d977d1c0 more on the pearl paper
Christian Urban <urbanc@in.tum.de>
parents: 2597
diff changeset
    96
      show "a \<in> supp S"
56b8d977d1c0 more on the pearl paper
Christian Urban <urbanc@in.tum.de>
parents: 2597
diff changeset
    97
  }
56b8d977d1c0 more on the pearl paper
Christian Urban <urbanc@in.tum.de>
parents: 2597
diff changeset
    98
  then show "supp S = UNIV" by auto
56b8d977d1c0 more on the pearl paper
Christian Urban <urbanc@in.tum.de>
parents: 2597
diff changeset
    99
qed
56b8d977d1c0 more on the pearl paper
Christian Urban <urbanc@in.tum.de>
parents: 2597
diff changeset
   100
*)
56b8d977d1c0 more on the pearl paper
Christian Urban <urbanc@in.tum.de>
parents: 2597
diff changeset
   101
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
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
   104
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
atom_decl name2
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
lemma sort_of_atom_name2:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
  "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
   109
unfolding atom_name2_def 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
using Rep_name2 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
by simp
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
text {* example swappings *}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   114
lemma
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
  fixes a b::"atom"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
  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
   117
  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
   118
using assms
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
by simp
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
lemma
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
  fixes a b::"name2"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
  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
   124
by simp
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
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
   127
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
datatype ty =
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
  TVar string
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
| Fun ty ty ("_ \<rightarrow> _")
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
primrec
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
  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
   134
where
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   135
  "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
   136
| "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
   137
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   138
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
   139
  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
   140
apply(induct x arbitrary: y)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   141
apply(case_tac [!] y)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   142
apply(simp_all)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   143
done
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
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
   146
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   147
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
   148
  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
   149
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   150
instantiation var :: at_base
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   151
begin
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   152
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   153
definition
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
  "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
   155
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   156
definition
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
  "atom a = Rep_var a"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
instance
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   160
apply (rule at_base_class)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   161
apply (rule type_definition_var)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   162
apply (rule atom_var_def)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   163
apply (rule permute_var_def)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   164
done
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   165
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   166
end
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   167
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   168
text {* Constructor for variables. *}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   169
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   170
definition
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   171
  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
   172
where
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   173
  "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
   174
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
lemma Var_eq_iff [simp]:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
  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
   177
  unfolding Var_def
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   178
  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
   179
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   180
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
   181
  "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
   182
  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
   183
  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
   184
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
lemma 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   186
  assumes "\<alpha> \<noteq> \<beta>" 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   187
  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
   188
  using assms by simp
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   190
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
   191
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   192
definition
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   193
  ty_of :: "var \<Rightarrow> ty"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   194
where
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
  "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
   196
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   197
text {*
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   198
  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
   199
  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
   200
*}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   202
lemma ty_of_Var [simp]:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   203
  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
   204
  unfolding ty_of_def
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   205
  unfolding sort_of_atom_var
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
  apply (rule inv_f_f)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   207
  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
   208
  done
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   209
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   210
lemma ty_of_permute [simp]:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   211
  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
   212
  unfolding ty_of_def
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   213
  unfolding atom_eqvt [symmetric]
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   214
  by simp
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   215
2556
8ed62410236e added a test about subtyping; disabled two tests, because of problem with function package
Christian Urban <urbanc@in.tum.de>
parents: 2470
diff changeset
   216
8ed62410236e added a test about subtyping; disabled two tests, because of problem with function package
Christian Urban <urbanc@in.tum.de>
parents: 2470
diff changeset
   217
section {* Tests with subtyping and automatic coercions *}
8ed62410236e added a test about subtyping; disabled two tests, because of problem with function package
Christian Urban <urbanc@in.tum.de>
parents: 2470
diff changeset
   218
2597
0f289a52edbe updated to changes in Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 2568
diff changeset
   219
declare [[coercion_enabled]]
2556
8ed62410236e added a test about subtyping; disabled two tests, because of problem with function package
Christian Urban <urbanc@in.tum.de>
parents: 2470
diff changeset
   220
8ed62410236e added a test about subtyping; disabled two tests, because of problem with function package
Christian Urban <urbanc@in.tum.de>
parents: 2470
diff changeset
   221
atom_decl var1
8ed62410236e added a test about subtyping; disabled two tests, because of problem with function package
Christian Urban <urbanc@in.tum.de>
parents: 2470
diff changeset
   222
atom_decl var2
8ed62410236e added a test about subtyping; disabled two tests, because of problem with function package
Christian Urban <urbanc@in.tum.de>
parents: 2470
diff changeset
   223
8ed62410236e added a test about subtyping; disabled two tests, because of problem with function package
Christian Urban <urbanc@in.tum.de>
parents: 2470
diff changeset
   224
declare [[coercion "atom::var1\<Rightarrow>atom"]]
8ed62410236e added a test about subtyping; disabled two tests, because of problem with function package
Christian Urban <urbanc@in.tum.de>
parents: 2470
diff changeset
   225
8ed62410236e added a test about subtyping; disabled two tests, because of problem with function package
Christian Urban <urbanc@in.tum.de>
parents: 2470
diff changeset
   226
declare [[coercion "atom::var2\<Rightarrow>atom"]]
8ed62410236e added a test about subtyping; disabled two tests, because of problem with function package
Christian Urban <urbanc@in.tum.de>
parents: 2470
diff changeset
   227
8ed62410236e added a test about subtyping; disabled two tests, because of problem with function package
Christian Urban <urbanc@in.tum.de>
parents: 2470
diff changeset
   228
lemma
8ed62410236e added a test about subtyping; disabled two tests, because of problem with function package
Christian Urban <urbanc@in.tum.de>
parents: 2470
diff changeset
   229
  fixes a::"var1" and b::"var2"
8ed62410236e added a test about subtyping; disabled two tests, because of problem with function package
Christian Urban <urbanc@in.tum.de>
parents: 2470
diff changeset
   230
  shows "a \<sharp> t \<and> b \<sharp> t"
8ed62410236e added a test about subtyping; disabled two tests, because of problem with function package
Christian Urban <urbanc@in.tum.de>
parents: 2470
diff changeset
   231
oops
2782
2cb34b1e7e19 made the subtyping work again
Christian Urban <urbanc@in.tum.de>
parents: 2781
diff changeset
   232
2556
8ed62410236e added a test about subtyping; disabled two tests, because of problem with function package
Christian Urban <urbanc@in.tum.de>
parents: 2470
diff changeset
   233
8ed62410236e added a test about subtyping; disabled two tests, because of problem with function package
Christian Urban <urbanc@in.tum.de>
parents: 2470
diff changeset
   234
(* does not yet work: it needs image as
8ed62410236e added a test about subtyping; disabled two tests, because of problem with function package
Christian Urban <urbanc@in.tum.de>
parents: 2470
diff changeset
   235
   coercion map *)
8ed62410236e added a test about subtyping; disabled two tests, because of problem with function package
Christian Urban <urbanc@in.tum.de>
parents: 2470
diff changeset
   236
8ed62410236e added a test about subtyping; disabled two tests, because of problem with function package
Christian Urban <urbanc@in.tum.de>
parents: 2470
diff changeset
   237
lemma
8ed62410236e added a test about subtyping; disabled two tests, because of problem with function package
Christian Urban <urbanc@in.tum.de>
parents: 2470
diff changeset
   238
  fixes as::"var1 set"
8ed62410236e added a test about subtyping; disabled two tests, because of problem with function package
Christian Urban <urbanc@in.tum.de>
parents: 2470
diff changeset
   239
  shows "atom ` as \<sharp>* t"
8ed62410236e added a test about subtyping; disabled two tests, because of problem with function package
Christian Urban <urbanc@in.tum.de>
parents: 2470
diff changeset
   240
oops
8ed62410236e added a test about subtyping; disabled two tests, because of problem with function package
Christian Urban <urbanc@in.tum.de>
parents: 2470
diff changeset
   241
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   242
end