Nominal-General/Nominal2_Atoms.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Tue, 15 Jun 2010 07:58:33 +0200
changeset 2265 9c44db3eef95
parent 2129 f38adea0591c
permissions -rw-r--r--
Remove only reference to 'equivp'.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
     1
(*  Title:      Nominal2_Atoms
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
     2
    Authors:    Brian Huffman, Christian Urban
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
     3
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
     4
    Definitions for concrete atom types. 
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
     5
*)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
     6
theory Nominal2_Atoms
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
     7
imports Nominal2_Base
1933
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
     8
        Nominal2_Eqvt
1079
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
     9
uses ("nominal_atoms.ML")
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    10
begin
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    11
1972
40db835442a0 deleted left-over code
Christian Urban <urbanc@in.tum.de>
parents: 1971
diff changeset
    12
section {* Infrastructure for concrete atom types *}
1933
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
    13
1941
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
    14
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    15
section {* A swapping operation for concrete atoms *}
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    16
  
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    17
definition
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    18
  flip :: "'a::at_base \<Rightarrow> 'a \<Rightarrow> perm" ("'(_ \<leftrightarrow> _')")
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    19
where
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    20
  "(a \<leftrightarrow> b) = (atom a \<rightleftharpoons> atom b)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    21
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    22
lemma flip_self [simp]: "(a \<leftrightarrow> a) = 0"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    23
  unfolding flip_def by (rule swap_self)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    24
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    25
lemma flip_commute: "(a \<leftrightarrow> b) = (b \<leftrightarrow> a)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    26
  unfolding flip_def by (rule swap_commute)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    27
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    28
lemma minus_flip [simp]: "- (a \<leftrightarrow> b) = (a \<leftrightarrow> b)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    29
  unfolding flip_def by (rule minus_swap)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    30
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    31
lemma add_flip_cancel: "(a \<leftrightarrow> b) + (a \<leftrightarrow> b) = 0"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    32
  unfolding flip_def by (rule swap_cancel)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    33
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    34
lemma permute_flip_cancel [simp]: "(a \<leftrightarrow> b) \<bullet> (a \<leftrightarrow> b) \<bullet> x = x"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    35
  unfolding permute_plus [symmetric] add_flip_cancel by simp
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    36
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    37
lemma permute_flip_cancel2 [simp]: "(a \<leftrightarrow> b) \<bullet> (b \<leftrightarrow> a) \<bullet> x = x"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    38
  by (simp add: flip_commute)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    39
2129
f38adea0591c added flip_eqvt and swap_eqvt to the equivariance lists
Christian Urban <urbanc@in.tum.de>
parents: 1972
diff changeset
    40
lemma flip_eqvt[eqvt]: 
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    41
  fixes a b c::"'a::at_base"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    42
  shows "p \<bullet> (a \<leftrightarrow> b) = (p \<bullet> a \<leftrightarrow> p \<bullet> b)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    43
  unfolding flip_def
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    44
  by (simp add: swap_eqvt atom_eqvt)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    45
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    46
lemma flip_at_base_simps [simp]:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    47
  shows "sort_of (atom a) = sort_of (atom b) \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> a = b"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    48
  and   "sort_of (atom a) = sort_of (atom b) \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> b = a"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    49
  and   "\<lbrakk>a \<noteq> c; b \<noteq> c\<rbrakk> \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> c = c"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    50
  and   "sort_of (atom a) \<noteq> sort_of (atom b) \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> x = x"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    51
  unfolding flip_def
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    52
  unfolding atom_eq_iff [symmetric]
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    53
  unfolding atom_eqvt [symmetric]
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    54
  by simp_all
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    55
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    56
text {* the following two lemmas do not hold for at_base, 
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    57
  only for single sort atoms from at *}
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    58
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    59
lemma permute_flip_at:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    60
  fixes a b c::"'a::at"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    61
  shows "(a \<leftrightarrow> b) \<bullet> c = (if c = a then b else if c = b then a else c)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    62
  unfolding flip_def
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    63
  apply (rule atom_eq_iff [THEN iffD1])
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    64
  apply (subst atom_eqvt [symmetric])
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    65
  apply (simp add: swap_atom)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    66
  done
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    67
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    68
lemma flip_at_simps [simp]:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    69
  fixes a b::"'a::at"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    70
  shows "(a \<leftrightarrow> b) \<bullet> a = b" 
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    71
  and   "(a \<leftrightarrow> b) \<bullet> b = a"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    72
  unfolding permute_flip_at by simp_all
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    73
1499
21dda372fb11 simplified strong induction proof by using flip
Christian Urban <urbanc@in.tum.de>
parents: 1363
diff changeset
    74
lemma flip_fresh_fresh:
21dda372fb11 simplified strong induction proof by using flip
Christian Urban <urbanc@in.tum.de>
parents: 1363
diff changeset
    75
  fixes a b::"'a::at_base"
21dda372fb11 simplified strong induction proof by using flip
Christian Urban <urbanc@in.tum.de>
parents: 1363
diff changeset
    76
  assumes "atom a \<sharp> x" "atom b \<sharp> x"
21dda372fb11 simplified strong induction proof by using flip
Christian Urban <urbanc@in.tum.de>
parents: 1363
diff changeset
    77
  shows "(a \<leftrightarrow> b) \<bullet> x = x"
21dda372fb11 simplified strong induction proof by using flip
Christian Urban <urbanc@in.tum.de>
parents: 1363
diff changeset
    78
using assms
21dda372fb11 simplified strong induction proof by using flip
Christian Urban <urbanc@in.tum.de>
parents: 1363
diff changeset
    79
by (simp add: flip_def swap_fresh_fresh)
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    80
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    81
subsection {* Syntax for coercing at-elements to the atom-type *}
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    82
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    83
syntax
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    84
  "_atom_constrain" :: "logic \<Rightarrow> type \<Rightarrow> logic" ("_:::_" [4, 0] 3)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    85
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    86
translations
1569
1694f32b480a some tuning
Christian Urban <urbanc@in.tum.de>
parents: 1499
diff changeset
    87
  "_atom_constrain a t" => "CONST atom (_constrain a t)"
1694f32b480a some tuning
Christian Urban <urbanc@in.tum.de>
parents: 1499
diff changeset
    88
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    89
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    90
subsection {* A lemma for proving instances of class @{text at}. *}
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    91
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    92
setup {* Sign.add_const_constraint (@{const_name "permute"}, NONE) *}
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    93
setup {* Sign.add_const_constraint (@{const_name "atom"}, NONE) *}
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    94
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    95
text {*
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    96
  New atom types are defined as subtypes of @{typ atom}.
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    97
*}
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    98
1079
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
    99
lemma exists_eq_simple_sort: 
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   100
  shows "\<exists>a. a \<in> {a. sort_of a = s}"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   101
  by (rule_tac x="Atom s 0" in exI, simp)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   102
1079
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   103
lemma exists_eq_sort: 
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   104
  shows "\<exists>a. a \<in> {a. sort_of a \<in> range sort_fun}"
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   105
  by (rule_tac x="Atom (sort_fun x) y" in exI, simp)
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   106
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   107
lemma at_base_class:
1079
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   108
  fixes sort_fun :: "'b \<Rightarrow>atom_sort"
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   109
  fixes Rep :: "'a \<Rightarrow> atom" and Abs :: "atom \<Rightarrow> 'a"
1079
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   110
  assumes type: "type_definition Rep Abs {a. sort_of a \<in> range sort_fun}"
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   111
  assumes atom_def: "\<And>a. atom a = Rep a"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   112
  assumes permute_def: "\<And>p a. p \<bullet> a = Abs (p \<bullet> Rep a)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   113
  shows "OFCLASS('a, at_base_class)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   114
proof
1079
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   115
  interpret type_definition Rep Abs "{a. sort_of a \<in> range sort_fun}" by (rule type)
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   116
  have sort_of_Rep: "\<And>a. sort_of (Rep a) \<in> range sort_fun" using Rep by simp
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   117
  fix a b :: 'a and p p1 p2 :: perm
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   118
  show "0 \<bullet> a = a"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   119
    unfolding permute_def by (simp add: Rep_inverse)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   120
  show "(p1 + p2) \<bullet> a = p1 \<bullet> p2 \<bullet> a"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   121
    unfolding permute_def by (simp add: Abs_inverse sort_of_Rep)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   122
  show "atom a = atom b \<longleftrightarrow> a = b"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   123
    unfolding atom_def by (simp add: Rep_inject)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   124
  show "p \<bullet> atom a = atom (p \<bullet> a)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   125
    unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   126
qed
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   127
1079
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   128
(*
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   129
lemma at_class:
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   130
  fixes s :: atom_sort
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   131
  fixes Rep :: "'a \<Rightarrow> atom" and Abs :: "atom \<Rightarrow> 'a"
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   132
  assumes type: "type_definition Rep Abs {a. sort_of a \<in> range (\<lambda>x::unit. s)}"
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   133
  assumes atom_def: "\<And>a. atom a = Rep a"
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   134
  assumes permute_def: "\<And>p a. p \<bullet> a = Abs (p \<bullet> Rep a)"
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   135
  shows "OFCLASS('a, at_class)"
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   136
proof
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   137
  interpret type_definition Rep Abs "{a. sort_of a \<in> range (\<lambda>x::unit. s)}" by (rule type)
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   138
  have sort_of_Rep: "\<And>a. sort_of (Rep a) = s" using Rep by (simp add: image_def)
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   139
  fix a b :: 'a and p p1 p2 :: perm
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   140
  show "0 \<bullet> a = a"
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   141
    unfolding permute_def by (simp add: Rep_inverse)
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   142
  show "(p1 + p2) \<bullet> a = p1 \<bullet> p2 \<bullet> a"
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   143
    unfolding permute_def by (simp add: Abs_inverse sort_of_Rep)
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   144
  show "sort_of (atom a) = sort_of (atom b)"
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   145
    unfolding atom_def by (simp add: sort_of_Rep)
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   146
  show "atom a = atom b \<longleftrightarrow> a = b"
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   147
    unfolding atom_def by (simp add: Rep_inject)
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   148
  show "p \<bullet> atom a = atom (p \<bullet> a)"
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   149
    unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep)
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   150
qed
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   151
*)
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   152
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   153
lemma at_class:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   154
  fixes s :: atom_sort
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   155
  fixes Rep :: "'a \<Rightarrow> atom" and Abs :: "atom \<Rightarrow> 'a"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   156
  assumes type: "type_definition Rep Abs {a. sort_of a = s}"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   157
  assumes atom_def: "\<And>a. atom a = Rep a"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   158
  assumes permute_def: "\<And>p a. p \<bullet> a = Abs (p \<bullet> Rep a)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   159
  shows "OFCLASS('a, at_class)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   160
proof
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   161
  interpret type_definition Rep Abs "{a. sort_of a = s}" by (rule type)
1079
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   162
  have sort_of_Rep: "\<And>a. sort_of (Rep a) = s" using Rep by (simp add: image_def)
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   163
  fix a b :: 'a and p p1 p2 :: perm
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   164
  show "0 \<bullet> a = a"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   165
    unfolding permute_def by (simp add: Rep_inverse)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   166
  show "(p1 + p2) \<bullet> a = p1 \<bullet> p2 \<bullet> a"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   167
    unfolding permute_def by (simp add: Abs_inverse sort_of_Rep)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   168
  show "sort_of (atom a) = sort_of (atom b)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   169
    unfolding atom_def by (simp add: sort_of_Rep)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   170
  show "atom a = atom b \<longleftrightarrow> a = b"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   171
    unfolding atom_def by (simp add: Rep_inject)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   172
  show "p \<bullet> atom a = atom (p \<bullet> a)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   173
    unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   174
qed
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   175
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   176
setup {* Sign.add_const_constraint
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   177
  (@{const_name "permute"}, SOME @{typ "perm \<Rightarrow> 'a::pt \<Rightarrow> 'a"}) *}
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   178
setup {* Sign.add_const_constraint
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   179
  (@{const_name "atom"}, SOME @{typ "'a::at_base \<Rightarrow> atom"}) *}
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   180
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   181
section {* Automation for creating concrete atom types *}
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   182
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   183
text {* at the moment only single-sort concrete atoms are supported *}
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   184
1079
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   185
use "nominal_atoms.ML"
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   186
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   187
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   188
end