src/HOL/Nominal/Nominal.thy
author wenzelm
Tue, 11 Oct 2016 20:31:13 +0200
changeset 64149 1380bf90d986
parent 63167 0909deb8059b
child 66453 cc19f7ca2ed6
permissions -rw-r--r--
force fresh build;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19477
diff changeset
     1
theory Nominal 
58372
bfd497f2f4c2 moved 'old_datatype' out of 'Main' (but put it in 'HOL-Proofs' because of the inductive realizer)
blanchet
parents: 58318
diff changeset
     2
imports "~~/src/HOL/Library/Infinite_Set" "~~/src/HOL/Library/Old_Datatype"
46950
d0181abdbdac declare command keywords via theory header, including strict checking outside Pure;
wenzelm
parents: 46947
diff changeset
     3
keywords
d0181abdbdac declare command keywords via theory header, including strict checking outside Pure;
wenzelm
parents: 46947
diff changeset
     4
  "atom_decl" "nominal_datatype" "equivariance" :: thy_decl and
d0181abdbdac declare command keywords via theory header, including strict checking outside Pure;
wenzelm
parents: 46947
diff changeset
     5
  "nominal_primrec" "nominal_inductive" "nominal_inductive2" :: thy_goal and
d0181abdbdac declare command keywords via theory header, including strict checking outside Pure;
wenzelm
parents: 46947
diff changeset
     6
  "avoids"
44689
f247fc952f31 tuned specifications
haftmann
parents: 44683
diff changeset
     7
begin
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
     8
61260
e6f03fae14d5 explicit indication of overloaded typedefs;
wenzelm
parents: 60585
diff changeset
     9
declare [[typedef_overloaded]]
e6f03fae14d5 explicit indication of overloaded typedefs;
wenzelm
parents: 60585
diff changeset
    10
e6f03fae14d5 explicit indication of overloaded typedefs;
wenzelm
parents: 60585
diff changeset
    11
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
    12
section \<open>Permutations\<close>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    13
(*======================*)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    14
41798
c3aa3c87ef21 modernized specifications;
wenzelm
parents: 41562
diff changeset
    15
type_synonym 
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    16
  'x prm = "('x \<times> 'x) list"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    17
30990
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
    18
(* polymorphic constants for permutation and swapping *)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    19
consts 
18491
1ce410ff9941 Tuned syntax for perm.
berghofe
parents: 18431
diff changeset
    20
  perm :: "'x prm \<Rightarrow> 'a \<Rightarrow> 'a"     (infixr "\<bullet>" 80)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    21
  swap :: "('x \<times> 'x) \<Rightarrow> 'x \<Rightarrow> 'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    22
30983
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    23
(* a "private" copy of the option type used in the abstraction function *)
58310
91ea607a34d8 updated news
blanchet
parents: 58238
diff changeset
    24
datatype 'a noption = nSome 'a | nNone
58238
a701907d621e ported old Nominal to use new datatypes
blanchet
parents: 56230
diff changeset
    25
a701907d621e ported old Nominal to use new datatypes
blanchet
parents: 56230
diff changeset
    26
datatype_compat noption
30983
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    27
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    28
(* a "private" copy of the product type used in the nominal induct method *)
58310
91ea607a34d8 updated news
blanchet
parents: 58238
diff changeset
    29
datatype ('a, 'b) nprod = nPair 'a 'b
58238
a701907d621e ported old Nominal to use new datatypes
blanchet
parents: 56230
diff changeset
    30
a701907d621e ported old Nominal to use new datatypes
blanchet
parents: 56230
diff changeset
    31
datatype_compat nprod
30983
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    32
24544
da7de38392df trivial cleaning up
urbanc
parents: 23755
diff changeset
    33
(* an auxiliary constant for the decision procedure involving *) 
30983
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    34
(* permutations (to avoid loops when using perm-compositions)  *)
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
    35
definition
44683
daeb538c57bf tuned specifications and proofs
haftmann
parents: 44567
diff changeset
    36
  "perm_aux pi x = pi\<bullet>x"
19477
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
    37
30990
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
    38
(* overloaded permutation operations *)
30983
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    39
overloading
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    40
  perm_fun    \<equiv> "perm :: 'x prm \<Rightarrow> ('a\<Rightarrow>'b) \<Rightarrow> ('a\<Rightarrow>'b)"   (unchecked)
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    41
  perm_bool   \<equiv> "perm :: 'x prm \<Rightarrow> bool \<Rightarrow> bool"           (unchecked)
45961
5cefe17916a6 treatment of type constructor `set`
haftmann
parents: 45694
diff changeset
    42
  perm_set    \<equiv> "perm :: 'x prm \<Rightarrow> 'a set \<Rightarrow> 'a set"           (unchecked)
30983
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    43
  perm_unit   \<equiv> "perm :: 'x prm \<Rightarrow> unit \<Rightarrow> unit"           (unchecked)
44689
f247fc952f31 tuned specifications
haftmann
parents: 44683
diff changeset
    44
  perm_prod   \<equiv> "perm :: 'x prm \<Rightarrow> ('a\<times>'b) \<Rightarrow> ('a\<times>'b)"    (unchecked)
30983
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    45
  perm_list   \<equiv> "perm :: 'x prm \<Rightarrow> 'a list \<Rightarrow> 'a list"     (unchecked)
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    46
  perm_option \<equiv> "perm :: 'x prm \<Rightarrow> 'a option \<Rightarrow> 'a option" (unchecked)
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    47
  perm_char   \<equiv> "perm :: 'x prm \<Rightarrow> char \<Rightarrow> char"           (unchecked)
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    48
  perm_nat    \<equiv> "perm :: 'x prm \<Rightarrow> nat \<Rightarrow> nat"             (unchecked)
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    49
  perm_int    \<equiv> "perm :: 'x prm \<Rightarrow> int \<Rightarrow> int"             (unchecked)
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    50
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    51
  perm_noption \<equiv> "perm :: 'x prm \<Rightarrow> 'a noption \<Rightarrow> 'a noption"   (unchecked)
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    52
  perm_nprod   \<equiv> "perm :: 'x prm \<Rightarrow> ('a, 'b) nprod \<Rightarrow> ('a, 'b) nprod" (unchecked)
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    53
begin
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    54
44838
haftmann
parents: 44835
diff changeset
    55
definition perm_fun :: "'x prm \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b" where
44833
haftmann
parents: 44696
diff changeset
    56
  "perm_fun pi f = (\<lambda>x. pi \<bullet> f (rev pi \<bullet> x))"
44683
daeb538c57bf tuned specifications and proofs
haftmann
parents: 44567
diff changeset
    57
44689
f247fc952f31 tuned specifications
haftmann
parents: 44683
diff changeset
    58
definition perm_bool :: "'x prm \<Rightarrow> bool \<Rightarrow> bool" where
44833
haftmann
parents: 44696
diff changeset
    59
  "perm_bool pi b = b"
30983
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    60
45961
5cefe17916a6 treatment of type constructor `set`
haftmann
parents: 45694
diff changeset
    61
definition perm_set :: "'x prm \<Rightarrow> 'a set \<Rightarrow> 'a set" where
46179
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
    62
  "perm_set pi X = {pi \<bullet> x | x. x \<in> X}"
45961
5cefe17916a6 treatment of type constructor `set`
haftmann
parents: 45694
diff changeset
    63
44683
daeb538c57bf tuned specifications and proofs
haftmann
parents: 44567
diff changeset
    64
primrec perm_unit :: "'x prm \<Rightarrow> unit \<Rightarrow> unit"  where 
30983
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    65
  "perm_unit pi () = ()"
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    66
  
44683
daeb538c57bf tuned specifications and proofs
haftmann
parents: 44567
diff changeset
    67
primrec perm_prod :: "'x prm \<Rightarrow> ('a\<times>'b) \<Rightarrow> ('a\<times>'b)" where
44833
haftmann
parents: 44696
diff changeset
    68
  "perm_prod pi (x, y) = (pi\<bullet>x, pi\<bullet>y)"
30983
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    69
44683
daeb538c57bf tuned specifications and proofs
haftmann
parents: 44567
diff changeset
    70
primrec perm_list :: "'x prm \<Rightarrow> 'a list \<Rightarrow> 'a list" where
30983
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    71
  nil_eqvt:  "perm_list pi []     = []"
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    72
| cons_eqvt: "perm_list pi (x#xs) = (pi\<bullet>x)#(pi\<bullet>xs)"
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    73
44683
daeb538c57bf tuned specifications and proofs
haftmann
parents: 44567
diff changeset
    74
primrec perm_option :: "'x prm \<Rightarrow> 'a option \<Rightarrow> 'a option" where
30983
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    75
  some_eqvt:  "perm_option pi (Some x) = Some (pi\<bullet>x)"
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    76
| none_eqvt:  "perm_option pi None     = None"
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    77
44683
daeb538c57bf tuned specifications and proofs
haftmann
parents: 44567
diff changeset
    78
definition perm_char :: "'x prm \<Rightarrow> char \<Rightarrow> char" where
44833
haftmann
parents: 44696
diff changeset
    79
  "perm_char pi c = c"
44683
daeb538c57bf tuned specifications and proofs
haftmann
parents: 44567
diff changeset
    80
daeb538c57bf tuned specifications and proofs
haftmann
parents: 44567
diff changeset
    81
definition perm_nat :: "'x prm \<Rightarrow> nat \<Rightarrow> nat" where
44833
haftmann
parents: 44696
diff changeset
    82
  "perm_nat pi i = i"
44683
daeb538c57bf tuned specifications and proofs
haftmann
parents: 44567
diff changeset
    83
daeb538c57bf tuned specifications and proofs
haftmann
parents: 44567
diff changeset
    84
definition perm_int :: "'x prm \<Rightarrow> int \<Rightarrow> int" where
44833
haftmann
parents: 44696
diff changeset
    85
  "perm_int pi i = i"
44683
daeb538c57bf tuned specifications and proofs
haftmann
parents: 44567
diff changeset
    86
daeb538c57bf tuned specifications and proofs
haftmann
parents: 44567
diff changeset
    87
primrec perm_noption :: "'x prm \<Rightarrow> 'a noption \<Rightarrow> 'a noption" where
30983
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    88
  nsome_eqvt:  "perm_noption pi (nSome x) = nSome (pi\<bullet>x)"
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    89
| nnone_eqvt:  "perm_noption pi nNone     = nNone"
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    90
44683
daeb538c57bf tuned specifications and proofs
haftmann
parents: 44567
diff changeset
    91
primrec perm_nprod :: "'x prm \<Rightarrow> ('a, 'b) nprod \<Rightarrow> ('a, 'b) nprod" where
30983
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    92
  "perm_nprod pi (nPair x y) = nPair (pi\<bullet>x) (pi\<bullet>y)"
44683
daeb538c57bf tuned specifications and proofs
haftmann
parents: 44567
diff changeset
    93
30983
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    94
end
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    95
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
    96
(* permutations on booleans *)
44689
f247fc952f31 tuned specifications
haftmann
parents: 44683
diff changeset
    97
lemmas perm_bool = perm_bool_def
f247fc952f31 tuned specifications
haftmann
parents: 44683
diff changeset
    98
f247fc952f31 tuned specifications
haftmann
parents: 44683
diff changeset
    99
lemma true_eqvt [simp]:
f247fc952f31 tuned specifications
haftmann
parents: 44683
diff changeset
   100
  "pi \<bullet> True \<longleftrightarrow> True"
f247fc952f31 tuned specifications
haftmann
parents: 44683
diff changeset
   101
  by (simp add: perm_bool_def)
f247fc952f31 tuned specifications
haftmann
parents: 44683
diff changeset
   102
f247fc952f31 tuned specifications
haftmann
parents: 44683
diff changeset
   103
lemma false_eqvt [simp]:
f247fc952f31 tuned specifications
haftmann
parents: 44683
diff changeset
   104
  "pi \<bullet> False \<longleftrightarrow> False"
f247fc952f31 tuned specifications
haftmann
parents: 44683
diff changeset
   105
  by (simp add: perm_bool_def)
18264
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
   106
19972
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
   107
lemma perm_boolI:
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
   108
  assumes a: "P"
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
   109
  shows "pi\<bullet>P"
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
   110
  using a by (simp add: perm_bool)
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
   111
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
   112
lemma perm_boolE:
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
   113
  assumes a: "pi\<bullet>P"
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
   114
  shows "P"
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
   115
  using a by (simp add: perm_bool)
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
   116
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
   117
lemma if_eqvt:
21010
7fe928722821 added the missing freshness-lemmas for nat, int, char and string and
urbanc
parents: 20809
diff changeset
   118
  fixes pi::"'a prm"
7fe928722821 added the missing freshness-lemmas for nat, int, char and string and
urbanc
parents: 20809
diff changeset
   119
  shows "pi\<bullet>(if b then c1 else c2) = (if (pi\<bullet>b) then (pi\<bullet>c1) else (pi\<bullet>c2))"
30983
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
   120
  by (simp add: perm_fun_def)
21010
7fe928722821 added the missing freshness-lemmas for nat, int, char and string and
urbanc
parents: 20809
diff changeset
   121
22514
4dfa84906915 moving lemmas into appropriate sections
urbanc
parents: 22511
diff changeset
   122
lemma imp_eqvt:
4dfa84906915 moving lemmas into appropriate sections
urbanc
parents: 22511
diff changeset
   123
  shows "pi\<bullet>(A\<longrightarrow>B) = ((pi\<bullet>A)\<longrightarrow>(pi\<bullet>B))"
4dfa84906915 moving lemmas into appropriate sections
urbanc
parents: 22511
diff changeset
   124
  by (simp add: perm_bool)
4dfa84906915 moving lemmas into appropriate sections
urbanc
parents: 22511
diff changeset
   125
4dfa84906915 moving lemmas into appropriate sections
urbanc
parents: 22511
diff changeset
   126
lemma conj_eqvt:
4dfa84906915 moving lemmas into appropriate sections
urbanc
parents: 22511
diff changeset
   127
  shows "pi\<bullet>(A\<and>B) = ((pi\<bullet>A)\<and>(pi\<bullet>B))"
4dfa84906915 moving lemmas into appropriate sections
urbanc
parents: 22511
diff changeset
   128
  by (simp add: perm_bool)
4dfa84906915 moving lemmas into appropriate sections
urbanc
parents: 22511
diff changeset
   129
4dfa84906915 moving lemmas into appropriate sections
urbanc
parents: 22511
diff changeset
   130
lemma disj_eqvt:
4dfa84906915 moving lemmas into appropriate sections
urbanc
parents: 22511
diff changeset
   131
  shows "pi\<bullet>(A\<or>B) = ((pi\<bullet>A)\<or>(pi\<bullet>B))"
4dfa84906915 moving lemmas into appropriate sections
urbanc
parents: 22511
diff changeset
   132
  by (simp add: perm_bool)
4dfa84906915 moving lemmas into appropriate sections
urbanc
parents: 22511
diff changeset
   133
4dfa84906915 moving lemmas into appropriate sections
urbanc
parents: 22511
diff changeset
   134
lemma neg_eqvt:
4dfa84906915 moving lemmas into appropriate sections
urbanc
parents: 22511
diff changeset
   135
  shows "pi\<bullet>(\<not> A) = (\<not> (pi\<bullet>A))"
4dfa84906915 moving lemmas into appropriate sections
urbanc
parents: 22511
diff changeset
   136
  by (simp add: perm_bool)
4dfa84906915 moving lemmas into appropriate sections
urbanc
parents: 22511
diff changeset
   137
26806
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
   138
(* permutation on sets *)
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
   139
lemma empty_eqvt:
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
   140
  shows "pi\<bullet>{} = {}"
46179
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
   141
  by (simp add: perm_set_def)
26806
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
   142
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
   143
lemma union_eqvt:
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
   144
  shows "(pi\<bullet>(X\<union>Y)) = (pi\<bullet>X) \<union> (pi\<bullet>Y)"
46179
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
   145
  by (auto simp add: perm_set_def)
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
   146
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
   147
lemma insert_eqvt:
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
   148
  shows "pi\<bullet>(insert x X) = insert (pi\<bullet>x) (pi\<bullet>X)"
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
   149
  by (auto simp add: perm_set_def)
26806
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
   150
30983
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
   151
(* permutations on products *)
26806
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
   152
lemma fst_eqvt:
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
   153
  "pi\<bullet>(fst x) = fst (pi\<bullet>x)"
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
   154
 by (cases x) simp
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
   155
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
   156
lemma snd_eqvt:
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
   157
  "pi\<bullet>(snd x) = snd (pi\<bullet>x)"
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
   158
 by (cases x) simp
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
   159
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
   160
(* permutation on lists *)
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
   161
lemma append_eqvt:
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
   162
  fixes pi :: "'x prm"
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
   163
  and   l1 :: "'a list"
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
   164
  and   l2 :: "'a list"
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
   165
  shows "pi\<bullet>(l1@l2) = (pi\<bullet>l1)@(pi\<bullet>l2)"
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
   166
  by (induct l1) auto
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
   167
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
   168
lemma rev_eqvt:
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
   169
  fixes pi :: "'x prm"
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
   170
  and   l  :: "'a list"
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
   171
  shows "pi\<bullet>(rev l) = rev (pi\<bullet>l)"
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
   172
  by (induct l) (simp_all add: append_eqvt)
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
   173
46179
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
   174
lemma set_eqvt:
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
   175
  fixes pi :: "'x prm"
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
   176
  and   xs :: "'a list"
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
   177
  shows "pi\<bullet>(set xs) = set (pi\<bullet>xs)"
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
   178
by (induct xs) (auto simp add: empty_eqvt insert_eqvt)
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
   179
30983
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
   180
(* permutation on characters and strings *)
23050
722f58379538 added lemma for permutations on strings
urbanc
parents: 22846
diff changeset
   181
lemma perm_string:
722f58379538 added lemma for permutations on strings
urbanc
parents: 22846
diff changeset
   182
  fixes s::"string"
722f58379538 added lemma for permutations on strings
urbanc
parents: 22846
diff changeset
   183
  shows "pi\<bullet>s = s"
30983
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
   184
  by (induct s)(auto simp add: perm_char_def)
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
   185
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   186
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
   187
section \<open>permutation equality\<close>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   188
(*==============================*)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   189
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
   190
definition prm_eq :: "'x prm \<Rightarrow> 'x prm \<Rightarrow> bool" (" _ \<triangleq> _ " [80,80] 80) where
44683
daeb538c57bf tuned specifications and proofs
haftmann
parents: 44567
diff changeset
   191
  "pi1 \<triangleq> pi2 \<longleftrightarrow> (\<forall>a::'x. pi1\<bullet>a = pi2\<bullet>a)"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   192
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
   193
section \<open>Support, Freshness and Supports\<close>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   194
(*========================================*)
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
   195
definition supp :: "'a \<Rightarrow> ('x set)" where  
44683
daeb538c57bf tuned specifications and proofs
haftmann
parents: 44567
diff changeset
   196
   "supp x = {a . (infinite {b . [(a,b)]\<bullet>x \<noteq> x})}"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   197
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
   198
definition fresh :: "'x \<Rightarrow> 'a \<Rightarrow> bool" ("_ \<sharp> _" [80,80] 80) where
44683
daeb538c57bf tuned specifications and proofs
haftmann
parents: 44567
diff changeset
   199
   "a \<sharp> x \<longleftrightarrow> a \<notin> supp x"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   200
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
   201
definition supports :: "'x set \<Rightarrow> 'a \<Rightarrow> bool" (infixl "supports" 80) where
44683
daeb538c57bf tuned specifications and proofs
haftmann
parents: 44567
diff changeset
   202
   "S supports x \<longleftrightarrow> (\<forall>a b. (a\<notin>S \<and> b\<notin>S \<longrightarrow> [(a,b)]\<bullet>x=x))"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   203
30990
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
   204
(* lemmas about supp *)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   205
lemma supp_fresh_iff: 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   206
  fixes x :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   207
  shows "(supp x) = {a::'x. \<not>a\<sharp>x}"
30990
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
   208
  by (simp add: fresh_def)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
   209
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   210
lemma supp_unit:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   211
  shows "supp () = {}"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   212
  by (simp add: supp_def)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   213
18264
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
   214
lemma supp_set_empty:
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
   215
  shows "supp {} = {}"
26806
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
   216
  by (force simp add: supp_def empty_eqvt)
18264
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
   217
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   218
lemma supp_prod: 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   219
  fixes x :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   220
  and   y :: "'b"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   221
  shows "(supp (x,y)) = (supp x)\<union>(supp y)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   222
  by  (force simp add: supp_def Collect_imp_eq Collect_neg_eq)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   223
18600
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   224
lemma supp_nprod: 
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   225
  fixes x :: "'a"
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   226
  and   y :: "'b"
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   227
  shows "(supp (nPair x y)) = (supp x)\<union>(supp y)"
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   228
  by  (force simp add: supp_def Collect_imp_eq Collect_neg_eq)
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   229
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   230
lemma supp_list_nil:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   231
  shows "supp [] = {}"
44696
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
   232
  by (simp add: supp_def)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   233
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   234
lemma supp_list_cons:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   235
  fixes x  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   236
  and   xs :: "'a list"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   237
  shows "supp (x#xs) = (supp x)\<union>(supp xs)"
30990
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
   238
  by (auto simp add: supp_def Collect_imp_eq Collect_neg_eq)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   239
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   240
lemma supp_list_append:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   241
  fixes xs :: "'a list"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   242
  and   ys :: "'a list"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   243
  shows "supp (xs@ys) = (supp xs)\<union>(supp ys)"
30990
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
   244
  by (induct xs) (auto simp add: supp_list_nil supp_list_cons)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   245
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   246
lemma supp_list_rev:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   247
  fixes xs :: "'a list"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   248
  shows "supp (rev xs) = (supp xs)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   249
  by (induct xs, auto simp add: supp_list_append supp_list_cons supp_list_nil)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   250
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   251
lemma supp_bool:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   252
  fixes x  :: "bool"
30983
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
   253
  shows "supp x = {}"
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
   254
  by (cases "x") (simp_all add: supp_def)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   255
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   256
lemma supp_some:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   257
  fixes x :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   258
  shows "supp (Some x) = (supp x)"
30983
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
   259
  by (simp add: supp_def)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   260
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   261
lemma supp_none:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   262
  fixes x :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   263
  shows "supp (None) = {}"
30983
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
   264
  by (simp add: supp_def)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   265
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   266
lemma supp_int:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   267
  fixes i::"int"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   268
  shows "supp (i) = {}"
30983
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
   269
  by (simp add: supp_def perm_int_def)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   270
20388
b5a61270ea5a added missing supp_nat lemma
urbanc
parents: 19986
diff changeset
   271
lemma supp_nat:
b5a61270ea5a added missing supp_nat lemma
urbanc
parents: 19986
diff changeset
   272
  fixes n::"nat"
30983
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
   273
  shows "(supp n) = {}"
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
   274
  by (simp add: supp_def perm_nat_def)
20388
b5a61270ea5a added missing supp_nat lemma
urbanc
parents: 19986
diff changeset
   275
18627
f0acb66858b4 added the lemmas supp_char and supp_string
urbanc
parents: 18600
diff changeset
   276
lemma supp_char:
f0acb66858b4 added the lemmas supp_char and supp_string
urbanc
parents: 18600
diff changeset
   277
  fixes c::"char"
30983
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
   278
  shows "(supp c) = {}"
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
   279
  by (simp add: supp_def perm_char_def)
18627
f0acb66858b4 added the lemmas supp_char and supp_string
urbanc
parents: 18600
diff changeset
   280
  
f0acb66858b4 added the lemmas supp_char and supp_string
urbanc
parents: 18600
diff changeset
   281
lemma supp_string:
f0acb66858b4 added the lemmas supp_char and supp_string
urbanc
parents: 18600
diff changeset
   282
  fixes s::"string"
30983
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
   283
  shows "(supp s) = {}"
e54777ab68bd adapted permutation functions to new overloading syntax (the functions are still "unchecked" because they are used in conjunction with type-classes)
Christian Urban <urbanc@in.tum.de>
parents: 30242
diff changeset
   284
  by (simp add: supp_def perm_string)
18627
f0acb66858b4 added the lemmas supp_char and supp_string
urbanc
parents: 18600
diff changeset
   285
30990
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
   286
(* lemmas about freshness *)
18264
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
   287
lemma fresh_set_empty:
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
   288
  shows "a\<sharp>{}"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
   289
  by (simp add: fresh_def supp_set_empty)
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
   290
19858
d319e39a2e0e added lemma fresh_unit to Nominal.thy
urbanc
parents: 19856
diff changeset
   291
lemma fresh_unit:
d319e39a2e0e added lemma fresh_unit to Nominal.thy
urbanc
parents: 19856
diff changeset
   292
  shows "a\<sharp>()"
d319e39a2e0e added lemma fresh_unit to Nominal.thy
urbanc
parents: 19856
diff changeset
   293
  by (simp add: fresh_def supp_unit)
d319e39a2e0e added lemma fresh_unit to Nominal.thy
urbanc
parents: 19856
diff changeset
   294
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   295
lemma fresh_prod:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   296
  fixes a :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   297
  and   x :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   298
  and   y :: "'b"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   299
  shows "a\<sharp>(x,y) = (a\<sharp>x \<and> a\<sharp>y)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   300
  by (simp add: fresh_def supp_prod)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   301
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   302
lemma fresh_list_nil:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   303
  fixes a :: "'x"
18264
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
   304
  shows "a\<sharp>[]"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   305
  by (simp add: fresh_def supp_list_nil) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   306
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   307
lemma fresh_list_cons:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   308
  fixes a :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   309
  and   x :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   310
  and   xs :: "'a list"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   311
  shows "a\<sharp>(x#xs) = (a\<sharp>x \<and> a\<sharp>xs)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   312
  by (simp add: fresh_def supp_list_cons)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   313
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   314
lemma fresh_list_append:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   315
  fixes a :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   316
  and   xs :: "'a list"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   317
  and   ys :: "'a list"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   318
  shows "a\<sharp>(xs@ys) = (a\<sharp>xs \<and> a\<sharp>ys)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   319
  by (simp add: fresh_def supp_list_append)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   320
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   321
lemma fresh_list_rev:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   322
  fixes a :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   323
  and   xs :: "'a list"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   324
  shows "a\<sharp>(rev xs) = a\<sharp>xs"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   325
  by (simp add: fresh_def supp_list_rev)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   326
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   327
lemma fresh_none:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   328
  fixes a :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   329
  shows "a\<sharp>None"
22831
18f4014e1259 tuned some of the proofs and added the lemma fresh_bool
urbanc
parents: 22829
diff changeset
   330
  by (simp add: fresh_def supp_none)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   331
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   332
lemma fresh_some:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   333
  fixes a :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   334
  and   x :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   335
  shows "a\<sharp>(Some x) = a\<sharp>x"
22831
18f4014e1259 tuned some of the proofs and added the lemma fresh_bool
urbanc
parents: 22829
diff changeset
   336
  by (simp add: fresh_def supp_some)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   337
21010
7fe928722821 added the missing freshness-lemmas for nat, int, char and string and
urbanc
parents: 20809
diff changeset
   338
lemma fresh_int:
7fe928722821 added the missing freshness-lemmas for nat, int, char and string and
urbanc
parents: 20809
diff changeset
   339
  fixes a :: "'x"
7fe928722821 added the missing freshness-lemmas for nat, int, char and string and
urbanc
parents: 20809
diff changeset
   340
  and   i :: "int"
7fe928722821 added the missing freshness-lemmas for nat, int, char and string and
urbanc
parents: 20809
diff changeset
   341
  shows "a\<sharp>i"
22831
18f4014e1259 tuned some of the proofs and added the lemma fresh_bool
urbanc
parents: 22829
diff changeset
   342
  by (simp add: fresh_def supp_int)
21010
7fe928722821 added the missing freshness-lemmas for nat, int, char and string and
urbanc
parents: 20809
diff changeset
   343
7fe928722821 added the missing freshness-lemmas for nat, int, char and string and
urbanc
parents: 20809
diff changeset
   344
lemma fresh_nat:
7fe928722821 added the missing freshness-lemmas for nat, int, char and string and
urbanc
parents: 20809
diff changeset
   345
  fixes a :: "'x"
7fe928722821 added the missing freshness-lemmas for nat, int, char and string and
urbanc
parents: 20809
diff changeset
   346
  and   n :: "nat"
7fe928722821 added the missing freshness-lemmas for nat, int, char and string and
urbanc
parents: 20809
diff changeset
   347
  shows "a\<sharp>n"
22831
18f4014e1259 tuned some of the proofs and added the lemma fresh_bool
urbanc
parents: 22829
diff changeset
   348
  by (simp add: fresh_def supp_nat)
21010
7fe928722821 added the missing freshness-lemmas for nat, int, char and string and
urbanc
parents: 20809
diff changeset
   349
7fe928722821 added the missing freshness-lemmas for nat, int, char and string and
urbanc
parents: 20809
diff changeset
   350
lemma fresh_char:
7fe928722821 added the missing freshness-lemmas for nat, int, char and string and
urbanc
parents: 20809
diff changeset
   351
  fixes a :: "'x"
7fe928722821 added the missing freshness-lemmas for nat, int, char and string and
urbanc
parents: 20809
diff changeset
   352
  and   c :: "char"
7fe928722821 added the missing freshness-lemmas for nat, int, char and string and
urbanc
parents: 20809
diff changeset
   353
  shows "a\<sharp>c"
22831
18f4014e1259 tuned some of the proofs and added the lemma fresh_bool
urbanc
parents: 22829
diff changeset
   354
  by (simp add: fresh_def supp_char)
21010
7fe928722821 added the missing freshness-lemmas for nat, int, char and string and
urbanc
parents: 20809
diff changeset
   355
7fe928722821 added the missing freshness-lemmas for nat, int, char and string and
urbanc
parents: 20809
diff changeset
   356
lemma fresh_string:
7fe928722821 added the missing freshness-lemmas for nat, int, char and string and
urbanc
parents: 20809
diff changeset
   357
  fixes a :: "'x"
7fe928722821 added the missing freshness-lemmas for nat, int, char and string and
urbanc
parents: 20809
diff changeset
   358
  and   s :: "string"
7fe928722821 added the missing freshness-lemmas for nat, int, char and string and
urbanc
parents: 20809
diff changeset
   359
  shows "a\<sharp>s"
22831
18f4014e1259 tuned some of the proofs and added the lemma fresh_bool
urbanc
parents: 22829
diff changeset
   360
  by (simp add: fresh_def supp_string)
18f4014e1259 tuned some of the proofs and added the lemma fresh_bool
urbanc
parents: 22829
diff changeset
   361
18f4014e1259 tuned some of the proofs and added the lemma fresh_bool
urbanc
parents: 22829
diff changeset
   362
lemma fresh_bool:
18f4014e1259 tuned some of the proofs and added the lemma fresh_bool
urbanc
parents: 22829
diff changeset
   363
  fixes a :: "'x"
18f4014e1259 tuned some of the proofs and added the lemma fresh_bool
urbanc
parents: 22829
diff changeset
   364
  and   b :: "bool"
18f4014e1259 tuned some of the proofs and added the lemma fresh_bool
urbanc
parents: 22829
diff changeset
   365
  shows "a\<sharp>b"
18f4014e1259 tuned some of the proofs and added the lemma fresh_bool
urbanc
parents: 22829
diff changeset
   366
  by (simp add: fresh_def supp_bool)
21010
7fe928722821 added the missing freshness-lemmas for nat, int, char and string and
urbanc
parents: 20809
diff changeset
   367
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
   368
text \<open>Normalization of freshness results; cf.\ \<open>nominal_induct\<close>\<close>
21377
c29146dc14f1 replaced exists_fresh lemma with a version formulated with obtains;
urbanc
parents: 21318
diff changeset
   369
lemma fresh_unit_elim: 
c29146dc14f1 replaced exists_fresh lemma with a version formulated with obtains;
urbanc
parents: 21318
diff changeset
   370
  shows "(a\<sharp>() \<Longrightarrow> PROP C) \<equiv> PROP C"
18294
bbfd64cc91ab fresh_unit_elim and fresh_prod_elim -- for nominal_induct;
wenzelm
parents: 18268
diff changeset
   371
  by (simp add: fresh_def supp_unit)
bbfd64cc91ab fresh_unit_elim and fresh_prod_elim -- for nominal_induct;
wenzelm
parents: 18268
diff changeset
   372
21377
c29146dc14f1 replaced exists_fresh lemma with a version formulated with obtains;
urbanc
parents: 21318
diff changeset
   373
lemma fresh_prod_elim: 
c29146dc14f1 replaced exists_fresh lemma with a version formulated with obtains;
urbanc
parents: 21318
diff changeset
   374
  shows "(a\<sharp>(x,y) \<Longrightarrow> PROP C) \<equiv> (a\<sharp>x \<Longrightarrow> a\<sharp>y \<Longrightarrow> PROP C)"
18294
bbfd64cc91ab fresh_unit_elim and fresh_prod_elim -- for nominal_induct;
wenzelm
parents: 18268
diff changeset
   375
  by rule (simp_all add: fresh_prod)
bbfd64cc91ab fresh_unit_elim and fresh_prod_elim -- for nominal_induct;
wenzelm
parents: 18268
diff changeset
   376
21405
26b51f724fe6 added an intro lemma for freshness of products; set up
urbanc
parents: 21377
diff changeset
   377
(* this rule needs to be added before the fresh_prodD is *)
26b51f724fe6 added an intro lemma for freshness of products; set up
urbanc
parents: 21377
diff changeset
   378
(* added to the simplifier with mksimps                  *) 
26b51f724fe6 added an intro lemma for freshness of products; set up
urbanc
parents: 21377
diff changeset
   379
lemma [simp]:
26b51f724fe6 added an intro lemma for freshness of products; set up
urbanc
parents: 21377
diff changeset
   380
  shows "a\<sharp>x1 \<Longrightarrow> a\<sharp>x2 \<Longrightarrow> a\<sharp>(x1,x2)"
26b51f724fe6 added an intro lemma for freshness of products; set up
urbanc
parents: 21377
diff changeset
   381
  by (simp add: fresh_prod)
26b51f724fe6 added an intro lemma for freshness of products; set up
urbanc
parents: 21377
diff changeset
   382
21318
edb595802d22 added fresh_prodD, which is included fresh_prodD into mksimps setup;
wenzelm
parents: 21010
diff changeset
   383
lemma fresh_prodD:
21377
c29146dc14f1 replaced exists_fresh lemma with a version formulated with obtains;
urbanc
parents: 21318
diff changeset
   384
  shows "a\<sharp>(x,y) \<Longrightarrow> a\<sharp>x"
c29146dc14f1 replaced exists_fresh lemma with a version formulated with obtains;
urbanc
parents: 21318
diff changeset
   385
  and   "a\<sharp>(x,y) \<Longrightarrow> a\<sharp>y"
21318
edb595802d22 added fresh_prodD, which is included fresh_prodD into mksimps setup;
wenzelm
parents: 21010
diff changeset
   386
  by (simp_all add: fresh_prod)
edb595802d22 added fresh_prodD, which is included fresh_prodD into mksimps setup;
wenzelm
parents: 21010
diff changeset
   387
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
   388
ML \<open>
26342
0f65fa163304 more antiquotations;
wenzelm
parents: 26090
diff changeset
   389
  val mksimps_pairs = (@{const_name Nominal.fresh}, @{thms fresh_prodD}) :: mksimps_pairs;
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
   390
\<close>
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
   391
declaration \<open>fn _ =>
45625
750c5a47400b modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 44838
diff changeset
   392
  Simplifier.map_ss (Simplifier.set_mksimps (mksimps mksimps_pairs))
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
   393
\<close>
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
   394
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
   395
section \<open>Abstract Properties for Permutations and  Atoms\<close>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   396
(*=========================================================*)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   397
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   398
(* properties for being a permutation type *)
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
   399
definition
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   400
  "pt TYPE('a) TYPE('x) \<equiv> 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   401
     (\<forall>(x::'a). ([]::'x prm)\<bullet>x = x) \<and> 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   402
     (\<forall>(pi1::'x prm) (pi2::'x prm) (x::'a). (pi1@pi2)\<bullet>x = pi1\<bullet>(pi2\<bullet>x)) \<and> 
18295
dd50de393330 changed \<sim> of permutation equality to \<triangleq>
urbanc
parents: 18294
diff changeset
   403
     (\<forall>(pi1::'x prm) (pi2::'x prm) (x::'a). pi1 \<triangleq> pi2 \<longrightarrow> pi1\<bullet>x = pi2\<bullet>x)"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   404
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   405
(* properties for being an atom type *)
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
   406
definition
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   407
  "at TYPE('x) \<equiv> 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   408
     (\<forall>(x::'x). ([]::'x prm)\<bullet>x = x) \<and>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   409
     (\<forall>(a::'x) (b::'x) (pi::'x prm) (x::'x). ((a,b)#(pi::'x prm))\<bullet>x = swap (a,b) (pi\<bullet>x)) \<and> 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   410
     (\<forall>(a::'x) (b::'x) (c::'x). swap (a,b) c = (if a=c then b else (if b=c then a else c))) \<and> 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   411
     (infinite (UNIV::'x set))"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   412
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   413
(* property of two atom-types being disjoint *)
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
   414
definition
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   415
  "disjoint TYPE('x) TYPE('y) \<equiv> 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   416
       (\<forall>(pi::'x prm)(x::'y). pi\<bullet>x = x) \<and> 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   417
       (\<forall>(pi::'y prm)(x::'x). pi\<bullet>x = x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   418
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   419
(* composition property of two permutation on a type 'a *)
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
   420
definition
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   421
  "cp TYPE ('a) TYPE('x) TYPE('y) \<equiv> 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   422
      (\<forall>(pi2::'y prm) (pi1::'x prm) (x::'a) . pi1\<bullet>(pi2\<bullet>x) = (pi1\<bullet>pi2)\<bullet>(pi1\<bullet>x))" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   423
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   424
(* property of having finite support *)
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
   425
definition
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   426
  "fs TYPE('a) TYPE('x) \<equiv> \<forall>(x::'a). finite ((supp x)::'x set)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   427
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
   428
section \<open>Lemmas about the atom-type properties\<close>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   429
(*==============================================*)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   430
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   431
lemma at1: 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   432
  fixes x::"'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   433
  assumes a: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   434
  shows "([]::'x prm)\<bullet>x = x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   435
  using a by (simp add: at_def)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   436
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   437
lemma at2: 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   438
  fixes a ::"'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   439
  and   b ::"'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   440
  and   x ::"'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   441
  and   pi::"'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   442
  assumes a: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   443
  shows "((a,b)#pi)\<bullet>x = swap (a,b) (pi\<bullet>x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   444
  using a by (simp only: at_def)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   445
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   446
lemma at3: 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   447
  fixes a ::"'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   448
  and   b ::"'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   449
  and   c ::"'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   450
  assumes a: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   451
  shows "swap (a,b) c = (if a=c then b else (if b=c then a else c))"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   452
  using a by (simp only: at_def)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   453
30990
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
   454
(* rules to calculate simple permutations *)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   455
lemmas at_calc = at2 at1 at3
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   456
22610
c8b5133045f3 tuned slightly the previous commit
urbanc
parents: 22609
diff changeset
   457
lemma at_swap_simps:
c8b5133045f3 tuned slightly the previous commit
urbanc
parents: 22609
diff changeset
   458
  fixes a ::"'x"
c8b5133045f3 tuned slightly the previous commit
urbanc
parents: 22609
diff changeset
   459
  and   b ::"'x"
c8b5133045f3 tuned slightly the previous commit
urbanc
parents: 22609
diff changeset
   460
  assumes a: "at TYPE('x)"
c8b5133045f3 tuned slightly the previous commit
urbanc
parents: 22609
diff changeset
   461
  shows "[(a,b)]\<bullet>a = b"
c8b5133045f3 tuned slightly the previous commit
urbanc
parents: 22609
diff changeset
   462
  and   "[(a,b)]\<bullet>b = a"
27374
2a3c22fd95ab added a lemma to at_swap_simps
urbanc
parents: 27228
diff changeset
   463
  and   "\<lbrakk>a\<noteq>c; b\<noteq>c\<rbrakk> \<Longrightarrow> [(a,b)]\<bullet>c = c"
22610
c8b5133045f3 tuned slightly the previous commit
urbanc
parents: 22609
diff changeset
   464
  using a by (simp_all add: at_calc)
c8b5133045f3 tuned slightly the previous commit
urbanc
parents: 22609
diff changeset
   465
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   466
lemma at4: 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   467
  assumes a: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   468
  shows "infinite (UNIV::'x set)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   469
  using a by (simp add: at_def)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   470
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   471
lemma at_append:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   472
  fixes pi1 :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   473
  and   pi2 :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   474
  and   c   :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   475
  assumes at: "at TYPE('x)" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   476
  shows "(pi1@pi2)\<bullet>c = pi1\<bullet>(pi2\<bullet>c)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   477
proof (induct pi1)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   478
  case Nil show ?case by (simp add: at1[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   479
next
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   480
  case (Cons x xs)
18053
2719a6b7d95e some minor tweaks in some proofs (nothing extraordinary)
urbanc
parents: 18048
diff changeset
   481
  have "(xs@pi2)\<bullet>c  =  xs\<bullet>(pi2\<bullet>c)" by fact
2719a6b7d95e some minor tweaks in some proofs (nothing extraordinary)
urbanc
parents: 18048
diff changeset
   482
  also have "(x#xs)@pi2 = x#(xs@pi2)" by simp
2719a6b7d95e some minor tweaks in some proofs (nothing extraordinary)
urbanc
parents: 18048
diff changeset
   483
  ultimately show ?case by (cases "x", simp add:  at2[OF at])
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   484
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   485
 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   486
lemma at_swap:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   487
  fixes a :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   488
  and   b :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   489
  and   c :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   490
  assumes at: "at TYPE('x)" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   491
  shows "swap (a,b) (swap (a,b) c) = c"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   492
  by (auto simp add: at3[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   493
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   494
lemma at_rev_pi:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   495
  fixes pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   496
  and   c  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   497
  assumes at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   498
  shows "(rev pi)\<bullet>(pi\<bullet>c) = c"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   499
proof(induct pi)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   500
  case Nil show ?case by (simp add: at1[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   501
next
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   502
  case (Cons x xs) thus ?case 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   503
    by (cases "x", simp add: at2[OF at] at_append[OF at] at1[OF at] at_swap[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   504
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   505
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   506
lemma at_pi_rev:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   507
  fixes pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   508
  and   x  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   509
  assumes at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   510
  shows "pi\<bullet>((rev pi)\<bullet>x) = x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   511
  by (rule at_rev_pi[OF at, of "rev pi" _,simplified])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   512
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   513
lemma at_bij1: 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   514
  fixes pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   515
  and   x  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   516
  and   y  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   517
  assumes at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   518
  and     a:  "(pi\<bullet>x) = y"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   519
  shows   "x=(rev pi)\<bullet>y"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   520
proof -
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   521
  from a have "y=(pi\<bullet>x)" by (rule sym)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   522
  thus ?thesis by (simp only: at_rev_pi[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   523
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   524
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   525
lemma at_bij2: 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   526
  fixes pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   527
  and   x  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   528
  and   y  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   529
  assumes at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   530
  and     a:  "((rev pi)\<bullet>x) = y"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   531
  shows   "x=pi\<bullet>y"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   532
proof -
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   533
  from a have "y=((rev pi)\<bullet>x)" by (rule sym)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   534
  thus ?thesis by (simp only: at_pi_rev[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   535
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   536
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   537
lemma at_bij:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   538
  fixes pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   539
  and   x  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   540
  and   y  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   541
  assumes at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   542
  shows "(pi\<bullet>x = pi\<bullet>y) = (x=y)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   543
proof 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   544
  assume "pi\<bullet>x = pi\<bullet>y" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   545
  hence  "x=(rev pi)\<bullet>(pi\<bullet>y)" by (rule at_bij1[OF at]) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   546
  thus "x=y" by (simp only: at_rev_pi[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   547
next
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   548
  assume "x=y"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   549
  thus "pi\<bullet>x = pi\<bullet>y" by simp
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   550
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   551
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   552
lemma at_supp:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   553
  fixes x :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   554
  assumes at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   555
  shows "supp x = {x}"
29903
2c0046b26f80 more finiteness changes
nipkow
parents: 29128
diff changeset
   556
by(auto simp: supp_def Collect_conj_eq Collect_imp_eq at_calc[OF at] at4[OF at])
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   557
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   558
lemma at_fresh:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   559
  fixes a :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   560
  and   b :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   561
  assumes at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   562
  shows "(a\<sharp>b) = (a\<noteq>b)" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   563
  by (simp add: at_supp[OF at] fresh_def)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   564
26766
0e2a29a1065c polished the proof for atm_prm_fresh and more lemmas for fresh_star
urbanc
parents: 26522
diff changeset
   565
lemma at_prm_fresh1:
0e2a29a1065c polished the proof for atm_prm_fresh and more lemmas for fresh_star
urbanc
parents: 26522
diff changeset
   566
  fixes c :: "'x"
0e2a29a1065c polished the proof for atm_prm_fresh and more lemmas for fresh_star
urbanc
parents: 26522
diff changeset
   567
  and   pi:: "'x prm"
0e2a29a1065c polished the proof for atm_prm_fresh and more lemmas for fresh_star
urbanc
parents: 26522
diff changeset
   568
  assumes at: "at TYPE('x)"
0e2a29a1065c polished the proof for atm_prm_fresh and more lemmas for fresh_star
urbanc
parents: 26522
diff changeset
   569
  and     a: "c\<sharp>pi" 
0e2a29a1065c polished the proof for atm_prm_fresh and more lemmas for fresh_star
urbanc
parents: 26522
diff changeset
   570
  shows "\<forall>(a,b)\<in>set pi. c\<noteq>a \<and> c\<noteq>b"
0e2a29a1065c polished the proof for atm_prm_fresh and more lemmas for fresh_star
urbanc
parents: 26522
diff changeset
   571
using a by (induct pi) (auto simp add: fresh_list_cons fresh_prod at_fresh[OF at])
0e2a29a1065c polished the proof for atm_prm_fresh and more lemmas for fresh_star
urbanc
parents: 26522
diff changeset
   572
0e2a29a1065c polished the proof for atm_prm_fresh and more lemmas for fresh_star
urbanc
parents: 26522
diff changeset
   573
lemma at_prm_fresh2:
0e2a29a1065c polished the proof for atm_prm_fresh and more lemmas for fresh_star
urbanc
parents: 26522
diff changeset
   574
  fixes c :: "'x"
0e2a29a1065c polished the proof for atm_prm_fresh and more lemmas for fresh_star
urbanc
parents: 26522
diff changeset
   575
  and   pi:: "'x prm"
0e2a29a1065c polished the proof for atm_prm_fresh and more lemmas for fresh_star
urbanc
parents: 26522
diff changeset
   576
  assumes at: "at TYPE('x)"
0e2a29a1065c polished the proof for atm_prm_fresh and more lemmas for fresh_star
urbanc
parents: 26522
diff changeset
   577
  and     a: "\<forall>(a,b)\<in>set pi. c\<noteq>a \<and> c\<noteq>b" 
0e2a29a1065c polished the proof for atm_prm_fresh and more lemmas for fresh_star
urbanc
parents: 26522
diff changeset
   578
  shows "pi\<bullet>c = c"
0e2a29a1065c polished the proof for atm_prm_fresh and more lemmas for fresh_star
urbanc
parents: 26522
diff changeset
   579
using a  by(induct pi) (auto simp add: at1[OF at] at2[OF at] at3[OF at])
0e2a29a1065c polished the proof for atm_prm_fresh and more lemmas for fresh_star
urbanc
parents: 26522
diff changeset
   580
19107
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   581
lemma at_prm_fresh:
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   582
  fixes c :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   583
  and   pi:: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   584
  assumes at: "at TYPE('x)"
19107
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   585
  and     a: "c\<sharp>pi" 
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   586
  shows "pi\<bullet>c = c"
26766
0e2a29a1065c polished the proof for atm_prm_fresh and more lemmas for fresh_star
urbanc
parents: 26522
diff changeset
   587
by (rule at_prm_fresh2[OF at], rule at_prm_fresh1[OF at, OF a])
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   588
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   589
lemma at_prm_rev_eq:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   590
  fixes pi1 :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   591
  and   pi2 :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   592
  assumes at: "at TYPE('x)"
19107
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   593
  shows "((rev pi1) \<triangleq> (rev pi2)) = (pi1 \<triangleq> pi2)"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   594
proof (simp add: prm_eq_def, auto)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   595
  fix x
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   596
  assume "\<forall>x::'x. (rev pi1)\<bullet>x = (rev pi2)\<bullet>x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   597
  hence "(rev (pi1::'x prm))\<bullet>(pi2\<bullet>(x::'x)) = (rev (pi2::'x prm))\<bullet>(pi2\<bullet>x)" by simp
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   598
  hence "(rev (pi1::'x prm))\<bullet>((pi2::'x prm)\<bullet>x) = (x::'x)" by (simp add: at_rev_pi[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   599
  hence "(pi2::'x prm)\<bullet>x = (pi1::'x prm)\<bullet>x" by (simp add: at_bij2[OF at])
18295
dd50de393330 changed \<sim> of permutation equality to \<triangleq>
urbanc
parents: 18294
diff changeset
   600
  thus "pi1\<bullet>x  =  pi2\<bullet>x" by simp
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   601
next
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   602
  fix x
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   603
  assume "\<forall>x::'x. pi1\<bullet>x = pi2\<bullet>x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   604
  hence "(pi1::'x prm)\<bullet>((rev pi2)\<bullet>x) = (pi2::'x prm)\<bullet>((rev pi2)\<bullet>(x::'x))" by simp
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   605
  hence "(pi1::'x prm)\<bullet>((rev pi2)\<bullet>(x::'x)) = x" by (simp add: at_pi_rev[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   606
  hence "(rev pi2)\<bullet>x = (rev pi1)\<bullet>(x::'x)" by (simp add: at_bij1[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   607
  thus "(rev pi1)\<bullet>x = (rev pi2)\<bullet>(x::'x)" by simp
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   608
qed
19107
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   609
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   610
lemma at_prm_eq_append:
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   611
  fixes pi1 :: "'x prm"
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   612
  and   pi2 :: "'x prm"
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   613
  and   pi3 :: "'x prm"
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   614
  assumes at: "at TYPE('x)"
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   615
  and     a: "pi1 \<triangleq> pi2"
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   616
  shows "(pi3@pi1) \<triangleq> (pi3@pi2)"
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   617
using a by (simp add: prm_eq_def at_append[OF at] at_bij[OF at])
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   618
19325
35177b864f80 tuned some proofs
urbanc
parents: 19216
diff changeset
   619
lemma at_prm_eq_append':
35177b864f80 tuned some proofs
urbanc
parents: 19216
diff changeset
   620
  fixes pi1 :: "'x prm"
35177b864f80 tuned some proofs
urbanc
parents: 19216
diff changeset
   621
  and   pi2 :: "'x prm"
35177b864f80 tuned some proofs
urbanc
parents: 19216
diff changeset
   622
  and   pi3 :: "'x prm"
35177b864f80 tuned some proofs
urbanc
parents: 19216
diff changeset
   623
  assumes at: "at TYPE('x)"
35177b864f80 tuned some proofs
urbanc
parents: 19216
diff changeset
   624
  and     a: "pi1 \<triangleq> pi2"
35177b864f80 tuned some proofs
urbanc
parents: 19216
diff changeset
   625
  shows "(pi1@pi3) \<triangleq> (pi2@pi3)"
35177b864f80 tuned some proofs
urbanc
parents: 19216
diff changeset
   626
using a by (simp add: prm_eq_def at_append[OF at])
35177b864f80 tuned some proofs
urbanc
parents: 19216
diff changeset
   627
19107
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   628
lemma at_prm_eq_trans:
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   629
  fixes pi1 :: "'x prm"
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   630
  and   pi2 :: "'x prm"
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   631
  and   pi3 :: "'x prm"
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   632
  assumes a1: "pi1 \<triangleq> pi2"
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   633
  and     a2: "pi2 \<triangleq> pi3"  
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   634
  shows "pi1 \<triangleq> pi3"
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   635
using a1 a2 by (auto simp add: prm_eq_def)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   636
  
19107
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   637
lemma at_prm_eq_refl:
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   638
  fixes pi :: "'x prm"
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   639
  shows "pi \<triangleq> pi"
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   640
by (simp add: prm_eq_def)
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   641
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   642
lemma at_prm_rev_eq1:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   643
  fixes pi1 :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   644
  and   pi2 :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   645
  assumes at: "at TYPE('x)"
18295
dd50de393330 changed \<sim> of permutation equality to \<triangleq>
urbanc
parents: 18294
diff changeset
   646
  shows "pi1 \<triangleq> pi2 \<Longrightarrow> (rev pi1) \<triangleq> (rev pi2)"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   647
  by (simp add: at_prm_rev_eq[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   648
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   649
lemma at_ds1:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   650
  fixes a  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   651
  assumes at: "at TYPE('x)"
18295
dd50de393330 changed \<sim> of permutation equality to \<triangleq>
urbanc
parents: 18294
diff changeset
   652
  shows "[(a,a)] \<triangleq> []"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   653
  by (force simp add: prm_eq_def at_calc[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   654
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   655
lemma at_ds2: 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   656
  fixes pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   657
  and   a  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   658
  and   b  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   659
  assumes at: "at TYPE('x)"
19107
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   660
  shows "([(a,b)]@pi) \<triangleq> (pi@[((rev pi)\<bullet>a,(rev pi)\<bullet>b)])"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   661
  by (force simp add: prm_eq_def at_append[OF at] at_bij[OF at] at_pi_rev[OF at] 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   662
      at_rev_pi[OF at] at_calc[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   663
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   664
lemma at_ds3: 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   665
  fixes a  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   666
  and   b  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   667
  and   c  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   668
  assumes at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   669
  and     a:  "distinct [a,b,c]"
18295
dd50de393330 changed \<sim> of permutation equality to \<triangleq>
urbanc
parents: 18294
diff changeset
   670
  shows "[(a,c),(b,c),(a,c)] \<triangleq> [(a,b)]"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   671
  using a by (force simp add: prm_eq_def at_calc[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   672
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   673
lemma at_ds4: 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   674
  fixes a  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   675
  and   b  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   676
  and   pi  :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   677
  assumes at: "at TYPE('x)"
18295
dd50de393330 changed \<sim> of permutation equality to \<triangleq>
urbanc
parents: 18294
diff changeset
   678
  shows "(pi@[(a,(rev pi)\<bullet>b)]) \<triangleq> ([(pi\<bullet>a,b)]@pi)"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   679
  by (force simp add: prm_eq_def at_append[OF at] at_calc[OF at] at_bij[OF at] 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   680
      at_pi_rev[OF at] at_rev_pi[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   681
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   682
lemma at_ds5: 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   683
  fixes a  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   684
  and   b  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   685
  assumes at: "at TYPE('x)"
18295
dd50de393330 changed \<sim> of permutation equality to \<triangleq>
urbanc
parents: 18294
diff changeset
   686
  shows "[(a,b)] \<triangleq> [(b,a)]"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   687
  by (force simp add: prm_eq_def at_calc[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   688
19164
0eccb98b1fdb added initialisation-code for finite_guess
urbanc
parents: 19140
diff changeset
   689
lemma at_ds5': 
0eccb98b1fdb added initialisation-code for finite_guess
urbanc
parents: 19140
diff changeset
   690
  fixes a  :: "'x"
0eccb98b1fdb added initialisation-code for finite_guess
urbanc
parents: 19140
diff changeset
   691
  and   b  :: "'x"
0eccb98b1fdb added initialisation-code for finite_guess
urbanc
parents: 19140
diff changeset
   692
  assumes at: "at TYPE('x)"
0eccb98b1fdb added initialisation-code for finite_guess
urbanc
parents: 19140
diff changeset
   693
  shows "[(a,b),(b,a)] \<triangleq> []"
0eccb98b1fdb added initialisation-code for finite_guess
urbanc
parents: 19140
diff changeset
   694
  by (force simp add: prm_eq_def at_calc[OF at])
0eccb98b1fdb added initialisation-code for finite_guess
urbanc
parents: 19140
diff changeset
   695
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   696
lemma at_ds6: 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   697
  fixes a  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   698
  and   b  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   699
  and   c  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   700
  assumes at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   701
  and     a: "distinct [a,b,c]"
18295
dd50de393330 changed \<sim> of permutation equality to \<triangleq>
urbanc
parents: 18294
diff changeset
   702
  shows "[(a,c),(a,b)] \<triangleq> [(b,c),(a,c)]"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   703
  using a by (force simp add: prm_eq_def at_calc[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   704
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   705
lemma at_ds7:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   706
  fixes pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   707
  assumes at: "at TYPE('x)"
18295
dd50de393330 changed \<sim> of permutation equality to \<triangleq>
urbanc
parents: 18294
diff changeset
   708
  shows "((rev pi)@pi) \<triangleq> []"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   709
  by (simp add: prm_eq_def at1[OF at] at_append[OF at] at_rev_pi[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   710
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   711
lemma at_ds8_aux:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   712
  fixes pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   713
  and   a  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   714
  and   b  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   715
  and   c  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   716
  assumes at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   717
  shows "pi\<bullet>(swap (a,b) c) = swap (pi\<bullet>a,pi\<bullet>b) (pi\<bullet>c)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   718
  by (force simp add: at_calc[OF at] at_bij[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   719
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   720
lemma at_ds8: 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   721
  fixes pi1 :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   722
  and   pi2 :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   723
  and   a  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   724
  and   b  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   725
  assumes at: "at TYPE('x)"
18295
dd50de393330 changed \<sim> of permutation equality to \<triangleq>
urbanc
parents: 18294
diff changeset
   726
  shows "(pi1@pi2) \<triangleq> ((pi1\<bullet>pi2)@pi1)"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   727
apply(induct_tac pi2)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   728
apply(simp add: prm_eq_def)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   729
apply(auto simp add: prm_eq_def)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   730
apply(simp add: at2[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   731
apply(drule_tac x="aa" in spec)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   732
apply(drule sym)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   733
apply(simp)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   734
apply(simp add: at_append[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   735
apply(simp add: at2[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   736
apply(simp add: at_ds8_aux[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   737
done
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   738
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   739
lemma at_ds9: 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   740
  fixes pi1 :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   741
  and   pi2 :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   742
  and   a  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   743
  and   b  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   744
  assumes at: "at TYPE('x)"
18295
dd50de393330 changed \<sim> of permutation equality to \<triangleq>
urbanc
parents: 18294
diff changeset
   745
  shows " ((rev pi2)@(rev pi1)) \<triangleq> ((rev pi1)@(rev (pi1\<bullet>pi2)))"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   746
apply(induct_tac pi2)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   747
apply(simp add: prm_eq_def)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   748
apply(auto simp add: prm_eq_def)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   749
apply(simp add: at_append[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   750
apply(simp add: at2[OF at] at1[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   751
apply(drule_tac x="swap(pi1\<bullet>a,pi1\<bullet>b) aa" in spec)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   752
apply(drule sym)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   753
apply(simp)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   754
apply(simp add: at_ds8_aux[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   755
apply(simp add: at_rev_pi[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   756
done
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   757
19107
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   758
lemma at_ds10:
19132
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
   759
  fixes pi :: "'x prm"
19107
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   760
  and   a  :: "'x"
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   761
  and   b  :: "'x"
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   762
  assumes at: "at TYPE('x)"
19132
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
   763
  and     a:  "b\<sharp>(rev pi)"
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
   764
  shows "([(pi\<bullet>a,b)]@pi) \<triangleq> (pi@[(a,b)])"
19107
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   765
using a
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   766
apply -
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   767
apply(rule at_prm_eq_trans)
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   768
apply(rule at_ds2[OF at])
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   769
apply(simp add: at_prm_fresh[OF at] at_rev_pi[OF at])
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   770
apply(rule at_prm_eq_refl)
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   771
done
b16a45c53884 added a few lemmas to do with permutation-equivalence for the
urbanc
parents: 19045
diff changeset
   772
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
   773
\<comment>"there always exists an atom that is not being in a finite set"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   774
lemma ex_in_inf:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   775
  fixes   A::"'x set"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   776
  assumes at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   777
  and     fs: "finite A"
21377
c29146dc14f1 replaced exists_fresh lemma with a version formulated with obtains;
urbanc
parents: 21318
diff changeset
   778
  obtains c::"'x" where "c\<notin>A"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   779
proof -
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   780
  from  fs at4[OF at] have "infinite ((UNIV::'x set) - A)" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   781
    by (simp add: Diff_infinite_finite)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   782
  hence "((UNIV::'x set) - A) \<noteq> ({}::'x set)" by (force simp only:)
21377
c29146dc14f1 replaced exists_fresh lemma with a version formulated with obtains;
urbanc
parents: 21318
diff changeset
   783
  then obtain c::"'x" where "c\<in>((UNIV::'x set) - A)" by force
c29146dc14f1 replaced exists_fresh lemma with a version formulated with obtains;
urbanc
parents: 21318
diff changeset
   784
  then have "c\<notin>A" by simp
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 41413
diff changeset
   785
  then show ?thesis ..
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   786
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   787
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
   788
text \<open>there always exists a fresh name for an object with finite support\<close>
21377
c29146dc14f1 replaced exists_fresh lemma with a version formulated with obtains;
urbanc
parents: 21318
diff changeset
   789
lemma at_exists_fresh': 
c29146dc14f1 replaced exists_fresh lemma with a version formulated with obtains;
urbanc
parents: 21318
diff changeset
   790
  fixes  x :: "'a"
c29146dc14f1 replaced exists_fresh lemma with a version formulated with obtains;
urbanc
parents: 21318
diff changeset
   791
  assumes at: "at TYPE('x)"
c29146dc14f1 replaced exists_fresh lemma with a version formulated with obtains;
urbanc
parents: 21318
diff changeset
   792
  and     fs: "finite ((supp x)::'x set)"
c29146dc14f1 replaced exists_fresh lemma with a version formulated with obtains;
urbanc
parents: 21318
diff changeset
   793
  shows "\<exists>c::'x. c\<sharp>x"
c29146dc14f1 replaced exists_fresh lemma with a version formulated with obtains;
urbanc
parents: 21318
diff changeset
   794
  by (auto simp add: fresh_def intro: ex_in_inf[OF at, OF fs])
c29146dc14f1 replaced exists_fresh lemma with a version formulated with obtains;
urbanc
parents: 21318
diff changeset
   795
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   796
lemma at_exists_fresh: 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   797
  fixes  x :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   798
  assumes at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   799
  and     fs: "finite ((supp x)::'x set)"
21377
c29146dc14f1 replaced exists_fresh lemma with a version formulated with obtains;
urbanc
parents: 21318
diff changeset
   800
  obtains c::"'x" where  "c\<sharp>x"
c29146dc14f1 replaced exists_fresh lemma with a version formulated with obtains;
urbanc
parents: 21318
diff changeset
   801
  by (auto intro: ex_in_inf[OF at, OF fs] simp add: fresh_def)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   802
21377
c29146dc14f1 replaced exists_fresh lemma with a version formulated with obtains;
urbanc
parents: 21318
diff changeset
   803
lemma at_finite_select: 
30990
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
   804
  fixes S::"'a set"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
   805
  assumes a: "at TYPE('a)"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
   806
  and     b: "finite S" 
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
   807
  shows "\<exists>x. x \<notin> S" 
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
   808
  using a b
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
   809
  apply(drule_tac S="UNIV::'a set" in Diff_infinite_finite)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
   810
  apply(simp add: at_def)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
   811
  apply(subgoal_tac "UNIV - S \<noteq> {}")
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
   812
  apply(simp only: ex_in_conv [symmetric])
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
   813
  apply(blast)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
   814
  apply(rule notI)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
   815
  apply(simp)
18657
0a37df3bb99d Added theorem at_finite_select.
berghofe
parents: 18656
diff changeset
   816
  done
0a37df3bb99d Added theorem at_finite_select.
berghofe
parents: 18656
diff changeset
   817
19140
5a98cdf99079 replaced the lemma at_two by at_different;
urbanc
parents: 19132
diff changeset
   818
lemma at_different:
19132
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
   819
  assumes at: "at TYPE('x)"
19140
5a98cdf99079 replaced the lemma at_two by at_different;
urbanc
parents: 19132
diff changeset
   820
  shows "\<exists>(b::'x). a\<noteq>b"
19132
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
   821
proof -
19140
5a98cdf99079 replaced the lemma at_two by at_different;
urbanc
parents: 19132
diff changeset
   822
  have "infinite (UNIV::'x set)" by (rule at4[OF at])
5a98cdf99079 replaced the lemma at_two by at_different;
urbanc
parents: 19132
diff changeset
   823
  hence inf2: "infinite (UNIV-{a})" by (rule infinite_remove)
19132
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
   824
  have "(UNIV-{a}) \<noteq> ({}::'x set)" 
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
   825
  proof (rule_tac ccontr, drule_tac notnotD)
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
   826
    assume "UNIV-{a} = ({}::'x set)"
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
   827
    with inf2 have "infinite ({}::'x set)" by simp
19869
eba1b9e7c458 removal of the obsolete "infinite_nonempty"
paulson
parents: 19858
diff changeset
   828
    then show "False" by auto
19132
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
   829
  qed
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
   830
  hence "\<exists>(b::'x). b\<in>(UNIV-{a})" by blast
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
   831
  then obtain b::"'x" where mem2: "b\<in>(UNIV-{a})" by blast
19140
5a98cdf99079 replaced the lemma at_two by at_different;
urbanc
parents: 19132
diff changeset
   832
  from mem2 have "a\<noteq>b" by blast
5a98cdf99079 replaced the lemma at_two by at_different;
urbanc
parents: 19132
diff changeset
   833
  then show "\<exists>(b::'x). a\<noteq>b" by blast
19132
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
   834
qed
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
   835
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
   836
\<comment>"the at-props imply the pt-props"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   837
lemma at_pt_inst:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   838
  assumes at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   839
  shows "pt TYPE('x) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   840
apply(auto simp only: pt_def)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   841
apply(simp only: at1[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   842
apply(simp only: at_append[OF at]) 
18053
2719a6b7d95e some minor tweaks in some proofs (nothing extraordinary)
urbanc
parents: 18048
diff changeset
   843
apply(simp only: prm_eq_def)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   844
done
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   845
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
   846
section \<open>finite support properties\<close>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   847
(*===================================*)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   848
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   849
lemma fs1:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   850
  fixes x :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   851
  assumes a: "fs TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   852
  shows "finite ((supp x)::'x set)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   853
  using a by (simp add: fs_def)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   854
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   855
lemma fs_at_inst:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   856
  fixes a :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   857
  assumes at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   858
  shows "fs TYPE('x) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   859
apply(simp add: fs_def) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   860
apply(simp add: at_supp[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   861
done
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   862
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   863
lemma fs_unit_inst:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   864
  shows "fs TYPE(unit) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   865
apply(simp add: fs_def)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   866
apply(simp add: supp_unit)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   867
done
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   868
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   869
lemma fs_prod_inst:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   870
  assumes fsa: "fs TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   871
  and     fsb: "fs TYPE('b) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   872
  shows "fs TYPE('a\<times>'b) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   873
apply(unfold fs_def)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   874
apply(auto simp add: supp_prod)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   875
apply(rule fs1[OF fsa])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   876
apply(rule fs1[OF fsb])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   877
done
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   878
18600
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   879
lemma fs_nprod_inst:
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   880
  assumes fsa: "fs TYPE('a) TYPE('x)"
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   881
  and     fsb: "fs TYPE('b) TYPE('x)"
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   882
  shows "fs TYPE(('a,'b) nprod) TYPE('x)"
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   883
apply(unfold fs_def, rule allI)
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   884
apply(case_tac x)
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   885
apply(auto simp add: supp_nprod)
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   886
apply(rule fs1[OF fsa])
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   887
apply(rule fs1[OF fsb])
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   888
done
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   889
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   890
lemma fs_list_inst:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   891
  assumes fs: "fs TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   892
  shows "fs TYPE('a list) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   893
apply(simp add: fs_def, rule allI)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   894
apply(induct_tac x)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   895
apply(simp add: supp_list_nil)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   896
apply(simp add: supp_list_cons)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   897
apply(rule fs1[OF fs])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   898
done
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   899
18431
a59c79a3544c improved the finite-support proof
urbanc
parents: 18351
diff changeset
   900
lemma fs_option_inst:
a59c79a3544c improved the finite-support proof
urbanc
parents: 18351
diff changeset
   901
  assumes fs: "fs TYPE('a) TYPE('x)"
a59c79a3544c improved the finite-support proof
urbanc
parents: 18351
diff changeset
   902
  shows "fs TYPE('a option) TYPE('x)"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   903
apply(simp add: fs_def, rule allI)
18431
a59c79a3544c improved the finite-support proof
urbanc
parents: 18351
diff changeset
   904
apply(case_tac x)
a59c79a3544c improved the finite-support proof
urbanc
parents: 18351
diff changeset
   905
apply(simp add: supp_none)
a59c79a3544c improved the finite-support proof
urbanc
parents: 18351
diff changeset
   906
apply(simp add: supp_some)
a59c79a3544c improved the finite-support proof
urbanc
parents: 18351
diff changeset
   907
apply(rule fs1[OF fs])
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   908
done
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   909
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
   910
section \<open>Lemmas about the permutation properties\<close>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   911
(*=================================================*)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   912
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   913
lemma pt1:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   914
  fixes x::"'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   915
  assumes a: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   916
  shows "([]::'x prm)\<bullet>x = x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   917
  using a by (simp add: pt_def)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   918
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   919
lemma pt2: 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   920
  fixes pi1::"'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   921
  and   pi2::"'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   922
  and   x  ::"'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   923
  assumes a: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   924
  shows "(pi1@pi2)\<bullet>x = pi1\<bullet>(pi2\<bullet>x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   925
  using a by (simp add: pt_def)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   926
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   927
lemma pt3:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   928
  fixes pi1::"'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   929
  and   pi2::"'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   930
  and   x  ::"'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   931
  assumes a: "pt TYPE('a) TYPE('x)"
18295
dd50de393330 changed \<sim> of permutation equality to \<triangleq>
urbanc
parents: 18294
diff changeset
   932
  shows "pi1 \<triangleq> pi2 \<Longrightarrow> pi1\<bullet>x = pi2\<bullet>x"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   933
  using a by (simp add: pt_def)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   934
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   935
lemma pt3_rev:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   936
  fixes pi1::"'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   937
  and   pi2::"'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   938
  and   x  ::"'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   939
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   940
  and     at: "at TYPE('x)"
18295
dd50de393330 changed \<sim> of permutation equality to \<triangleq>
urbanc
parents: 18294
diff changeset
   941
  shows "pi1 \<triangleq> pi2 \<Longrightarrow> (rev pi1)\<bullet>x = (rev pi2)\<bullet>x"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   942
  by (rule pt3[OF pt], simp add: at_prm_rev_eq[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   943
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
   944
section \<open>composition properties\<close>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   945
(* ============================== *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   946
lemma cp1:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   947
  fixes pi1::"'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   948
  and   pi2::"'y prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   949
  and   x  ::"'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   950
  assumes cp: "cp TYPE ('a) TYPE('x) TYPE('y)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   951
  shows "pi1\<bullet>(pi2\<bullet>x) = (pi1\<bullet>pi2)\<bullet>(pi1\<bullet>x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   952
  using cp by (simp add: cp_def)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   953
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   954
lemma cp_pt_inst:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   955
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   956
  and     at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   957
  shows "cp TYPE('a) TYPE('x) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   958
apply(auto simp add: cp_def pt2[OF pt,symmetric])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   959
apply(rule pt3[OF pt])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   960
apply(rule at_ds8[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   961
done
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   962
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
   963
section \<open>disjointness properties\<close>
19638
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
   964
(*=================================*)
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
   965
lemma dj_perm_forget:
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
   966
  fixes pi::"'y prm"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
   967
  and   x ::"'x"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
   968
  assumes dj: "disjoint TYPE('x) TYPE('y)"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
   969
  shows "pi\<bullet>x=x" 
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
   970
  using dj by (simp_all add: disjoint_def)
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
   971
28371
471a356fdea9 Added some more lemmas that are useful in proof of strong induction rule.
berghofe
parents: 28322
diff changeset
   972
lemma dj_perm_set_forget:
471a356fdea9 Added some more lemmas that are useful in proof of strong induction rule.
berghofe
parents: 28322
diff changeset
   973
  fixes pi::"'y prm"
471a356fdea9 Added some more lemmas that are useful in proof of strong induction rule.
berghofe
parents: 28322
diff changeset
   974
  and   x ::"'x set"
471a356fdea9 Added some more lemmas that are useful in proof of strong induction rule.
berghofe
parents: 28322
diff changeset
   975
  assumes dj: "disjoint TYPE('x) TYPE('y)"
44833
haftmann
parents: 44696
diff changeset
   976
  shows "pi\<bullet>x=x" 
45961
5cefe17916a6 treatment of type constructor `set`
haftmann
parents: 45694
diff changeset
   977
  using dj by (simp_all add: perm_set_def disjoint_def)
28371
471a356fdea9 Added some more lemmas that are useful in proof of strong induction rule.
berghofe
parents: 28322
diff changeset
   978
19638
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
   979
lemma dj_perm_perm_forget:
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
   980
  fixes pi1::"'x prm"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
   981
  and   pi2::"'y prm"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
   982
  assumes dj: "disjoint TYPE('x) TYPE('y)"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
   983
  shows "pi2\<bullet>pi1=pi1"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
   984
  using dj by (induct pi1, auto simp add: disjoint_def)
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
   985
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
   986
lemma dj_cp:
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
   987
  fixes pi1::"'x prm"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
   988
  and   pi2::"'y prm"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
   989
  and   x  ::"'a"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
   990
  assumes cp: "cp TYPE ('a) TYPE('x) TYPE('y)"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
   991
  and     dj: "disjoint TYPE('y) TYPE('x)"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
   992
  shows "pi1\<bullet>(pi2\<bullet>x) = (pi2)\<bullet>(pi1\<bullet>x)"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
   993
  by (simp add: cp1[OF cp] dj_perm_perm_forget[OF dj])
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
   994
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
   995
lemma dj_supp:
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
   996
  fixes a::"'x"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
   997
  assumes dj: "disjoint TYPE('x) TYPE('y)"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
   998
  shows "(supp a) = ({}::'y set)"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
   999
apply(simp add: supp_def dj_perm_forget[OF dj])
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
  1000
done
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
  1001
19972
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
  1002
lemma at_fresh_ineq:
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
  1003
  fixes a :: "'x"
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
  1004
  and   b :: "'y"
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
  1005
  assumes dj: "disjoint TYPE('y) TYPE('x)"
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
  1006
  shows "a\<sharp>b" 
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
  1007
  by (simp add: fresh_def dj_supp[OF dj])
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
  1008
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  1009
section \<open>permutation type instances\<close>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1010
(* ===================================*)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1011
44696
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1012
lemma pt_fun_inst:
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1013
  assumes pta: "pt TYPE('a) TYPE('x)"
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1014
  and     ptb: "pt TYPE('b) TYPE('x)"
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1015
  and     at:  "at TYPE('x)"
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1016
  shows  "pt TYPE('a\<Rightarrow>'b) TYPE('x)"
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1017
apply(auto simp only: pt_def)
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1018
apply(simp_all add: perm_fun_def)
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1019
apply(simp add: pt1[OF pta] pt1[OF ptb])
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1020
apply(simp add: pt2[OF pta] pt2[OF ptb])
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1021
apply(subgoal_tac "(rev pi1) \<triangleq> (rev pi2)")(*A*)
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1022
apply(simp add: pt3[OF pta] pt3[OF ptb])
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1023
(*A*)
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1024
apply(simp add: at_prm_rev_eq[OF at])
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1025
done
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1026
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1027
lemma pt_bool_inst:
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1028
  shows  "pt TYPE(bool) TYPE('x)"
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1029
  by (simp add: pt_def perm_bool_def)
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1030
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1031
lemma pt_set_inst:
46179
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  1032
  assumes pt: "pt TYPE('a) TYPE('x)"
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  1033
  shows  "pt TYPE('a set) TYPE('x)"
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  1034
apply(simp add: pt_def)
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  1035
apply(simp_all add: perm_set_def)
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  1036
apply(simp add: pt1[OF pt])
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  1037
apply(force simp add: pt2[OF pt] pt3[OF pt])
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  1038
done
44696
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1039
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1040
lemma pt_unit_inst:
44833
haftmann
parents: 44696
diff changeset
  1041
  shows "pt TYPE(unit) TYPE('x)"
44696
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1042
  by (simp add: pt_def)
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1043
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1044
lemma pt_prod_inst:
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1045
  assumes pta: "pt TYPE('a) TYPE('x)"
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1046
  and     ptb: "pt TYPE('b) TYPE('x)"
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1047
  shows  "pt TYPE('a \<times> 'b) TYPE('x)"
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1048
  apply(auto simp add: pt_def)
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1049
  apply(rule pt1[OF pta])
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1050
  apply(rule pt1[OF ptb])
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1051
  apply(rule pt2[OF pta])
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1052
  apply(rule pt2[OF ptb])
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1053
  apply(rule pt3[OF pta],assumption)
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1054
  apply(rule pt3[OF ptb],assumption)
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1055
  done
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1056
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1057
lemma pt_list_nil: 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1058
  fixes xs :: "'a list"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1059
  assumes pt: "pt TYPE('a) TYPE ('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1060
  shows "([]::'x prm)\<bullet>xs = xs" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1061
apply(induct_tac xs)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1062
apply(simp_all add: pt1[OF pt])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1063
done
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1064
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1065
lemma pt_list_append: 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1066
  fixes pi1 :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1067
  and   pi2 :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1068
  and   xs  :: "'a list"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1069
  assumes pt: "pt TYPE('a) TYPE ('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1070
  shows "(pi1@pi2)\<bullet>xs = pi1\<bullet>(pi2\<bullet>xs)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1071
apply(induct_tac xs)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1072
apply(simp_all add: pt2[OF pt])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1073
done
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1074
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1075
lemma pt_list_prm_eq: 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1076
  fixes pi1 :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1077
  and   pi2 :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1078
  and   xs  :: "'a list"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1079
  assumes pt: "pt TYPE('a) TYPE ('x)"
18295
dd50de393330 changed \<sim> of permutation equality to \<triangleq>
urbanc
parents: 18294
diff changeset
  1080
  shows "pi1 \<triangleq> pi2  \<Longrightarrow> pi1\<bullet>xs = pi2\<bullet>xs"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1081
apply(induct_tac xs)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1082
apply(simp_all add: prm_eq_def pt3[OF pt])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1083
done
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1084
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1085
lemma pt_list_inst:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1086
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1087
  shows  "pt TYPE('a list) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1088
apply(auto simp only: pt_def)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1089
apply(rule pt_list_nil[OF pt])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1090
apply(rule pt_list_append[OF pt])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1091
apply(rule pt_list_prm_eq[OF pt],assumption)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1092
done
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1093
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1094
lemma pt_option_inst:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1095
  assumes pta: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1096
  shows  "pt TYPE('a option) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1097
apply(auto simp only: pt_def)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1098
apply(case_tac "x")
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1099
apply(simp_all add: pt1[OF pta])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1100
apply(case_tac "x")
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1101
apply(simp_all add: pt2[OF pta])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1102
apply(case_tac "x")
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1103
apply(simp_all add: pt3[OF pta])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1104
done
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1105
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1106
lemma pt_noption_inst:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1107
  assumes pta: "pt TYPE('a) TYPE('x)"
18579
002d371401f5 changed the name of the type "nOption" to "noption".
urbanc
parents: 18578
diff changeset
  1108
  shows  "pt TYPE('a noption) TYPE('x)"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1109
apply(auto simp only: pt_def)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1110
apply(case_tac "x")
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1111
apply(simp_all add: pt1[OF pta])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1112
apply(case_tac "x")
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1113
apply(simp_all add: pt2[OF pta])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1114
apply(case_tac "x")
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1115
apply(simp_all add: pt3[OF pta])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1116
done
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1117
44696
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1118
lemma pt_nprod_inst:
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1119
  assumes pta: "pt TYPE('a) TYPE('x)"
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1120
  and     ptb: "pt TYPE('b) TYPE('x)"
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1121
  shows  "pt TYPE(('a,'b) nprod) TYPE('x)"
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1122
  apply(auto simp add: pt_def)
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1123
  apply(case_tac x)
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1124
  apply(simp add: pt1[OF pta] pt1[OF ptb])
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1125
  apply(case_tac x)
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1126
  apply(simp add: pt2[OF pta] pt2[OF ptb])
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1127
  apply(case_tac x)
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1128
  apply(simp add: pt3[OF pta] pt3[OF ptb])
4e99277c81f2 pseudo-definition for perms on sets; tuned
haftmann
parents: 44689
diff changeset
  1129
  done
24544
da7de38392df trivial cleaning up
urbanc
parents: 23755
diff changeset
  1130
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  1131
section \<open>further lemmas for permutation types\<close>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1132
(*==============================================*)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1133
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1134
lemma pt_rev_pi:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1135
  fixes pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1136
  and   x  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1137
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1138
  and     at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1139
  shows "(rev pi)\<bullet>(pi\<bullet>x) = x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1140
proof -
18295
dd50de393330 changed \<sim> of permutation equality to \<triangleq>
urbanc
parents: 18294
diff changeset
  1141
  have "((rev pi)@pi) \<triangleq> ([]::'x prm)" by (simp add: at_ds7[OF at])
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1142
  hence "((rev pi)@pi)\<bullet>(x::'a) = ([]::'x prm)\<bullet>x" by (simp add: pt3[OF pt]) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1143
  thus ?thesis by (simp add: pt1[OF pt] pt2[OF pt])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1144
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1145
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1146
lemma pt_pi_rev:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1147
  fixes pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1148
  and   x  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1149
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1150
  and     at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1151
  shows "pi\<bullet>((rev pi)\<bullet>x) = x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1152
  by (simp add: pt_rev_pi[OF pt, OF at,of "rev pi" "x",simplified])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1153
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1154
lemma pt_bij1: 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1155
  fixes pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1156
  and   x  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1157
  and   y  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1158
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1159
  and     at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1160
  and     a:  "(pi\<bullet>x) = y"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1161
  shows   "x=(rev pi)\<bullet>y"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1162
proof -
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1163
  from a have "y=(pi\<bullet>x)" by (rule sym)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1164
  thus ?thesis by (simp only: pt_rev_pi[OF pt, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1165
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1166
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1167
lemma pt_bij2: 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1168
  fixes pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1169
  and   x  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1170
  and   y  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1171
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1172
  and     at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1173
  and     a:  "x = (rev pi)\<bullet>y"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1174
  shows   "(pi\<bullet>x)=y"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1175
  using a by (simp add: pt_pi_rev[OF pt, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1176
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1177
lemma pt_bij:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1178
  fixes pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1179
  and   x  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1180
  and   y  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1181
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1182
  and     at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1183
  shows "(pi\<bullet>x = pi\<bullet>y) = (x=y)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1184
proof 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1185
  assume "pi\<bullet>x = pi\<bullet>y" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1186
  hence  "x=(rev pi)\<bullet>(pi\<bullet>y)" by (rule pt_bij1[OF pt, OF at]) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1187
  thus "x=y" by (simp only: pt_rev_pi[OF pt, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1188
next
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1189
  assume "x=y"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1190
  thus "pi\<bullet>x = pi\<bullet>y" by simp
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1191
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1192
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  1193
lemma pt_eq_eqvt:
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  1194
  fixes pi :: "'x prm"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  1195
  and   x  :: "'a"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  1196
  and   y  :: "'a"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  1197
  assumes pt: "pt TYPE('a) TYPE('x)"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  1198
  and     at: "at TYPE('x)"
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22808
diff changeset
  1199
  shows "pi\<bullet>(x=y) = (pi\<bullet>x = pi\<bullet>y)"
30990
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  1200
  using pt at
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  1201
  by (auto simp add: pt_bij perm_bool)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  1202
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1203
lemma pt_bij3:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1204
  fixes pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1205
  and   x  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1206
  and   y  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1207
  assumes a:  "x=y"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1208
  shows "(pi\<bullet>x = pi\<bullet>y)"
30990
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  1209
  using a by simp 
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1210
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1211
lemma pt_bij4:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1212
  fixes pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1213
  and   x  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1214
  and   y  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1215
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1216
  and     at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1217
  and     a:  "pi\<bullet>x = pi\<bullet>y"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1218
  shows "x = y"
30990
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  1219
  using a by (simp add: pt_bij[OF pt, OF at])
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1220
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1221
lemma pt_swap_bij:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1222
  fixes a  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1223
  and   b  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1224
  and   x  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1225
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1226
  and     at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1227
  shows "[(a,b)]\<bullet>([(a,b)]\<bullet>x) = x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1228
  by (rule pt_bij2[OF pt, OF at], simp)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1229
19164
0eccb98b1fdb added initialisation-code for finite_guess
urbanc
parents: 19140
diff changeset
  1230
lemma pt_swap_bij':
0eccb98b1fdb added initialisation-code for finite_guess
urbanc
parents: 19140
diff changeset
  1231
  fixes a  :: "'x"
0eccb98b1fdb added initialisation-code for finite_guess
urbanc
parents: 19140
diff changeset
  1232
  and   b  :: "'x"
0eccb98b1fdb added initialisation-code for finite_guess
urbanc
parents: 19140
diff changeset
  1233
  and   x  :: "'a"
0eccb98b1fdb added initialisation-code for finite_guess
urbanc
parents: 19140
diff changeset
  1234
  assumes pt: "pt TYPE('a) TYPE('x)"
0eccb98b1fdb added initialisation-code for finite_guess
urbanc
parents: 19140
diff changeset
  1235
  and     at: "at TYPE('x)"
0eccb98b1fdb added initialisation-code for finite_guess
urbanc
parents: 19140
diff changeset
  1236
  shows "[(a,b)]\<bullet>([(b,a)]\<bullet>x) = x"
0eccb98b1fdb added initialisation-code for finite_guess
urbanc
parents: 19140
diff changeset
  1237
apply(simp add: pt2[OF pt,symmetric])
0eccb98b1fdb added initialisation-code for finite_guess
urbanc
parents: 19140
diff changeset
  1238
apply(rule trans)
0eccb98b1fdb added initialisation-code for finite_guess
urbanc
parents: 19140
diff changeset
  1239
apply(rule pt3[OF pt])
0eccb98b1fdb added initialisation-code for finite_guess
urbanc
parents: 19140
diff changeset
  1240
apply(rule at_ds5'[OF at])
0eccb98b1fdb added initialisation-code for finite_guess
urbanc
parents: 19140
diff changeset
  1241
apply(rule pt1[OF pt])
0eccb98b1fdb added initialisation-code for finite_guess
urbanc
parents: 19140
diff changeset
  1242
done
0eccb98b1fdb added initialisation-code for finite_guess
urbanc
parents: 19140
diff changeset
  1243
24571
a6d0428dea8e some cleaning up to do with contexts
urbanc
parents: 24568
diff changeset
  1244
lemma pt_swap_bij'':
a6d0428dea8e some cleaning up to do with contexts
urbanc
parents: 24568
diff changeset
  1245
  fixes a  :: "'x"
a6d0428dea8e some cleaning up to do with contexts
urbanc
parents: 24568
diff changeset
  1246
  and   x  :: "'a"
a6d0428dea8e some cleaning up to do with contexts
urbanc
parents: 24568
diff changeset
  1247
  assumes pt: "pt TYPE('a) TYPE('x)"
a6d0428dea8e some cleaning up to do with contexts
urbanc
parents: 24568
diff changeset
  1248
  and     at: "at TYPE('x)"
a6d0428dea8e some cleaning up to do with contexts
urbanc
parents: 24568
diff changeset
  1249
  shows "[(a,a)]\<bullet>x = x"
a6d0428dea8e some cleaning up to do with contexts
urbanc
parents: 24568
diff changeset
  1250
apply(rule trans)
a6d0428dea8e some cleaning up to do with contexts
urbanc
parents: 24568
diff changeset
  1251
apply(rule pt3[OF pt])
a6d0428dea8e some cleaning up to do with contexts
urbanc
parents: 24568
diff changeset
  1252
apply(rule at_ds1[OF at])
a6d0428dea8e some cleaning up to do with contexts
urbanc
parents: 24568
diff changeset
  1253
apply(rule pt1[OF pt])
a6d0428dea8e some cleaning up to do with contexts
urbanc
parents: 24568
diff changeset
  1254
done
a6d0428dea8e some cleaning up to do with contexts
urbanc
parents: 24568
diff changeset
  1255
26806
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
  1256
lemma supp_singleton:
46179
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  1257
  shows "supp {x} = supp x"
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  1258
  by (force simp add: supp_def perm_set_def)
26806
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
  1259
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
  1260
lemma fresh_singleton:
46179
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  1261
  shows "a\<sharp>{x} = a\<sharp>x"
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  1262
  by (simp add: fresh_def supp_singleton)
26806
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
  1263
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1264
lemma pt_set_bij1:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1265
  fixes pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1266
  and   x  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1267
  and   X  :: "'a set"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1268
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1269
  and     at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1270
  shows "((pi\<bullet>x)\<in>X) = (x\<in>((rev pi)\<bullet>X))"
46179
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  1271
  by (force simp add: perm_set_def pt_rev_pi[OF pt, OF at] pt_pi_rev[OF pt, OF at])
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1272
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1273
lemma pt_set_bij1a:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1274
  fixes pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1275
  and   x  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1276
  and   X  :: "'a set"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1277
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1278
  and     at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1279
  shows "(x\<in>(pi\<bullet>X)) = (((rev pi)\<bullet>x)\<in>X)"
46179
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  1280
  by (force simp add: perm_set_def pt_rev_pi[OF pt, OF at] pt_pi_rev[OF pt, OF at])
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1281
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1282
lemma pt_set_bij:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1283
  fixes pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1284
  and   x  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1285
  and   X  :: "'a set"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1286
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1287
  and     at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1288
  shows "((pi\<bullet>x)\<in>(pi\<bullet>X)) = (x\<in>X)"
46179
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  1289
  by (simp add: perm_set_def pt_bij[OF pt, OF at])
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1290
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  1291
lemma pt_in_eqvt:
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  1292
  fixes pi :: "'x prm"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  1293
  and   x  :: "'a"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  1294
  and   X  :: "'a set"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  1295
  assumes pt: "pt TYPE('a) TYPE('x)"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  1296
  and     at: "at TYPE('x)"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  1297
  shows "pi\<bullet>(x\<in>X)=((pi\<bullet>x)\<in>(pi\<bullet>X))"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  1298
using assms
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  1299
by (auto simp add:  pt_set_bij perm_bool)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  1300
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1301
lemma pt_set_bij2:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1302
  fixes pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1303
  and   x  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1304
  and   X  :: "'a set"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1305
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1306
  and     at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1307
  and     a:  "x\<in>X"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1308
  shows "(pi\<bullet>x)\<in>(pi\<bullet>X)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1309
  using a by (simp add: pt_set_bij[OF pt, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1310
18264
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  1311
lemma pt_set_bij2a:
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  1312
  fixes pi :: "'x prm"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  1313
  and   x  :: "'a"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  1314
  and   X  :: "'a set"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  1315
  assumes pt: "pt TYPE('a) TYPE('x)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  1316
  and     at: "at TYPE('x)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  1317
  and     a:  "x\<in>((rev pi)\<bullet>X)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  1318
  shows "(pi\<bullet>x)\<in>X"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  1319
  using a by (simp add: pt_set_bij1[OF pt, OF at])
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  1320
26773
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26766
diff changeset
  1321
(* FIXME: is this lemma needed anywhere? *)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1322
lemma pt_set_bij3:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1323
  fixes pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1324
  and   x  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1325
  and   X  :: "'a set"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1326
  shows "pi\<bullet>(x\<in>X) = (x\<in>X)"
26773
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26766
diff changeset
  1327
by (simp add: perm_bool)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1328
18159
08282ca0402e added a few equivariance lemmas (they need to be automated
urbanc
parents: 18068
diff changeset
  1329
lemma pt_subseteq_eqvt:
08282ca0402e added a few equivariance lemmas (they need to be automated
urbanc
parents: 18068
diff changeset
  1330
  fixes pi :: "'x prm"
08282ca0402e added a few equivariance lemmas (they need to be automated
urbanc
parents: 18068
diff changeset
  1331
  and   Y  :: "'a set"
08282ca0402e added a few equivariance lemmas (they need to be automated
urbanc
parents: 18068
diff changeset
  1332
  and   X  :: "'a set"
08282ca0402e added a few equivariance lemmas (they need to be automated
urbanc
parents: 18068
diff changeset
  1333
  assumes pt: "pt TYPE('a) TYPE('x)"
08282ca0402e added a few equivariance lemmas (they need to be automated
urbanc
parents: 18068
diff changeset
  1334
  and     at: "at TYPE('x)"
26090
ec111fa4f8c5 added eqvt-flag to subseteq-lemma
urbanc
parents: 25950
diff changeset
  1335
  shows "(pi\<bullet>(X\<subseteq>Y)) = ((pi\<bullet>X)\<subseteq>(pi\<bullet>Y))"
46179
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  1336
by (auto simp add: perm_set_def perm_bool pt_bij[OF pt, OF at])
18159
08282ca0402e added a few equivariance lemmas (they need to be automated
urbanc
parents: 18068
diff changeset
  1337
19772
45897b49fdd2 added some further lemmas that deal with permutations and set-operators
urbanc
parents: 19771
diff changeset
  1338
lemma pt_set_diff_eqvt:
45897b49fdd2 added some further lemmas that deal with permutations and set-operators
urbanc
parents: 19771
diff changeset
  1339
  fixes X::"'a set"
45897b49fdd2 added some further lemmas that deal with permutations and set-operators
urbanc
parents: 19771
diff changeset
  1340
  and   Y::"'a set"
45897b49fdd2 added some further lemmas that deal with permutations and set-operators
urbanc
parents: 19771
diff changeset
  1341
  and   pi::"'x prm"
45897b49fdd2 added some further lemmas that deal with permutations and set-operators
urbanc
parents: 19771
diff changeset
  1342
  assumes pt: "pt TYPE('a) TYPE('x)"
45897b49fdd2 added some further lemmas that deal with permutations and set-operators
urbanc
parents: 19771
diff changeset
  1343
  and     at: "at TYPE('x)"
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22808
diff changeset
  1344
  shows "pi\<bullet>(X - Y) = (pi\<bullet>X) - (pi\<bullet>Y)"
46179
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  1345
  by (auto simp add: perm_set_def pt_bij[OF pt, OF at])
19772
45897b49fdd2 added some further lemmas that deal with permutations and set-operators
urbanc
parents: 19771
diff changeset
  1346
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22808
diff changeset
  1347
lemma pt_Collect_eqvt:
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22808
diff changeset
  1348
  fixes pi::"'x prm"
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22808
diff changeset
  1349
  assumes pt: "pt TYPE('a) TYPE('x)"
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22808
diff changeset
  1350
  and     at: "at TYPE('x)"
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22808
diff changeset
  1351
  shows "pi\<bullet>{x::'a. P x} = {x. P ((rev pi)\<bullet>x)}"
46179
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  1352
apply(auto simp add: perm_set_def pt_rev_pi[OF pt, OF at])
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22808
diff changeset
  1353
apply(rule_tac x="(rev pi)\<bullet>x" in exI)
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22808
diff changeset
  1354
apply(simp add: pt_pi_rev[OF pt, OF at])
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22808
diff changeset
  1355
done
19772
45897b49fdd2 added some further lemmas that deal with permutations and set-operators
urbanc
parents: 19771
diff changeset
  1356
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  1357
\<comment> "some helper lemmas for the pt_perm_supp_ineq lemma"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1358
lemma Collect_permI: 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1359
  fixes pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1360
  and   x  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1361
  assumes a: "\<forall>x. (P1 x = P2 x)" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1362
  shows "{pi\<bullet>x| x. P1 x} = {pi\<bullet>x| x. P2 x}"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1363
  using a by force
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1364
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1365
lemma Infinite_cong:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1366
  assumes a: "X = Y"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1367
  shows "infinite X = infinite Y"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1368
  using a by (simp)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1369
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1370
lemma pt_set_eq_ineq:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1371
  fixes pi :: "'y prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1372
  assumes pt: "pt TYPE('x) TYPE('y)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1373
  and     at: "at TYPE('y)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1374
  shows "{pi\<bullet>x| x::'x. P x} = {x::'x. P ((rev pi)\<bullet>x)}"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1375
  by (force simp only: pt_rev_pi[OF pt, OF at] pt_pi_rev[OF pt, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1376
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1377
lemma pt_inject_on_ineq:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1378
  fixes X  :: "'y set"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1379
  and   pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1380
  assumes pt: "pt TYPE('y) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1381
  and     at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1382
  shows "inj_on (perm pi) X"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1383
proof (unfold inj_on_def, intro strip)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1384
  fix x::"'y" and y::"'y"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1385
  assume "pi\<bullet>x = pi\<bullet>y"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1386
  thus "x=y" by (simp add: pt_bij[OF pt, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1387
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1388
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1389
lemma pt_set_finite_ineq: 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1390
  fixes X  :: "'x set"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1391
  and   pi :: "'y prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1392
  assumes pt: "pt TYPE('x) TYPE('y)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1393
  and     at: "at TYPE('y)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1394
  shows "finite (pi\<bullet>X) = finite X"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1395
proof -
46179
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  1396
  have image: "(pi\<bullet>X) = (perm pi ` X)" by (force simp only: perm_set_def)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1397
  show ?thesis
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1398
  proof (rule iffI)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1399
    assume "finite (pi\<bullet>X)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1400
    hence "finite (perm pi ` X)" using image by (simp)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1401
    thus "finite X" using pt_inject_on_ineq[OF pt, OF at] by (rule finite_imageD)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1402
  next
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1403
    assume "finite X"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1404
    hence "finite (perm pi ` X)" by (rule finite_imageI)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1405
    thus "finite (pi\<bullet>X)" using image by (simp)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1406
  qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1407
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1408
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1409
lemma pt_set_infinite_ineq: 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1410
  fixes X  :: "'x set"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1411
  and   pi :: "'y prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1412
  assumes pt: "pt TYPE('x) TYPE('y)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1413
  and     at: "at TYPE('y)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1414
  shows "infinite (pi\<bullet>X) = infinite X"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1415
using pt at by (simp add: pt_set_finite_ineq)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1416
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1417
lemma pt_perm_supp_ineq:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1418
  fixes  pi  :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1419
  and    x   :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1420
  assumes pta: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1421
  and     ptb: "pt TYPE('y) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1422
  and     at:  "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1423
  and     cp:  "cp TYPE('a) TYPE('x) TYPE('y)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1424
  shows "(pi\<bullet>((supp x)::'y set)) = supp (pi\<bullet>x)" (is "?LHS = ?RHS")
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1425
proof -
46179
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  1426
  have "?LHS = {pi\<bullet>a | a. infinite {b. [(a,b)]\<bullet>x \<noteq> x}}" by (simp add: supp_def perm_set_def)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1427
  also have "\<dots> = {pi\<bullet>a | a. infinite {pi\<bullet>b | b. [(a,b)]\<bullet>x \<noteq> x}}" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1428
  proof (rule Collect_permI, rule allI, rule iffI)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1429
    fix a
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1430
    assume "infinite {b::'y. [(a,b)]\<bullet>x  \<noteq> x}"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1431
    hence "infinite (pi\<bullet>{b::'y. [(a,b)]\<bullet>x \<noteq> x})" by (simp add: pt_set_infinite_ineq[OF ptb, OF at])
46179
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  1432
    thus "infinite {pi\<bullet>b |b::'y. [(a,b)]\<bullet>x  \<noteq> x}" by (simp add: perm_set_def)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1433
  next
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1434
    fix a
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1435
    assume "infinite {pi\<bullet>b |b::'y. [(a,b)]\<bullet>x \<noteq> x}"
46179
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  1436
    hence "infinite (pi\<bullet>{b::'y. [(a,b)]\<bullet>x \<noteq> x})" by (simp add: perm_set_def)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1437
    thus "infinite {b::'y. [(a,b)]\<bullet>x  \<noteq> x}" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1438
      by (simp add: pt_set_infinite_ineq[OF ptb, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1439
  qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1440
  also have "\<dots> = {a. infinite {b::'y. [((rev pi)\<bullet>a,(rev pi)\<bullet>b)]\<bullet>x \<noteq> x}}" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1441
    by (simp add: pt_set_eq_ineq[OF ptb, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1442
  also have "\<dots> = {a. infinite {b. pi\<bullet>([((rev pi)\<bullet>a,(rev pi)\<bullet>b)]\<bullet>x) \<noteq> (pi\<bullet>x)}}"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1443
    by (simp add: pt_bij[OF pta, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1444
  also have "\<dots> = {a. infinite {b. [(a,b)]\<bullet>(pi\<bullet>x) \<noteq> (pi\<bullet>x)}}"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1445
  proof (rule Collect_cong, rule Infinite_cong, rule Collect_cong)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1446
    fix a::"'y" and b::"'y"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1447
    have "pi\<bullet>(([((rev pi)\<bullet>a,(rev pi)\<bullet>b)])\<bullet>x) = [(a,b)]\<bullet>(pi\<bullet>x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1448
      by (simp add: cp1[OF cp] pt_pi_rev[OF ptb, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1449
    thus "(pi\<bullet>([((rev pi)\<bullet>a,(rev pi)\<bullet>b)]\<bullet>x) \<noteq>  pi\<bullet>x) = ([(a,b)]\<bullet>(pi\<bullet>x) \<noteq> pi\<bullet>x)" by simp
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1450
  qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1451
  finally show "?LHS = ?RHS" by (simp add: supp_def) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1452
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1453
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1454
lemma pt_perm_supp:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1455
  fixes  pi  :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1456
  and    x   :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1457
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1458
  and     at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1459
  shows "(pi\<bullet>((supp x)::'x set)) = supp (pi\<bullet>x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1460
apply(rule pt_perm_supp_ineq)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1461
apply(rule pt)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1462
apply(rule at_pt_inst)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1463
apply(rule at)+
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1464
apply(rule cp_pt_inst)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1465
apply(rule pt)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1466
apply(rule at)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1467
done
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1468
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1469
lemma pt_supp_finite_pi:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1470
  fixes  pi  :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1471
  and    x   :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1472
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1473
  and     at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1474
  and     f: "finite ((supp x)::'x set)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1475
  shows "finite ((supp (pi\<bullet>x))::'x set)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1476
apply(simp add: pt_perm_supp[OF pt, OF at, symmetric])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1477
apply(simp add: pt_set_finite_ineq[OF at_pt_inst[OF at], OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1478
apply(rule f)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1479
done
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1480
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1481
lemma pt_fresh_left_ineq:  
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1482
  fixes  pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1483
  and     x :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1484
  and     a :: "'y"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1485
  assumes pta: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1486
  and     ptb: "pt TYPE('y) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1487
  and     at:  "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1488
  and     cp:  "cp TYPE('a) TYPE('x) TYPE('y)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1489
  shows "a\<sharp>(pi\<bullet>x) = ((rev pi)\<bullet>a)\<sharp>x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1490
apply(simp add: fresh_def)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1491
apply(simp add: pt_set_bij1[OF ptb, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1492
apply(simp add: pt_perm_supp_ineq[OF pta, OF ptb, OF at, OF cp])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1493
done
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1494
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1495
lemma pt_fresh_right_ineq:  
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1496
  fixes  pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1497
  and     x :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1498
  and     a :: "'y"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1499
  assumes pta: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1500
  and     ptb: "pt TYPE('y) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1501
  and     at:  "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1502
  and     cp:  "cp TYPE('a) TYPE('x) TYPE('y)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1503
  shows "(pi\<bullet>a)\<sharp>x = a\<sharp>((rev pi)\<bullet>x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1504
apply(simp add: fresh_def)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1505
apply(simp add: pt_set_bij1[OF ptb, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1506
apply(simp add: pt_perm_supp_ineq[OF pta, OF ptb, OF at, OF cp])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1507
done
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1508
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1509
lemma pt_fresh_bij_ineq:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1510
  fixes  pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1511
  and     x :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1512
  and     a :: "'y"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1513
  assumes pta: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1514
  and     ptb: "pt TYPE('y) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1515
  and     at:  "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1516
  and     cp:  "cp TYPE('a) TYPE('x) TYPE('y)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1517
  shows "(pi\<bullet>a)\<sharp>(pi\<bullet>x) = a\<sharp>x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1518
apply(simp add: pt_fresh_left_ineq[OF pta, OF ptb, OF at, OF cp])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1519
apply(simp add: pt_rev_pi[OF ptb, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1520
done
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1521
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1522
lemma pt_fresh_left:  
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1523
  fixes  pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1524
  and     x :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1525
  and     a :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1526
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1527
  and     at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1528
  shows "a\<sharp>(pi\<bullet>x) = ((rev pi)\<bullet>a)\<sharp>x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1529
apply(rule pt_fresh_left_ineq)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1530
apply(rule pt)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1531
apply(rule at_pt_inst)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1532
apply(rule at)+
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1533
apply(rule cp_pt_inst)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1534
apply(rule pt)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1535
apply(rule at)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1536
done
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1537
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1538
lemma pt_fresh_right:  
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1539
  fixes  pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1540
  and     x :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1541
  and     a :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1542
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1543
  and     at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1544
  shows "(pi\<bullet>a)\<sharp>x = a\<sharp>((rev pi)\<bullet>x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1545
apply(rule pt_fresh_right_ineq)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1546
apply(rule pt)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1547
apply(rule at_pt_inst)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1548
apply(rule at)+
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1549
apply(rule cp_pt_inst)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1550
apply(rule pt)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1551
apply(rule at)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1552
done
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1553
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1554
lemma pt_fresh_bij:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1555
  fixes  pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1556
  and     x :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1557
  and     a :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1558
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1559
  and     at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1560
  shows "(pi\<bullet>a)\<sharp>(pi\<bullet>x) = a\<sharp>x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1561
apply(rule pt_fresh_bij_ineq)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1562
apply(rule pt)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1563
apply(rule at_pt_inst)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1564
apply(rule at)+
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1565
apply(rule cp_pt_inst)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1566
apply(rule pt)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1567
apply(rule at)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1568
done
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1569
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1570
lemma pt_fresh_bij1:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1571
  fixes  pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1572
  and     x :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1573
  and     a :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1574
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1575
  and     at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1576
  and     a:  "a\<sharp>x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1577
  shows "(pi\<bullet>a)\<sharp>(pi\<bullet>x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1578
using a by (simp add: pt_fresh_bij[OF pt, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1579
19566
63e18ed22fda added the lemma pt_fresh_bij2
urbanc
parents: 19562
diff changeset
  1580
lemma pt_fresh_bij2:
63e18ed22fda added the lemma pt_fresh_bij2
urbanc
parents: 19562
diff changeset
  1581
  fixes  pi :: "'x prm"
63e18ed22fda added the lemma pt_fresh_bij2
urbanc
parents: 19562
diff changeset
  1582
  and     x :: "'a"
63e18ed22fda added the lemma pt_fresh_bij2
urbanc
parents: 19562
diff changeset
  1583
  and     a :: "'x"
63e18ed22fda added the lemma pt_fresh_bij2
urbanc
parents: 19562
diff changeset
  1584
  assumes pt: "pt TYPE('a) TYPE('x)"
63e18ed22fda added the lemma pt_fresh_bij2
urbanc
parents: 19562
diff changeset
  1585
  and     at: "at TYPE('x)"
63e18ed22fda added the lemma pt_fresh_bij2
urbanc
parents: 19562
diff changeset
  1586
  and     a:  "(pi\<bullet>a)\<sharp>(pi\<bullet>x)"
63e18ed22fda added the lemma pt_fresh_bij2
urbanc
parents: 19562
diff changeset
  1587
  shows  "a\<sharp>x"
63e18ed22fda added the lemma pt_fresh_bij2
urbanc
parents: 19562
diff changeset
  1588
using a by (simp add: pt_fresh_bij[OF pt, OF at])
63e18ed22fda added the lemma pt_fresh_bij2
urbanc
parents: 19562
diff changeset
  1589
19972
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
  1590
lemma pt_fresh_eqvt:
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
  1591
  fixes  pi :: "'x prm"
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
  1592
  and     x :: "'a"
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
  1593
  and     a :: "'x"
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
  1594
  assumes pt: "pt TYPE('a) TYPE('x)"
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
  1595
  and     at: "at TYPE('x)"
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
  1596
  shows "pi\<bullet>(a\<sharp>x) = (pi\<bullet>a)\<sharp>(pi\<bullet>x)"
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
  1597
  by (simp add: perm_bool pt_fresh_bij[OF pt, OF at])
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
  1598
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1599
lemma pt_perm_fresh1:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1600
  fixes a :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1601
  and   b :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1602
  and   x :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1603
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1604
  and     at: "at TYPE ('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1605
  and     a1: "\<not>(a\<sharp>x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1606
  and     a2: "b\<sharp>x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1607
  shows "[(a,b)]\<bullet>x \<noteq> x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1608
proof
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1609
  assume neg: "[(a,b)]\<bullet>x = x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1610
  from a1 have a1':"a\<in>(supp x)" by (simp add: fresh_def) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1611
  from a2 have a2':"b\<notin>(supp x)" by (simp add: fresh_def) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1612
  from a1' a2' have a3: "a\<noteq>b" by force
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1613
  from a1' have "([(a,b)]\<bullet>a)\<in>([(a,b)]\<bullet>(supp x))" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1614
    by (simp only: pt_set_bij[OF at_pt_inst[OF at], OF at])
19325
35177b864f80 tuned some proofs
urbanc
parents: 19216
diff changeset
  1615
  hence "b\<in>([(a,b)]\<bullet>(supp x))" by (simp add: at_calc[OF at])
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1616
  hence "b\<in>(supp ([(a,b)]\<bullet>x))" by (simp add: pt_perm_supp[OF pt,OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1617
  with a2' neg show False by simp
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1618
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1619
19638
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
  1620
(* the next two lemmas are needed in the proof *)
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
  1621
(* of the structural induction principle       *)
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
  1622
lemma pt_fresh_aux:
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
  1623
  fixes a::"'x"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
  1624
  and   b::"'x"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
  1625
  and   c::"'x"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
  1626
  and   x::"'a"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
  1627
  assumes pt: "pt TYPE('a) TYPE('x)"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
  1628
  and     at: "at TYPE ('x)"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
  1629
  assumes a1: "c\<noteq>a" and  a2: "a\<sharp>x" and a3: "c\<sharp>x"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
  1630
  shows "c\<sharp>([(a,b)]\<bullet>x)"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
  1631
using a1 a2 a3 by (simp_all add: pt_fresh_left[OF pt, OF at] at_calc[OF at])
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
  1632
22786
d8d7a53ffb63 fixes last commit
narboux
parents: 22785
diff changeset
  1633
lemma pt_fresh_perm_app:
d8d7a53ffb63 fixes last commit
narboux
parents: 22785
diff changeset
  1634
  fixes pi :: "'x prm" 
d8d7a53ffb63 fixes last commit
narboux
parents: 22785
diff changeset
  1635
  and   a  :: "'x"
d8d7a53ffb63 fixes last commit
narboux
parents: 22785
diff changeset
  1636
  and   x  :: "'y"
d8d7a53ffb63 fixes last commit
narboux
parents: 22785
diff changeset
  1637
  assumes pt: "pt TYPE('y) TYPE('x)"
d8d7a53ffb63 fixes last commit
narboux
parents: 22785
diff changeset
  1638
  and     at: "at TYPE('x)"
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22808
diff changeset
  1639
  and     h1: "a\<sharp>pi"
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22808
diff changeset
  1640
  and     h2: "a\<sharp>x"
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22808
diff changeset
  1641
  shows "a\<sharp>(pi\<bullet>x)"
22786
d8d7a53ffb63 fixes last commit
narboux
parents: 22785
diff changeset
  1642
using assms
d8d7a53ffb63 fixes last commit
narboux
parents: 22785
diff changeset
  1643
proof -
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22808
diff changeset
  1644
  have "a\<sharp>(rev pi)"using h1 by (simp add: fresh_list_rev)
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22808
diff changeset
  1645
  then have "(rev pi)\<bullet>a = a" by (simp add: at_prm_fresh[OF at])
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22808
diff changeset
  1646
  then have "((rev pi)\<bullet>a)\<sharp>x" using h2 by simp
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22808
diff changeset
  1647
  thus "a\<sharp>(pi\<bullet>x)"  by (simp add: pt_fresh_right[OF pt, OF at])
22786
d8d7a53ffb63 fixes last commit
narboux
parents: 22785
diff changeset
  1648
qed
d8d7a53ffb63 fixes last commit
narboux
parents: 22785
diff changeset
  1649
d8d7a53ffb63 fixes last commit
narboux
parents: 22785
diff changeset
  1650
lemma pt_fresh_perm_app_ineq:
19638
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
  1651
  fixes pi::"'x prm"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
  1652
  and   c::"'y"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
  1653
  and   x::"'a"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
  1654
  assumes pta: "pt TYPE('a) TYPE('x)"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
  1655
  and     ptb: "pt TYPE('y) TYPE('x)"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
  1656
  and     at:  "at TYPE('x)"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
  1657
  and     cp:  "cp TYPE('a) TYPE('x) TYPE('y)"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
  1658
  and     dj:  "disjoint TYPE('y) TYPE('x)"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
  1659
  assumes a: "c\<sharp>x"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
  1660
  shows "c\<sharp>(pi\<bullet>x)"
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
  1661
using a by (simp add: pt_fresh_left_ineq[OF pta, OF ptb, OF at, OF cp] dj_perm_forget[OF dj])
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
  1662
22535
cbee450f88a6 added extended the lemma for equivariance of freshness
urbanc
parents: 22514
diff changeset
  1663
lemma pt_fresh_eqvt_ineq:
cbee450f88a6 added extended the lemma for equivariance of freshness
urbanc
parents: 22514
diff changeset
  1664
  fixes pi::"'x prm"
cbee450f88a6 added extended the lemma for equivariance of freshness
urbanc
parents: 22514
diff changeset
  1665
  and   c::"'y"
cbee450f88a6 added extended the lemma for equivariance of freshness
urbanc
parents: 22514
diff changeset
  1666
  and   x::"'a"
cbee450f88a6 added extended the lemma for equivariance of freshness
urbanc
parents: 22514
diff changeset
  1667
  assumes pta: "pt TYPE('a) TYPE('x)"
cbee450f88a6 added extended the lemma for equivariance of freshness
urbanc
parents: 22514
diff changeset
  1668
  and     ptb: "pt TYPE('y) TYPE('x)"
cbee450f88a6 added extended the lemma for equivariance of freshness
urbanc
parents: 22514
diff changeset
  1669
  and     at:  "at TYPE('x)"
cbee450f88a6 added extended the lemma for equivariance of freshness
urbanc
parents: 22514
diff changeset
  1670
  and     cp:  "cp TYPE('a) TYPE('x) TYPE('y)"
cbee450f88a6 added extended the lemma for equivariance of freshness
urbanc
parents: 22514
diff changeset
  1671
  and     dj:  "disjoint TYPE('y) TYPE('x)"
cbee450f88a6 added extended the lemma for equivariance of freshness
urbanc
parents: 22514
diff changeset
  1672
  shows "pi\<bullet>(c\<sharp>x) = (pi\<bullet>c)\<sharp>(pi\<bullet>x)"
cbee450f88a6 added extended the lemma for equivariance of freshness
urbanc
parents: 22514
diff changeset
  1673
by (simp add: pt_fresh_left_ineq[OF pta, OF ptb, OF at, OF cp] dj_perm_forget[OF dj] perm_bool)
cbee450f88a6 added extended the lemma for equivariance of freshness
urbanc
parents: 22514
diff changeset
  1674
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  1675
\<comment>"the co-set of a finite set is infinte"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1676
lemma finite_infinite:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1677
  assumes a: "finite {b::'x. P b}"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1678
  and     b: "infinite (UNIV::'x set)"        
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1679
  shows "infinite {b. \<not>P b}"
27687
224a18d1a3ac simplified a proof
urbanc
parents: 27374
diff changeset
  1680
proof -
224a18d1a3ac simplified a proof
urbanc
parents: 27374
diff changeset
  1681
  from a b have "infinite (UNIV - {b::'x. P b})" by (simp add: Diff_infinite_finite)
224a18d1a3ac simplified a proof
urbanc
parents: 27374
diff changeset
  1682
  moreover 
224a18d1a3ac simplified a proof
urbanc
parents: 27374
diff changeset
  1683
  have "{b::'x. \<not>P b} = UNIV - {b::'x. P b}" by auto
224a18d1a3ac simplified a proof
urbanc
parents: 27374
diff changeset
  1684
  ultimately show "infinite {b::'x. \<not>P b}" by simp
224a18d1a3ac simplified a proof
urbanc
parents: 27374
diff changeset
  1685
qed 
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1686
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1687
lemma pt_fresh_fresh:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1688
  fixes   x :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1689
  and     a :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1690
  and     b :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1691
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1692
  and     at: "at TYPE ('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1693
  and     a1: "a\<sharp>x" and a2: "b\<sharp>x" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1694
  shows "[(a,b)]\<bullet>x=x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1695
proof (cases "a=b")
19325
35177b864f80 tuned some proofs
urbanc
parents: 19216
diff changeset
  1696
  assume "a=b"
35177b864f80 tuned some proofs
urbanc
parents: 19216
diff changeset
  1697
  hence "[(a,b)] \<triangleq> []" by (simp add: at_ds1[OF at])
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1698
  hence "[(a,b)]\<bullet>x=([]::'x prm)\<bullet>x" by (rule pt3[OF pt])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1699
  thus ?thesis by (simp only: pt1[OF pt])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1700
next
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1701
  assume c2: "a\<noteq>b"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1702
  from a1 have f1: "finite {c. [(a,c)]\<bullet>x \<noteq> x}" by (simp add: fresh_def supp_def)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1703
  from a2 have f2: "finite {c. [(b,c)]\<bullet>x \<noteq> x}" by (simp add: fresh_def supp_def)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1704
  from f1 and f2 have f3: "finite {c. perm [(a,c)] x \<noteq> x \<or> perm [(b,c)] x \<noteq> x}" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1705
    by (force simp only: Collect_disj_eq)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1706
  have "infinite {c. [(a,c)]\<bullet>x = x \<and> [(b,c)]\<bullet>x = x}" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1707
    by (simp add: finite_infinite[OF f3,OF at4[OF at], simplified])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1708
  hence "infinite ({c. [(a,c)]\<bullet>x = x \<and> [(b,c)]\<bullet>x = x}-{a,b})" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1709
    by (force dest: Diff_infinite_finite)
29903
2c0046b26f80 more finiteness changes
nipkow
parents: 29128
diff changeset
  1710
  hence "({c. [(a,c)]\<bullet>x = x \<and> [(b,c)]\<bullet>x = x}-{a,b}) \<noteq> {}"
44683
daeb538c57bf tuned specifications and proofs
haftmann
parents: 44567
diff changeset
  1711
    by (metis finite_set set_empty2)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1712
  hence "\<exists>c. c\<in>({c. [(a,c)]\<bullet>x = x \<and> [(b,c)]\<bullet>x = x}-{a,b})" by (force)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1713
  then obtain c 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1714
    where eq1: "[(a,c)]\<bullet>x = x" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1715
      and eq2: "[(b,c)]\<bullet>x = x" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1716
      and ineq: "a\<noteq>c \<and> b\<noteq>c"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1717
    by (force)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1718
  hence "[(a,c)]\<bullet>([(b,c)]\<bullet>([(a,c)]\<bullet>x)) = x" by simp 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1719
  hence eq3: "[(a,c),(b,c),(a,c)]\<bullet>x = x" by (simp add: pt2[OF pt,symmetric])
18295
dd50de393330 changed \<sim> of permutation equality to \<triangleq>
urbanc
parents: 18294
diff changeset
  1720
  from c2 ineq have "[(a,c),(b,c),(a,c)] \<triangleq> [(a,b)]" by (simp add: at_ds3[OF at])
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1721
  hence "[(a,c),(b,c),(a,c)]\<bullet>x = [(a,b)]\<bullet>x" by (rule pt3[OF pt])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1722
  thus ?thesis using eq3 by simp
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1723
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1724
26773
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26766
diff changeset
  1725
lemma pt_pi_fresh_fresh:
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26766
diff changeset
  1726
  fixes   x :: "'a"
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26766
diff changeset
  1727
  and     pi :: "'x prm"
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26766
diff changeset
  1728
  assumes pt: "pt TYPE('a) TYPE('x)"
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26766
diff changeset
  1729
  and     at: "at TYPE ('x)"
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26766
diff changeset
  1730
  and     a:  "\<forall>(a,b)\<in>set pi. a\<sharp>x \<and> b\<sharp>x" 
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26766
diff changeset
  1731
  shows "pi\<bullet>x=x"
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26766
diff changeset
  1732
using a
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26766
diff changeset
  1733
proof (induct pi)
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26766
diff changeset
  1734
  case Nil
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26766
diff changeset
  1735
  show "([]::'x prm)\<bullet>x = x" by (rule pt1[OF pt])
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26766
diff changeset
  1736
next
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26766
diff changeset
  1737
  case (Cons ab pi)
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26766
diff changeset
  1738
  have a: "\<forall>(a,b)\<in>set (ab#pi). a\<sharp>x \<and> b\<sharp>x" by fact
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26766
diff changeset
  1739
  have ih: "(\<forall>(a,b)\<in>set pi. a\<sharp>x \<and> b\<sharp>x) \<Longrightarrow> pi\<bullet>x=x" by fact
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26766
diff changeset
  1740
  obtain a b where e: "ab=(a,b)" by (cases ab) (auto)
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26766
diff changeset
  1741
  from a have a': "a\<sharp>x" "b\<sharp>x" using e by auto
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26766
diff changeset
  1742
  have "(ab#pi)\<bullet>x = ([(a,b)]@pi)\<bullet>x" using e by simp
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26766
diff changeset
  1743
  also have "\<dots> = [(a,b)]\<bullet>(pi\<bullet>x)" by (simp only: pt2[OF pt])
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26766
diff changeset
  1744
  also have "\<dots> = [(a,b)]\<bullet>x" using ih a by simp
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26766
diff changeset
  1745
  also have "\<dots> = x" using a' by (simp add: pt_fresh_fresh[OF pt, OF at])
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26766
diff changeset
  1746
  finally show "(ab#pi)\<bullet>x = x" by simp
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26766
diff changeset
  1747
qed
ba8b1a8a12a7 added more infrastructure for fresh_star
urbanc
parents: 26766
diff changeset
  1748
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1749
lemma pt_perm_compose:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1750
  fixes pi1 :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1751
  and   pi2 :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1752
  and   x  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1753
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1754
  and     at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1755
  shows "pi2\<bullet>(pi1\<bullet>x) = (pi2\<bullet>pi1)\<bullet>(pi2\<bullet>x)" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1756
proof -
23393
31781b2de73d tuned proofs: avoid implicit prems;
wenzelm
parents: 23159
diff changeset
  1757
  have "(pi2@pi1) \<triangleq> ((pi2\<bullet>pi1)@pi2)" by (rule at_ds8 [OF at])
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1758
  hence "(pi2@pi1)\<bullet>x = ((pi2\<bullet>pi1)@pi2)\<bullet>x" by (rule pt3[OF pt])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1759
  thus ?thesis by (simp add: pt2[OF pt])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1760
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1761
19045
75786c2eb897 added lemma pt_perm_compose'
urbanc
parents: 18745
diff changeset
  1762
lemma pt_perm_compose':
75786c2eb897 added lemma pt_perm_compose'
urbanc
parents: 18745
diff changeset
  1763
  fixes pi1 :: "'x prm"
75786c2eb897 added lemma pt_perm_compose'
urbanc
parents: 18745
diff changeset
  1764
  and   pi2 :: "'x prm"
75786c2eb897 added lemma pt_perm_compose'
urbanc
parents: 18745
diff changeset
  1765
  and   x  :: "'a"
75786c2eb897 added lemma pt_perm_compose'
urbanc
parents: 18745
diff changeset
  1766
  assumes pt: "pt TYPE('a) TYPE('x)"
75786c2eb897 added lemma pt_perm_compose'
urbanc
parents: 18745
diff changeset
  1767
  and     at: "at TYPE('x)"
75786c2eb897 added lemma pt_perm_compose'
urbanc
parents: 18745
diff changeset
  1768
  shows "(pi2\<bullet>pi1)\<bullet>x = pi2\<bullet>(pi1\<bullet>((rev pi2)\<bullet>x))" 
75786c2eb897 added lemma pt_perm_compose'
urbanc
parents: 18745
diff changeset
  1769
proof -
75786c2eb897 added lemma pt_perm_compose'
urbanc
parents: 18745
diff changeset
  1770
  have "pi2\<bullet>(pi1\<bullet>((rev pi2)\<bullet>x)) = (pi2\<bullet>pi1)\<bullet>(pi2\<bullet>((rev pi2)\<bullet>x))"
75786c2eb897 added lemma pt_perm_compose'
urbanc
parents: 18745
diff changeset
  1771
    by (rule pt_perm_compose[OF pt, OF at])
75786c2eb897 added lemma pt_perm_compose'
urbanc
parents: 18745
diff changeset
  1772
  also have "\<dots> = (pi2\<bullet>pi1)\<bullet>x" by (simp add: pt_pi_rev[OF pt, OF at])
75786c2eb897 added lemma pt_perm_compose'
urbanc
parents: 18745
diff changeset
  1773
  finally have "pi2\<bullet>(pi1\<bullet>((rev pi2)\<bullet>x)) = (pi2\<bullet>pi1)\<bullet>x" by simp
75786c2eb897 added lemma pt_perm_compose'
urbanc
parents: 18745
diff changeset
  1774
  thus ?thesis by simp
75786c2eb897 added lemma pt_perm_compose'
urbanc
parents: 18745
diff changeset
  1775
qed
75786c2eb897 added lemma pt_perm_compose'
urbanc
parents: 18745
diff changeset
  1776
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1777
lemma pt_perm_compose_rev:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1778
  fixes pi1 :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1779
  and   pi2 :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1780
  and   x  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1781
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1782
  and     at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1783
  shows "(rev pi2)\<bullet>((rev pi1)\<bullet>x) = (rev pi1)\<bullet>(rev (pi1\<bullet>pi2)\<bullet>x)" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1784
proof -
18295
dd50de393330 changed \<sim> of permutation equality to \<triangleq>
urbanc
parents: 18294
diff changeset
  1785
  have "((rev pi2)@(rev pi1)) \<triangleq> ((rev pi1)@(rev (pi1\<bullet>pi2)))" by (rule at_ds9[OF at])
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1786
  hence "((rev pi2)@(rev pi1))\<bullet>x = ((rev pi1)@(rev (pi1\<bullet>pi2)))\<bullet>x" by (rule pt3[OF pt])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1787
  thus ?thesis by (simp add: pt2[OF pt])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1788
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1789
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  1790
section \<open>equivariance for some connectives\<close>
19972
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
  1791
lemma pt_all_eqvt:
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
  1792
  fixes  pi :: "'x prm"
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
  1793
  and     x :: "'a"
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
  1794
  assumes pt: "pt TYPE('a) TYPE('x)"
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
  1795
  and     at: "at TYPE('x)"
22715
381e6c45f13b improved the equivariance lemmas for the quantifiers; had to export the lemma eqvt_force_add and eqvt_force_del in the thmdecls
urbanc
parents: 22714
diff changeset
  1796
  shows "pi\<bullet>(\<forall>(x::'a). P x) = (\<forall>(x::'a). pi\<bullet>(P ((rev pi)\<bullet>x)))"
19972
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
  1797
apply(auto simp add: perm_bool perm_fun_def)
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
  1798
apply(drule_tac x="pi\<bullet>x" in spec)
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
  1799
apply(simp add: pt_rev_pi[OF pt, OF at])
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
  1800
done
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19869
diff changeset
  1801
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  1802
lemma pt_ex_eqvt:
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  1803
  fixes  pi :: "'x prm"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  1804
  and     x :: "'a"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  1805
  assumes pt: "pt TYPE('a) TYPE('x)"
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  1806
  and     at: "at TYPE('x)"
22715
381e6c45f13b improved the equivariance lemmas for the quantifiers; had to export the lemma eqvt_force_add and eqvt_force_del in the thmdecls
urbanc
parents: 22714
diff changeset
  1807
  shows "pi\<bullet>(\<exists>(x::'a). P x) = (\<exists>(x::'a). pi\<bullet>(P ((rev pi)\<bullet>x)))"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  1808
apply(auto simp add: perm_bool perm_fun_def)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  1809
apply(rule_tac x="pi\<bullet>x" in exI) 
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  1810
apply(simp add: pt_rev_pi[OF pt, OF at])
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  1811
done
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  1812
28011
90074908db16 added equivariance lemmas for ex1 and the
urbanc
parents: 27687
diff changeset
  1813
lemma pt_ex1_eqvt:
90074908db16 added equivariance lemmas for ex1 and the
urbanc
parents: 27687
diff changeset
  1814
  fixes  pi :: "'x prm"
90074908db16 added equivariance lemmas for ex1 and the
urbanc
parents: 27687
diff changeset
  1815
  and     x :: "'a"
90074908db16 added equivariance lemmas for ex1 and the
urbanc
parents: 27687
diff changeset
  1816
  assumes pt: "pt TYPE('a) TYPE('x)"
90074908db16 added equivariance lemmas for ex1 and the
urbanc
parents: 27687
diff changeset
  1817
  and     at: "at TYPE('x)"
90074908db16 added equivariance lemmas for ex1 and the
urbanc
parents: 27687
diff changeset
  1818
  shows  "(pi\<bullet>(\<exists>!x. P (x::'a))) = (\<exists>!x. pi\<bullet>(P (rev pi\<bullet>x)))"
90074908db16 added equivariance lemmas for ex1 and the
urbanc
parents: 27687
diff changeset
  1819
unfolding Ex1_def
90074908db16 added equivariance lemmas for ex1 and the
urbanc
parents: 27687
diff changeset
  1820
by (simp add: pt_ex_eqvt[OF pt at] conj_eqvt pt_all_eqvt[OF pt at] 
90074908db16 added equivariance lemmas for ex1 and the
urbanc
parents: 27687
diff changeset
  1821
              imp_eqvt pt_eq_eqvt[OF pt at] pt_pi_rev[OF pt at])
90074908db16 added equivariance lemmas for ex1 and the
urbanc
parents: 27687
diff changeset
  1822
90074908db16 added equivariance lemmas for ex1 and the
urbanc
parents: 27687
diff changeset
  1823
lemma pt_the_eqvt:
90074908db16 added equivariance lemmas for ex1 and the
urbanc
parents: 27687
diff changeset
  1824
  fixes  pi :: "'x prm"
90074908db16 added equivariance lemmas for ex1 and the
urbanc
parents: 27687
diff changeset
  1825
  assumes pt: "pt TYPE('a) TYPE('x)"
90074908db16 added equivariance lemmas for ex1 and the
urbanc
parents: 27687
diff changeset
  1826
  and     at: "at TYPE('x)"
90074908db16 added equivariance lemmas for ex1 and the
urbanc
parents: 27687
diff changeset
  1827
  and     unique: "\<exists>!x. P x"
90074908db16 added equivariance lemmas for ex1 and the
urbanc
parents: 27687
diff changeset
  1828
  shows "pi\<bullet>(THE(x::'a). P x) = (THE(x::'a). pi\<bullet>(P ((rev pi)\<bullet>x)))"
90074908db16 added equivariance lemmas for ex1 and the
urbanc
parents: 27687
diff changeset
  1829
  apply(rule the1_equality [symmetric])
90074908db16 added equivariance lemmas for ex1 and the
urbanc
parents: 27687
diff changeset
  1830
  apply(simp add: pt_ex1_eqvt[OF pt at,symmetric])
90074908db16 added equivariance lemmas for ex1 and the
urbanc
parents: 27687
diff changeset
  1831
  apply(simp add: perm_bool unique)
90074908db16 added equivariance lemmas for ex1 and the
urbanc
parents: 27687
diff changeset
  1832
  apply(simp add: perm_bool pt_rev_pi [OF pt at])
90074908db16 added equivariance lemmas for ex1 and the
urbanc
parents: 27687
diff changeset
  1833
  apply(rule theI'[OF unique])
90074908db16 added equivariance lemmas for ex1 and the
urbanc
parents: 27687
diff changeset
  1834
  done
90074908db16 added equivariance lemmas for ex1 and the
urbanc
parents: 27687
diff changeset
  1835
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  1836
section \<open>facts about supports\<close>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1837
(*==============================*)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1838
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1839
lemma supports_subset:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1840
  fixes x  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1841
  and   S1 :: "'x set"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1842
  and   S2 :: "'x set"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1843
  assumes  a: "S1 supports x"
18053
2719a6b7d95e some minor tweaks in some proofs (nothing extraordinary)
urbanc
parents: 18048
diff changeset
  1844
  and      b: "S1 \<subseteq> S2"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1845
  shows "S2 supports x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1846
  using a b
22808
a7daa74e2980 eliminated unnamed infixes, tuned syntax;
wenzelm
parents: 22786
diff changeset
  1847
  by (force simp add: supports_def)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1848
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1849
lemma supp_is_subset:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1850
  fixes S :: "'x set"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1851
  and   x :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1852
  assumes a1: "S supports x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1853
  and     a2: "finite S"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1854
  shows "(supp x)\<subseteq>S"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1855
proof (rule ccontr)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1856
  assume "\<not>(supp x \<subseteq> S)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1857
  hence "\<exists>a. a\<in>(supp x) \<and> a\<notin>S" by force
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1858
  then obtain a where b1: "a\<in>supp x" and b2: "a\<notin>S" by force
22808
a7daa74e2980 eliminated unnamed infixes, tuned syntax;
wenzelm
parents: 22786
diff changeset
  1859
  from a1 b2 have "\<forall>b. (b\<notin>S \<longrightarrow> ([(a,b)]\<bullet>x = x))" by (unfold supports_def, force)
19216
a45baf1ac094 tuned some of the proofs about fresh_fun
urbanc
parents: 19164
diff changeset
  1860
  hence "{b. [(a,b)]\<bullet>x \<noteq> x}\<subseteq>S" by force
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1861
  with a2 have "finite {b. [(a,b)]\<bullet>x \<noteq> x}" by (simp add: finite_subset)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1862
  hence "a\<notin>(supp x)" by (unfold supp_def, auto)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1863
  with b1 show False by simp
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1864
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1865
18264
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  1866
lemma supp_supports:
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  1867
  fixes x :: "'a"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  1868
  assumes  pt: "pt TYPE('a) TYPE('x)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  1869
  and      at: "at TYPE ('x)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  1870
  shows "((supp x)::'x set) supports x"
22808
a7daa74e2980 eliminated unnamed infixes, tuned syntax;
wenzelm
parents: 22786
diff changeset
  1871
proof (unfold supports_def, intro strip)
18264
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  1872
  fix a b
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  1873
  assume "(a::'x)\<notin>(supp x) \<and> (b::'x)\<notin>(supp x)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  1874
  hence "a\<sharp>x" and "b\<sharp>x" by (auto simp add: fresh_def)
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  1875
  thus "[(a,b)]\<bullet>x = x" by (rule pt_fresh_fresh[OF pt, OF at])
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  1876
qed
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  1877
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1878
lemma supports_finite:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1879
  fixes S :: "'x set"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1880
  and   x :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1881
  assumes a1: "S supports x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1882
  and     a2: "finite S"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1883
  shows "finite ((supp x)::'x set)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1884
proof -
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1885
  have "(supp x)\<subseteq>S" using a1 a2 by (rule supp_is_subset)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1886
  thus ?thesis using a2 by (simp add: finite_subset)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1887
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1888
  
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1889
lemma supp_is_inter:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1890
  fixes  x :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1891
  assumes  pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1892
  and      at: "at TYPE ('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1893
  and      fs: "fs TYPE('a) TYPE('x)"
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60580
diff changeset
  1894
  shows "((supp x)::'x set) = (\<Inter>{S. finite S \<and> S supports x})"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1895
proof (rule equalityI)
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60580
diff changeset
  1896
  show "((supp x)::'x set) \<subseteq> (\<Inter>{S. finite S \<and> S supports x})"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1897
  proof (clarify)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1898
    fix S c
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1899
    assume b: "c\<in>((supp x)::'x set)" and "finite (S::'x set)" and "S supports x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1900
    hence  "((supp x)::'x set)\<subseteq>S" by (simp add: supp_is_subset) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1901
    with b show "c\<in>S" by force
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1902
  qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1903
next
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60580
diff changeset
  1904
  show "(\<Inter>{S. finite S \<and> S supports x}) \<subseteq> ((supp x)::'x set)"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1905
  proof (clarify, simp)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1906
    fix c
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1907
    assume d: "\<forall>(S::'x set). finite S \<and> S supports x \<longrightarrow> c\<in>S"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1908
    have "((supp x)::'x set) supports x" by (rule supp_supports[OF pt, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1909
    with d fs1[OF fs] show "c\<in>supp x" by force
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1910
  qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1911
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1912
    
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1913
lemma supp_is_least_supports:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1914
  fixes S :: "'x set"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1915
  and   x :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1916
  assumes  pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1917
  and      at: "at TYPE ('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1918
  and      a1: "S supports x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1919
  and      a2: "finite S"
19477
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  1920
  and      a3: "\<forall>S'. (S' supports x) \<longrightarrow> S\<subseteq>S'"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1921
  shows "S = (supp x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1922
proof (rule equalityI)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1923
  show "((supp x)::'x set)\<subseteq>S" using a1 a2 by (rule supp_is_subset)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1924
next
19477
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  1925
  have "((supp x)::'x set) supports x" by (rule supp_supports[OF pt, OF at])
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  1926
  with a3 show "S\<subseteq>supp x" by force
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1927
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1928
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1929
lemma supports_set:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1930
  fixes S :: "'x set"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1931
  and   X :: "'a set"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1932
  assumes  pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1933
  and      at: "at TYPE ('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1934
  and      a: "\<forall>x\<in>X. (\<forall>(a::'x) (b::'x). a\<notin>S\<and>b\<notin>S \<longrightarrow> ([(a,b)]\<bullet>x)\<in>X)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1935
  shows  "S supports X"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1936
using a
22808
a7daa74e2980 eliminated unnamed infixes, tuned syntax;
wenzelm
parents: 22786
diff changeset
  1937
apply(auto simp add: supports_def)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1938
apply(simp add: pt_set_bij1a[OF pt, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1939
apply(force simp add: pt_swap_bij[OF pt, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1940
apply(simp add: pt_set_bij1a[OF pt, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1941
done
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1942
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1943
lemma supports_fresh:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1944
  fixes S :: "'x set"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1945
  and   a :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1946
  and   x :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1947
  assumes a1: "S supports x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1948
  and     a2: "finite S"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1949
  and     a3: "a\<notin>S"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1950
  shows "a\<sharp>x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1951
proof (simp add: fresh_def)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1952
  have "(supp x)\<subseteq>S" using a1 a2 by (rule supp_is_subset)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1953
  thus "a\<notin>(supp x)" using a3 by force
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1954
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1955
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1956
lemma at_fin_set_supports:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1957
  fixes X::"'x set"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1958
  assumes at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1959
  shows "X supports X"
19329
d6ddf304ec24 simplified the proof at_fin_set_supp
urbanc
parents: 19325
diff changeset
  1960
proof -
26806
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
  1961
  have "\<forall>a b. a\<notin>X \<and> b\<notin>X \<longrightarrow> [(a,b)]\<bullet>X = X"
46179
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  1962
    by (auto simp add: perm_set_def at_calc[OF at])
22808
a7daa74e2980 eliminated unnamed infixes, tuned syntax;
wenzelm
parents: 22786
diff changeset
  1963
  then show ?thesis by (simp add: supports_def)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1964
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1965
19329
d6ddf304ec24 simplified the proof at_fin_set_supp
urbanc
parents: 19325
diff changeset
  1966
lemma infinite_Collection:
d6ddf304ec24 simplified the proof at_fin_set_supp
urbanc
parents: 19325
diff changeset
  1967
  assumes a1:"infinite X"
d6ddf304ec24 simplified the proof at_fin_set_supp
urbanc
parents: 19325
diff changeset
  1968
  and     a2:"\<forall>b\<in>X. P(b)"
d6ddf304ec24 simplified the proof at_fin_set_supp
urbanc
parents: 19325
diff changeset
  1969
  shows "infinite {b\<in>X. P(b)}"
d6ddf304ec24 simplified the proof at_fin_set_supp
urbanc
parents: 19325
diff changeset
  1970
  using a1 a2 
d6ddf304ec24 simplified the proof at_fin_set_supp
urbanc
parents: 19325
diff changeset
  1971
  apply auto
d6ddf304ec24 simplified the proof at_fin_set_supp
urbanc
parents: 19325
diff changeset
  1972
  apply (subgoal_tac "infinite (X - {b\<in>X. P b})")
26806
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
  1973
  apply (simp add: set_diff_eq)
19329
d6ddf304ec24 simplified the proof at_fin_set_supp
urbanc
parents: 19325
diff changeset
  1974
  apply (simp add: Diff_infinite_finite)
d6ddf304ec24 simplified the proof at_fin_set_supp
urbanc
parents: 19325
diff changeset
  1975
  done
d6ddf304ec24 simplified the proof at_fin_set_supp
urbanc
parents: 19325
diff changeset
  1976
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1977
lemma at_fin_set_supp:
19329
d6ddf304ec24 simplified the proof at_fin_set_supp
urbanc
parents: 19325
diff changeset
  1978
  fixes X::"'x set" 
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1979
  assumes at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1980
  and     fs: "finite X"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1981
  shows "(supp X) = X"
19329
d6ddf304ec24 simplified the proof at_fin_set_supp
urbanc
parents: 19325
diff changeset
  1982
proof (rule subset_antisym)
d6ddf304ec24 simplified the proof at_fin_set_supp
urbanc
parents: 19325
diff changeset
  1983
  show "(supp X) \<subseteq> X" using at_fin_set_supports[OF at] using fs by (simp add: supp_is_subset)
d6ddf304ec24 simplified the proof at_fin_set_supp
urbanc
parents: 19325
diff changeset
  1984
next
d6ddf304ec24 simplified the proof at_fin_set_supp
urbanc
parents: 19325
diff changeset
  1985
  have inf: "infinite (UNIV-X)" using at4[OF at] fs by (auto simp add: Diff_infinite_finite)
d6ddf304ec24 simplified the proof at_fin_set_supp
urbanc
parents: 19325
diff changeset
  1986
  { fix a::"'x"
d6ddf304ec24 simplified the proof at_fin_set_supp
urbanc
parents: 19325
diff changeset
  1987
    assume asm: "a\<in>X"
26806
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26773
diff changeset
  1988
    hence "\<forall>b\<in>(UNIV-X). [(a,b)]\<bullet>X\<noteq>X"
46179
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  1989
      by (auto simp add: perm_set_def at_calc[OF at])
19329
d6ddf304ec24 simplified the proof at_fin_set_supp
urbanc
parents: 19325
diff changeset
  1990
    with inf have "infinite {b\<in>(UNIV-X). [(a,b)]\<bullet>X\<noteq>X}" by (rule infinite_Collection)
d6ddf304ec24 simplified the proof at_fin_set_supp
urbanc
parents: 19325
diff changeset
  1991
    hence "infinite {b. [(a,b)]\<bullet>X\<noteq>X}" by (rule_tac infinite_super, auto)
d6ddf304ec24 simplified the proof at_fin_set_supp
urbanc
parents: 19325
diff changeset
  1992
    hence "a\<in>(supp X)" by (simp add: supp_def)
d6ddf304ec24 simplified the proof at_fin_set_supp
urbanc
parents: 19325
diff changeset
  1993
  }
d6ddf304ec24 simplified the proof at_fin_set_supp
urbanc
parents: 19325
diff changeset
  1994
  then show "X\<subseteq>(supp X)" by blast
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1995
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1996
25950
a3067f6f08a2 Added lemma at_fin_set_fresh.
berghofe
parents: 24572
diff changeset
  1997
lemma at_fin_set_fresh:
a3067f6f08a2 Added lemma at_fin_set_fresh.
berghofe
parents: 24572
diff changeset
  1998
  fixes X::"'x set" 
a3067f6f08a2 Added lemma at_fin_set_fresh.
berghofe
parents: 24572
diff changeset
  1999
  assumes at: "at TYPE('x)"
a3067f6f08a2 Added lemma at_fin_set_fresh.
berghofe
parents: 24572
diff changeset
  2000
  and     fs: "finite X"
a3067f6f08a2 Added lemma at_fin_set_fresh.
berghofe
parents: 24572
diff changeset
  2001
  shows "(x \<sharp> X) = (x \<notin> X)"
a3067f6f08a2 Added lemma at_fin_set_fresh.
berghofe
parents: 24572
diff changeset
  2002
  by (simp add: at_fin_set_supp fresh_def at fs)
a3067f6f08a2 Added lemma at_fin_set_fresh.
berghofe
parents: 24572
diff changeset
  2003
30990
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2004
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  2005
section \<open>Permutations acting on Functions\<close>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2006
(*==========================================*)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2007
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2008
lemma pt_fun_app_eq:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2009
  fixes f  :: "'a\<Rightarrow>'b"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2010
  and   x  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2011
  and   pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2012
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2013
  and     at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2014
  shows "pi\<bullet>(f x) = (pi\<bullet>f)(pi\<bullet>x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2015
  by (simp add: perm_fun_def pt_rev_pi[OF pt, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2016
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2017
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  2018
\<comment>"sometimes pt_fun_app_eq does too much; this lemma 'corrects it'"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2019
lemma pt_perm:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2020
  fixes x  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2021
  and   pi1 :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2022
  and   pi2 :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2023
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2024
  and     at: "at TYPE ('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2025
  shows "(pi1\<bullet>perm pi2)(pi1\<bullet>x) = pi1\<bullet>(pi2\<bullet>x)" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2026
  by (simp add: pt_fun_app_eq[OF pt, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2027
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2028
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2029
lemma pt_fun_eq:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2030
  fixes f  :: "'a\<Rightarrow>'b"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2031
  and   pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2032
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2033
  and     at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2034
  shows "(pi\<bullet>f = f) = (\<forall> x. pi\<bullet>(f x) = f (pi\<bullet>x))" (is "?LHS = ?RHS")
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2035
proof
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2036
  assume a: "?LHS"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2037
  show "?RHS"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2038
  proof
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2039
    fix x
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2040
    have "pi\<bullet>(f x) = (pi\<bullet>f)(pi\<bullet>x)" by (simp add: pt_fun_app_eq[OF pt, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2041
    also have "\<dots> = f (pi\<bullet>x)" using a by simp
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2042
    finally show "pi\<bullet>(f x) = f (pi\<bullet>x)" by simp
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2043
  qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2044
next
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2045
  assume b: "?RHS"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2046
  show "?LHS"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2047
  proof (rule ccontr)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2048
    assume "(pi\<bullet>f) \<noteq> f"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  2049
    hence "\<exists>x. (pi\<bullet>f) x \<noteq> f x" by (simp add: fun_eq_iff)
19477
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2050
    then obtain x where b1: "(pi\<bullet>f) x \<noteq> f x" by force
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2051
    from b have "pi\<bullet>(f ((rev pi)\<bullet>x)) = f (pi\<bullet>((rev pi)\<bullet>x))" by force
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2052
    hence "(pi\<bullet>f)(pi\<bullet>((rev pi)\<bullet>x)) = f (pi\<bullet>((rev pi)\<bullet>x))" 
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2053
      by (simp add: pt_fun_app_eq[OF pt, OF at])
19477
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2054
    hence "(pi\<bullet>f) x = f x" by (simp add: pt_pi_rev[OF pt, OF at])
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2055
    with b1 show "False" by simp
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2056
  qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2057
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2058
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  2059
\<comment> "two helper lemmas for the equivariance of functions"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2060
lemma pt_swap_eq_aux:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2061
  fixes   y :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2062
  and    pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2063
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2064
  and     a: "\<forall>(a::'x) (b::'x). [(a,b)]\<bullet>y = y"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2065
  shows "pi\<bullet>y = y"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2066
proof(induct pi)
24544
da7de38392df trivial cleaning up
urbanc
parents: 23755
diff changeset
  2067
  case Nil show ?case by (simp add: pt1[OF pt])
da7de38392df trivial cleaning up
urbanc
parents: 23755
diff changeset
  2068
next
da7de38392df trivial cleaning up
urbanc
parents: 23755
diff changeset
  2069
  case (Cons x xs)
da7de38392df trivial cleaning up
urbanc
parents: 23755
diff changeset
  2070
  have ih: "xs\<bullet>y = y" by fact
da7de38392df trivial cleaning up
urbanc
parents: 23755
diff changeset
  2071
  obtain a b where p: "x=(a,b)" by force
da7de38392df trivial cleaning up
urbanc
parents: 23755
diff changeset
  2072
  have "((a,b)#xs)\<bullet>y = ([(a,b)]@xs)\<bullet>y" by simp
da7de38392df trivial cleaning up
urbanc
parents: 23755
diff changeset
  2073
  also have "\<dots> = [(a,b)]\<bullet>(xs\<bullet>y)" by (simp only: pt2[OF pt])
da7de38392df trivial cleaning up
urbanc
parents: 23755
diff changeset
  2074
  finally show ?case using a ih p by simp
da7de38392df trivial cleaning up
urbanc
parents: 23755
diff changeset
  2075
qed
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2076
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2077
lemma pt_swap_eq:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2078
  fixes   y :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2079
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2080
  shows "(\<forall>(a::'x) (b::'x). [(a,b)]\<bullet>y = y) = (\<forall>pi::'x prm. pi\<bullet>y = y)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2081
  by (force intro: pt_swap_eq_aux[OF pt])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2082
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2083
lemma pt_eqvt_fun1a:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2084
  fixes f     :: "'a\<Rightarrow>'b"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2085
  assumes pta: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2086
  and     ptb: "pt TYPE('b) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2087
  and     at:  "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2088
  and     a:   "((supp f)::'x set)={}"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2089
  shows "\<forall>(pi::'x prm). pi\<bullet>f = f" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2090
proof (intro strip)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2091
  fix pi
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2092
  have "\<forall>a b. a\<notin>((supp f)::'x set) \<and> b\<notin>((supp f)::'x set) \<longrightarrow> (([(a,b)]\<bullet>f) = f)" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2093
    by (intro strip, fold fresh_def, 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2094
      simp add: pt_fresh_fresh[OF pt_fun_inst[OF pta, OF ptb, OF at],OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2095
  with a have "\<forall>(a::'x) (b::'x). ([(a,b)]\<bullet>f) = f" by force
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2096
  hence "\<forall>(pi::'x prm). pi\<bullet>f = f" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2097
    by (simp add: pt_swap_eq[OF pt_fun_inst[OF pta, OF ptb, OF at]])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2098
  thus "(pi::'x prm)\<bullet>f = f" by simp
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2099
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2100
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2101
lemma pt_eqvt_fun1b:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2102
  fixes f     :: "'a\<Rightarrow>'b"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2103
  assumes a: "\<forall>(pi::'x prm). pi\<bullet>f = f"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2104
  shows "((supp f)::'x set)={}"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2105
using a by (simp add: supp_def)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2106
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2107
lemma pt_eqvt_fun1:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2108
  fixes f     :: "'a\<Rightarrow>'b"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2109
  assumes pta: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2110
  and     ptb: "pt TYPE('b) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2111
  and     at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2112
  shows "(((supp f)::'x set)={}) = (\<forall>(pi::'x prm). pi\<bullet>f = f)" (is "?LHS = ?RHS")
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2113
by (rule iffI, simp add: pt_eqvt_fun1a[OF pta, OF ptb, OF at], simp add: pt_eqvt_fun1b)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2114
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2115
lemma pt_eqvt_fun2a:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2116
  fixes f     :: "'a\<Rightarrow>'b"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2117
  assumes pta: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2118
  and     ptb: "pt TYPE('b) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2119
  and     at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2120
  assumes a: "((supp f)::'x set)={}"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2121
  shows "\<forall>(pi::'x prm) (x::'a). pi\<bullet>(f x) = f(pi\<bullet>x)" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2122
proof (intro strip)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2123
  fix pi x
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2124
  from a have b: "\<forall>(pi::'x prm). pi\<bullet>f = f" by (simp add: pt_eqvt_fun1[OF pta, OF ptb, OF at]) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2125
  have "(pi::'x prm)\<bullet>(f x) = (pi\<bullet>f)(pi\<bullet>x)" by (simp add: pt_fun_app_eq[OF pta, OF at]) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2126
  with b show "(pi::'x prm)\<bullet>(f x) = f (pi\<bullet>x)" by force 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2127
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2128
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2129
lemma pt_eqvt_fun2b:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2130
  fixes f     :: "'a\<Rightarrow>'b"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2131
  assumes pt1: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2132
  and     pt2: "pt TYPE('b) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2133
  and     at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2134
  assumes a: "\<forall>(pi::'x prm) (x::'a). pi\<bullet>(f x) = f(pi\<bullet>x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2135
  shows "((supp f)::'x set)={}"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2136
proof -
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2137
  from a have "\<forall>(pi::'x prm). pi\<bullet>f = f" by (simp add: pt_fun_eq[OF pt1, OF at, symmetric])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2138
  thus ?thesis by (simp add: supp_def)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2139
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2140
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2141
lemma pt_eqvt_fun2:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2142
  fixes f     :: "'a\<Rightarrow>'b"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2143
  assumes pta: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2144
  and     ptb: "pt TYPE('b) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2145
  and     at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2146
  shows "(((supp f)::'x set)={}) = (\<forall>(pi::'x prm) (x::'a). pi\<bullet>(f x) = f(pi\<bullet>x))" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2147
by (rule iffI, 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2148
    simp add: pt_eqvt_fun2a[OF pta, OF ptb, OF at], 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2149
    simp add: pt_eqvt_fun2b[OF pta, OF ptb, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2150
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2151
lemma pt_supp_fun_subset:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2152
  fixes f :: "'a\<Rightarrow>'b"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2153
  assumes pta: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2154
  and     ptb: "pt TYPE('b) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2155
  and     at: "at TYPE('x)" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2156
  and     f1: "finite ((supp f)::'x set)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2157
  and     f2: "finite ((supp x)::'x set)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2158
  shows "supp (f x) \<subseteq> (((supp f)\<union>(supp x))::'x set)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2159
proof -
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2160
  have s1: "((supp f)\<union>((supp x)::'x set)) supports (f x)"
22808
a7daa74e2980 eliminated unnamed infixes, tuned syntax;
wenzelm
parents: 22786
diff changeset
  2161
  proof (simp add: supports_def, fold fresh_def, auto)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2162
    fix a::"'x" and b::"'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2163
    assume "a\<sharp>f" and "b\<sharp>f"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2164
    hence a1: "[(a,b)]\<bullet>f = f" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2165
      by (rule pt_fresh_fresh[OF pt_fun_inst[OF pta, OF ptb, OF at], OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2166
    assume "a\<sharp>x" and "b\<sharp>x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2167
    hence a2: "[(a,b)]\<bullet>x = x" by (rule pt_fresh_fresh[OF pta, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2168
    from a1 a2 show "[(a,b)]\<bullet>(f x) = (f x)" by (simp add: pt_fun_app_eq[OF pta, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2169
  qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2170
  from f1 f2 have "finite ((supp f)\<union>((supp x)::'x set))" by force
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2171
  with s1 show ?thesis by (rule supp_is_subset)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2172
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2173
      
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2174
lemma pt_empty_supp_fun_subset:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2175
  fixes f :: "'a\<Rightarrow>'b"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2176
  assumes pta: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2177
  and     ptb: "pt TYPE('b) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2178
  and     at:  "at TYPE('x)" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2179
  and     e:   "(supp f)=({}::'x set)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2180
  shows "supp (f x) \<subseteq> ((supp x)::'x set)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2181
proof (unfold supp_def, auto)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2182
  fix a::"'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2183
  assume a1: "finite {b. [(a, b)]\<bullet>x \<noteq> x}"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2184
  assume "infinite {b. [(a, b)]\<bullet>(f x) \<noteq> f x}"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2185
  hence a2: "infinite {b. f ([(a, b)]\<bullet>x) \<noteq> f x}" using e
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2186
    by (simp add: pt_eqvt_fun2[OF pta, OF ptb, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2187
  have a3: "{b. f ([(a,b)]\<bullet>x) \<noteq> f x}\<subseteq>{b. [(a,b)]\<bullet>x \<noteq> x}" by force
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2188
  from a1 a2 a3 show False by (force dest: finite_subset)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2189
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2190
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  2191
section \<open>Facts about the support of finite sets of finitely supported things\<close>
18264
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2192
(*=============================================================================*)
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2193
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
  2194
definition X_to_Un_supp :: "('a set) \<Rightarrow> 'x set" where
18264
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2195
  "X_to_Un_supp X \<equiv> \<Union>x\<in>X. ((supp x)::'x set)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2196
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2197
lemma UNION_f_eqvt:
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2198
  fixes X::"('a set)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2199
  and   f::"'a \<Rightarrow> 'x set"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2200
  and   pi::"'x prm"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2201
  assumes pt: "pt TYPE('a) TYPE('x)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2202
  and     at: "at TYPE('x)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2203
  shows "pi\<bullet>(\<Union>x\<in>X. f x) = (\<Union>x\<in>(pi\<bullet>X). (pi\<bullet>f) x)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2204
proof -
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2205
  have pt_x: "pt TYPE('x) TYPE('x)" by (force intro: at_pt_inst at)
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2206
  show ?thesis
18351
6bab9cef50cf ISAR-fied two proofs
urbanc
parents: 18295
diff changeset
  2207
  proof (rule equalityI)
6bab9cef50cf ISAR-fied two proofs
urbanc
parents: 18295
diff changeset
  2208
    show "pi\<bullet>(\<Union>x\<in>X. f x) \<subseteq> (\<Union>x\<in>(pi\<bullet>X). (pi\<bullet>f) x)"
46179
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  2209
      apply(auto simp add: perm_set_def)
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22808
diff changeset
  2210
      apply(rule_tac x="pi\<bullet>xb" in exI)
18351
6bab9cef50cf ISAR-fied two proofs
urbanc
parents: 18295
diff changeset
  2211
      apply(rule conjI)
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22808
diff changeset
  2212
      apply(rule_tac x="xb" in exI)
18351
6bab9cef50cf ISAR-fied two proofs
urbanc
parents: 18295
diff changeset
  2213
      apply(simp)
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22808
diff changeset
  2214
      apply(subgoal_tac "(pi\<bullet>f) (pi\<bullet>xb) = pi\<bullet>(f xb)")(*A*)
18351
6bab9cef50cf ISAR-fied two proofs
urbanc
parents: 18295
diff changeset
  2215
      apply(simp)
6bab9cef50cf ISAR-fied two proofs
urbanc
parents: 18295
diff changeset
  2216
      apply(rule pt_set_bij2[OF pt_x, OF at])
6bab9cef50cf ISAR-fied two proofs
urbanc
parents: 18295
diff changeset
  2217
      apply(assumption)
6bab9cef50cf ISAR-fied two proofs
urbanc
parents: 18295
diff changeset
  2218
      (*A*)
6bab9cef50cf ISAR-fied two proofs
urbanc
parents: 18295
diff changeset
  2219
      apply(rule sym)
6bab9cef50cf ISAR-fied two proofs
urbanc
parents: 18295
diff changeset
  2220
      apply(rule pt_fun_app_eq[OF pt, OF at])
6bab9cef50cf ISAR-fied two proofs
urbanc
parents: 18295
diff changeset
  2221
      done
6bab9cef50cf ISAR-fied two proofs
urbanc
parents: 18295
diff changeset
  2222
  next
6bab9cef50cf ISAR-fied two proofs
urbanc
parents: 18295
diff changeset
  2223
    show "(\<Union>x\<in>(pi\<bullet>X). (pi\<bullet>f) x) \<subseteq> pi\<bullet>(\<Union>x\<in>X. f x)"
46179
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  2224
      apply(auto simp add: perm_set_def)
18351
6bab9cef50cf ISAR-fied two proofs
urbanc
parents: 18295
diff changeset
  2225
      apply(rule_tac x="(rev pi)\<bullet>x" in exI)
6bab9cef50cf ISAR-fied two proofs
urbanc
parents: 18295
diff changeset
  2226
      apply(rule conjI)
6bab9cef50cf ISAR-fied two proofs
urbanc
parents: 18295
diff changeset
  2227
      apply(simp add: pt_pi_rev[OF pt_x, OF at])
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22808
diff changeset
  2228
      apply(rule_tac x="xb" in bexI)
18351
6bab9cef50cf ISAR-fied two proofs
urbanc
parents: 18295
diff changeset
  2229
      apply(simp add: pt_set_bij1[OF pt_x, OF at])
6bab9cef50cf ISAR-fied two proofs
urbanc
parents: 18295
diff changeset
  2230
      apply(simp add: pt_fun_app_eq[OF pt, OF at])
6bab9cef50cf ISAR-fied two proofs
urbanc
parents: 18295
diff changeset
  2231
      apply(assumption)
6bab9cef50cf ISAR-fied two proofs
urbanc
parents: 18295
diff changeset
  2232
      done
6bab9cef50cf ISAR-fied two proofs
urbanc
parents: 18295
diff changeset
  2233
  qed
18264
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2234
qed
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2235
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2236
lemma X_to_Un_supp_eqvt:
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2237
  fixes X::"('a set)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2238
  and   pi::"'x prm"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2239
  assumes pt: "pt TYPE('a) TYPE('x)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2240
  and     at: "at TYPE('x)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2241
  shows "pi\<bullet>(X_to_Un_supp X) = ((X_to_Un_supp (pi\<bullet>X))::'x set)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2242
  apply(simp add: X_to_Un_supp_def)
45961
5cefe17916a6 treatment of type constructor `set`
haftmann
parents: 45694
diff changeset
  2243
  apply(simp add: UNION_f_eqvt[OF pt, OF at] perm_fun_def)
18264
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2244
  apply(simp add: pt_perm_supp[OF pt, OF at])
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2245
  apply(simp add: pt_pi_rev[OF pt, OF at])
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2246
  done
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2247
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2248
lemma Union_supports_set:
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2249
  fixes X::"('a set)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2250
  assumes pt: "pt TYPE('a) TYPE('x)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2251
  and     at: "at TYPE('x)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2252
  shows "(\<Union>x\<in>X. ((supp x)::'x set)) supports X"
22808
a7daa74e2980 eliminated unnamed infixes, tuned syntax;
wenzelm
parents: 22786
diff changeset
  2253
  apply(simp add: supports_def fresh_def[symmetric])
18264
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2254
  apply(rule allI)+
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2255
  apply(rule impI)
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2256
  apply(erule conjE)
46179
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  2257
  apply(simp add: perm_set_def)
18264
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2258
  apply(auto)
22829
f1db55c7534d tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents: 22808
diff changeset
  2259
  apply(subgoal_tac "[(a,b)]\<bullet>xa = xa")(*A*)
18264
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2260
  apply(simp)
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2261
  apply(rule pt_fresh_fresh[OF pt, OF at])
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2262
  apply(force)
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2263
  apply(force)
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2264
  apply(rule_tac x="x" in exI)
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2265
  apply(simp)
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2266
  apply(rule sym)
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2267
  apply(rule pt_fresh_fresh[OF pt, OF at])
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2268
  apply(force)+
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2269
  done
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2270
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2271
lemma Union_of_fin_supp_sets:
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2272
  fixes X::"('a set)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2273
  assumes fs: "fs TYPE('a) TYPE('x)" 
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2274
  and     fi: "finite X"   
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2275
  shows "finite (\<Union>x\<in>X. ((supp x)::'x set))"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2276
using fi by (induct, auto simp add: fs1[OF fs])
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2277
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2278
lemma Union_included_in_supp:
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2279
  fixes X::"('a set)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2280
  assumes pt: "pt TYPE('a) TYPE('x)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2281
  and     at: "at TYPE('x)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2282
  and     fs: "fs TYPE('a) TYPE('x)" 
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2283
  and     fi: "finite X"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2284
  shows "(\<Union>x\<in>X. ((supp x)::'x set)) \<subseteq> supp X"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2285
proof -
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2286
  have "supp ((X_to_Un_supp X)::'x set) \<subseteq> ((supp X)::'x set)"  
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2287
    apply(rule pt_empty_supp_fun_subset)
45961
5cefe17916a6 treatment of type constructor `set`
haftmann
parents: 45694
diff changeset
  2288
    apply(force intro: pt_set_inst at_pt_inst pt at)+
18264
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2289
    apply(rule pt_eqvt_fun2b)
45961
5cefe17916a6 treatment of type constructor `set`
haftmann
parents: 45694
diff changeset
  2290
    apply(force intro: pt_set_inst at_pt_inst pt at)+
18351
6bab9cef50cf ISAR-fied two proofs
urbanc
parents: 18295
diff changeset
  2291
    apply(rule allI)+
18264
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2292
    apply(rule X_to_Un_supp_eqvt[OF pt, OF at])
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2293
    done
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2294
  hence "supp (\<Union>x\<in>X. ((supp x)::'x set)) \<subseteq> ((supp X)::'x set)" by (simp add: X_to_Un_supp_def)
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2295
  moreover
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2296
  have "supp (\<Union>x\<in>X. ((supp x)::'x set)) = (\<Union>x\<in>X. ((supp x)::'x set))"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2297
    apply(rule at_fin_set_supp[OF at])
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2298
    apply(rule Union_of_fin_supp_sets[OF fs, OF fi])
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2299
    done
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2300
  ultimately show ?thesis by force
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2301
qed
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2302
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2303
lemma supp_of_fin_sets:
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2304
  fixes X::"('a set)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2305
  assumes pt: "pt TYPE('a) TYPE('x)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2306
  and     at: "at TYPE('x)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2307
  and     fs: "fs TYPE('a) TYPE('x)" 
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2308
  and     fi: "finite X"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2309
  shows "(supp X) = (\<Union>x\<in>X. ((supp x)::'x set))"
18351
6bab9cef50cf ISAR-fied two proofs
urbanc
parents: 18295
diff changeset
  2310
apply(rule equalityI)
18264
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2311
apply(rule supp_is_subset)
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2312
apply(rule Union_supports_set[OF pt, OF at])
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2313
apply(rule Union_of_fin_supp_sets[OF fs, OF fi])
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2314
apply(rule Union_included_in_supp[OF pt, OF at, OF fs, OF fi])
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2315
done
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2316
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2317
lemma supp_fin_union:
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2318
  fixes X::"('a set)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2319
  and   Y::"('a set)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2320
  assumes pt: "pt TYPE('a) TYPE('x)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2321
  and     at: "at TYPE('x)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2322
  and     fs: "fs TYPE('a) TYPE('x)" 
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2323
  and     f1: "finite X"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2324
  and     f2: "finite Y"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2325
  shows "(supp (X\<union>Y)) = (supp X)\<union>((supp Y)::'x set)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2326
using f1 f2 by (force simp add: supp_of_fin_sets[OF pt, OF at, OF fs])
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2327
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2328
lemma supp_fin_insert:
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2329
  fixes X::"('a set)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2330
  and   x::"'a"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2331
  assumes pt: "pt TYPE('a) TYPE('x)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2332
  and     at: "at TYPE('x)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2333
  and     fs: "fs TYPE('a) TYPE('x)" 
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2334
  and     f:  "finite X"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2335
  shows "(supp (insert x X)) = (supp x)\<union>((supp X)::'x set)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2336
proof -
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2337
  have "(supp (insert x X)) = ((supp ({x}\<union>(X::'a set)))::'x set)" by simp
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2338
  also have "\<dots> = (supp {x})\<union>(supp X)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2339
    by (rule supp_fin_union[OF pt, OF at, OF fs], simp_all add: f)
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2340
  finally show "(supp (insert x X)) = (supp x)\<union>((supp X)::'x set)" 
46179
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  2341
    by (simp add: supp_singleton)
18264
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2342
qed
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2343
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2344
lemma fresh_fin_union:
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2345
  fixes X::"('a set)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2346
  and   Y::"('a set)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2347
  and   a::"'x"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2348
  assumes pt: "pt TYPE('a) TYPE('x)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2349
  and     at: "at TYPE('x)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2350
  and     fs: "fs TYPE('a) TYPE('x)" 
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2351
  and     f1: "finite X"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2352
  and     f2: "finite Y"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2353
  shows "a\<sharp>(X\<union>Y) = (a\<sharp>X \<and> a\<sharp>Y)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2354
apply(simp add: fresh_def)
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2355
apply(simp add: supp_fin_union[OF pt, OF at, OF fs, OF f1, OF f2])
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2356
done
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2357
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2358
lemma fresh_fin_insert:
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2359
  fixes X::"('a set)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2360
  and   x::"'a"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2361
  and   a::"'x"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2362
  assumes pt: "pt TYPE('a) TYPE('x)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2363
  and     at: "at TYPE('x)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2364
  and     fs: "fs TYPE('a) TYPE('x)" 
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2365
  and     f:  "finite X"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2366
  shows "a\<sharp>(insert x X) = (a\<sharp>x \<and> a\<sharp>X)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2367
apply(simp add: fresh_def)
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2368
apply(simp add: supp_fin_insert[OF pt, OF at, OF fs, OF f])
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2369
done
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2370
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2371
lemma fresh_fin_insert1:
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2372
  fixes X::"('a set)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2373
  and   x::"'a"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2374
  and   a::"'x"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2375
  assumes pt: "pt TYPE('a) TYPE('x)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2376
  and     at: "at TYPE('x)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2377
  and     fs: "fs TYPE('a) TYPE('x)" 
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2378
  and     f:  "finite X"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2379
  and     a1:  "a\<sharp>x"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2380
  and     a2:  "a\<sharp>X"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2381
  shows "a\<sharp>(insert x X)"
30990
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2382
  using a1 a2
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2383
  by (simp add: fresh_fin_insert[OF pt, OF at, OF fs, OF f])
18264
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2384
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2385
lemma pt_list_set_supp:
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2386
  fixes xs :: "'a list"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2387
  assumes pt: "pt TYPE('a) TYPE('x)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2388
  and     at: "at TYPE('x)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2389
  and     fs: "fs TYPE('a) TYPE('x)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2390
  shows "supp (set xs) = ((supp xs)::'x set)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2391
proof -
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2392
  have "supp (set xs) = (\<Union>x\<in>(set xs). ((supp x)::'x set))"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2393
    by (rule supp_of_fin_sets[OF pt, OF at, OF fs], rule finite_set)
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2394
  also have "(\<Union>x\<in>(set xs). ((supp x)::'x set)) = (supp xs)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2395
  proof(induct xs)
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2396
    case Nil show ?case by (simp add: supp_list_nil)
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2397
  next
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2398
    case (Cons h t) thus ?case by (simp add: supp_list_cons)
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2399
  qed
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2400
  finally show ?thesis by simp
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2401
qed
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2402
    
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2403
lemma pt_list_set_fresh:
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2404
  fixes a :: "'x"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2405
  and   xs :: "'a list"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2406
  assumes pt: "pt TYPE('a) TYPE('x)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2407
  and     at: "at TYPE('x)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2408
  and     fs: "fs TYPE('a) TYPE('x)"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2409
  shows "a\<sharp>(set xs) = a\<sharp>xs"
3b808e24667b added the version of nominal.thy that contains
urbanc
parents: 18246
diff changeset
  2410
by (simp add: fresh_def pt_list_set_supp[OF pt, OF at, OF fs])
26847
9254cca608ef added at_set_avoiding lemmas
urbanc
parents: 26806
diff changeset
  2411
30990
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2412
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  2413
section \<open>generalisation of freshness to lists and sets of atoms\<close>
30990
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2414
(*================================================================*)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2415
 
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2416
consts
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2417
  fresh_star :: "'b \<Rightarrow> 'a \<Rightarrow> bool" ("_ \<sharp>* _" [100,100] 100)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2418
62145
5b946c81dfbf eliminated old defs;
wenzelm
parents: 61260
diff changeset
  2419
overloading fresh_star_set \<equiv> "fresh_star :: 'b set \<Rightarrow> 'a \<Rightarrow> bool"
5b946c81dfbf eliminated old defs;
wenzelm
parents: 61260
diff changeset
  2420
begin
5b946c81dfbf eliminated old defs;
wenzelm
parents: 61260
diff changeset
  2421
  definition fresh_star_set: "xs\<sharp>*c \<equiv> \<forall>x::'b\<in>xs. x\<sharp>(c::'a)"
5b946c81dfbf eliminated old defs;
wenzelm
parents: 61260
diff changeset
  2422
end
5b946c81dfbf eliminated old defs;
wenzelm
parents: 61260
diff changeset
  2423
5b946c81dfbf eliminated old defs;
wenzelm
parents: 61260
diff changeset
  2424
overloading frsh_star_list \<equiv> "fresh_star :: 'b list \<Rightarrow> 'a \<Rightarrow> bool"
5b946c81dfbf eliminated old defs;
wenzelm
parents: 61260
diff changeset
  2425
begin
5b946c81dfbf eliminated old defs;
wenzelm
parents: 61260
diff changeset
  2426
  definition fresh_star_list: "xs\<sharp>*c \<equiv> \<forall>x::'b\<in>set xs. x\<sharp>(c::'a)"
5b946c81dfbf eliminated old defs;
wenzelm
parents: 61260
diff changeset
  2427
end
30990
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2428
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2429
lemmas fresh_star_def = fresh_star_list fresh_star_set
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2430
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2431
lemma fresh_star_prod_set:
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2432
  fixes xs::"'a set"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2433
  shows "xs\<sharp>*(a,b) = (xs\<sharp>*a \<and> xs\<sharp>*b)"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2434
by (auto simp add: fresh_star_def fresh_prod)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2435
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2436
lemma fresh_star_prod_list:
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2437
  fixes xs::"'a list"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2438
  shows "xs\<sharp>*(a,b) = (xs\<sharp>*a \<and> xs\<sharp>*b)"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2439
  by (auto simp add: fresh_star_def fresh_prod)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2440
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2441
lemmas fresh_star_prod = fresh_star_prod_list fresh_star_prod_set
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2442
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2443
lemma fresh_star_set_eq: "set xs \<sharp>* c = xs \<sharp>* c"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2444
  by (simp add: fresh_star_def)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2445
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2446
lemma fresh_star_Un_elim:
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2447
  "((S \<union> T) \<sharp>* c \<Longrightarrow> PROP C) \<equiv> (S \<sharp>* c \<Longrightarrow> T \<sharp>* c \<Longrightarrow> PROP C)"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2448
  apply rule
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2449
  apply (simp_all add: fresh_star_def)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2450
  apply (erule meta_mp)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2451
  apply blast
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2452
  done
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2453
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2454
lemma fresh_star_insert_elim:
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2455
  "(insert x S \<sharp>* c \<Longrightarrow> PROP C) \<equiv> (x \<sharp> c \<Longrightarrow> S \<sharp>* c \<Longrightarrow> PROP C)"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2456
  by rule (simp_all add: fresh_star_def)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2457
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2458
lemma fresh_star_empty_elim:
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2459
  "({} \<sharp>* c \<Longrightarrow> PROP C) \<equiv> PROP C"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2460
  by (simp add: fresh_star_def)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2461
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  2462
text \<open>Normalization of freshness results; see \ \<open>nominal_induct\<close>\<close>
30990
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2463
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2464
lemma fresh_star_unit_elim: 
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2465
  shows "((a::'a set)\<sharp>*() \<Longrightarrow> PROP C) \<equiv> PROP C"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2466
  and "((b::'a list)\<sharp>*() \<Longrightarrow> PROP C) \<equiv> PROP C"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2467
  by (simp_all add: fresh_star_def fresh_def supp_unit)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2468
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2469
lemma fresh_star_prod_elim: 
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2470
  shows "((a::'a set)\<sharp>*(x,y) \<Longrightarrow> PROP C) \<equiv> (a\<sharp>*x \<Longrightarrow> a\<sharp>*y \<Longrightarrow> PROP C)"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2471
  and "((b::'a list)\<sharp>*(x,y) \<Longrightarrow> PROP C) \<equiv> (b\<sharp>*x \<Longrightarrow> b\<sharp>*y \<Longrightarrow> PROP C)"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2472
  by (rule, simp_all add: fresh_star_prod)+
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2473
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2474
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2475
lemma pt_fresh_star_bij_ineq:
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2476
  fixes  pi :: "'x prm"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2477
  and     x :: "'a"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2478
  and     a :: "'y set"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2479
  and     b :: "'y list"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2480
  assumes pta: "pt TYPE('a) TYPE('x)"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2481
  and     ptb: "pt TYPE('y) TYPE('x)"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2482
  and     at:  "at TYPE('x)"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2483
  and     cp:  "cp TYPE('a) TYPE('x) TYPE('y)"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2484
  shows "(pi\<bullet>a)\<sharp>*(pi\<bullet>x) = a\<sharp>*x"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2485
  and   "(pi\<bullet>b)\<sharp>*(pi\<bullet>x) = b\<sharp>*x"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2486
apply(unfold fresh_star_def)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2487
apply(auto)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2488
apply(drule_tac x="pi\<bullet>xa" in bspec)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2489
apply(erule pt_set_bij2[OF ptb, OF at])
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2490
apply(simp add: fresh_star_def pt_fresh_bij_ineq[OF pta, OF ptb, OF at, OF cp])
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2491
apply(drule_tac x="(rev pi)\<bullet>xa" in bspec)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2492
apply(simp add: pt_set_bij1[OF ptb, OF at])
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2493
apply(simp add: pt_fresh_left_ineq[OF pta, OF ptb, OF at, OF cp])
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2494
apply(drule_tac x="pi\<bullet>xa" in bspec)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2495
apply(simp add: pt_set_bij1[OF ptb, OF at])
46179
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  2496
apply(simp add: set_eqvt pt_rev_pi[OF pt_list_inst[OF ptb], OF at])
30990
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2497
apply(simp add: pt_fresh_bij_ineq[OF pta, OF ptb, OF at, OF cp])
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2498
apply(drule_tac x="(rev pi)\<bullet>xa" in bspec)
46179
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  2499
apply(simp add: pt_set_bij1[OF ptb, OF at] set_eqvt)
30990
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2500
apply(simp add: pt_fresh_left_ineq[OF pta, OF ptb, OF at, OF cp])
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2501
done
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2502
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2503
lemma pt_fresh_star_bij:
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2504
  fixes  pi :: "'x prm"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2505
  and     x :: "'a"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2506
  and     a :: "'x set"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2507
  and     b :: "'x list"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2508
  assumes pt: "pt TYPE('a) TYPE('x)"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2509
  and     at: "at TYPE('x)"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2510
  shows "(pi\<bullet>a)\<sharp>*(pi\<bullet>x) = a\<sharp>*x"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2511
  and   "(pi\<bullet>b)\<sharp>*(pi\<bullet>x) = b\<sharp>*x"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2512
apply(rule pt_fresh_star_bij_ineq(1))
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2513
apply(rule pt)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2514
apply(rule at_pt_inst)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2515
apply(rule at)+
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2516
apply(rule cp_pt_inst)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2517
apply(rule pt)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2518
apply(rule at)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2519
apply(rule pt_fresh_star_bij_ineq(2))
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2520
apply(rule pt)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2521
apply(rule at_pt_inst)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2522
apply(rule at)+
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2523
apply(rule cp_pt_inst)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2524
apply(rule pt)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2525
apply(rule at)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2526
done
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2527
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2528
lemma pt_fresh_star_eqvt:
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2529
  fixes  pi :: "'x prm"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2530
  and     x :: "'a"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2531
  and     a :: "'x set"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2532
  and     b :: "'x list"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2533
  assumes pt: "pt TYPE('a) TYPE('x)"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2534
  and     at: "at TYPE('x)"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2535
  shows "pi\<bullet>(a\<sharp>*x) = (pi\<bullet>a)\<sharp>*(pi\<bullet>x)"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2536
  and   "pi\<bullet>(b\<sharp>*x) = (pi\<bullet>b)\<sharp>*(pi\<bullet>x)"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2537
  by (simp_all add: perm_bool pt_fresh_star_bij[OF pt, OF at])
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2538
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2539
lemma pt_fresh_star_eqvt_ineq:
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2540
  fixes pi::"'x prm"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2541
  and   a::"'y set"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2542
  and   b::"'y list"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2543
  and   x::"'a"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2544
  assumes pta: "pt TYPE('a) TYPE('x)"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2545
  and     ptb: "pt TYPE('y) TYPE('x)"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2546
  and     at:  "at TYPE('x)"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2547
  and     cp:  "cp TYPE('a) TYPE('x) TYPE('y)"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2548
  and     dj:  "disjoint TYPE('y) TYPE('x)"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2549
  shows "pi\<bullet>(a\<sharp>*x) = (pi\<bullet>a)\<sharp>*(pi\<bullet>x)"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2550
  and   "pi\<bullet>(b\<sharp>*x) = (pi\<bullet>b)\<sharp>*(pi\<bullet>x)"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2551
  by (simp_all add: pt_fresh_star_bij_ineq[OF pta, OF ptb, OF at, OF cp] dj_perm_forget[OF dj] perm_bool)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2552
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2553
lemma pt_freshs_freshs:
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2554
  assumes pt: "pt TYPE('a) TYPE('x)"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2555
  and at: "at TYPE ('x)"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2556
  and pi: "set (pi::'x prm) \<subseteq> Xs \<times> Ys"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2557
  and Xs: "Xs \<sharp>* (x::'a)"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2558
  and Ys: "Ys \<sharp>* x"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2559
  shows "pi\<bullet>x = x"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2560
  using pi
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2561
proof (induct pi)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2562
  case Nil
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2563
  show ?case by (simp add: pt1 [OF pt])
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2564
next
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2565
  case (Cons p pi)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2566
  obtain a b where p: "p = (a, b)" by (cases p)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2567
  with Cons Xs Ys have "a \<sharp> x" "b \<sharp> x"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2568
    by (simp_all add: fresh_star_def)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2569
  with Cons p show ?case
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2570
    by (simp add: pt_fresh_fresh [OF pt at]
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2571
      pt2 [OF pt, of "[(a, b)]" pi, simplified])
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2572
qed
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2573
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2574
lemma pt_fresh_star_pi: 
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2575
  fixes x::"'a"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2576
  and   pi::"'x prm"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2577
  assumes pt: "pt TYPE('a) TYPE('x)"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2578
  and     at: "at TYPE('x)"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2579
  and     a: "((supp x)::'x set)\<sharp>* pi"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2580
  shows "pi\<bullet>x = x"
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2581
using a
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2582
apply(induct pi)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2583
apply(auto simp add: fresh_star_def fresh_list_cons fresh_prod pt1[OF pt])
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2584
apply(subgoal_tac "((a,b)#pi)\<bullet>x = ([(a,b)]@pi)\<bullet>x")
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2585
apply(simp only: pt2[OF pt])
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2586
apply(rule pt_fresh_fresh[OF pt at])
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2587
apply(simp add: fresh_def at_supp[OF at])
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2588
apply(blast)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2589
apply(simp add: fresh_def at_supp[OF at])
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2590
apply(blast)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2591
apply(simp add: pt2[OF pt])
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2592
done
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2593
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  2594
section \<open>Infrastructure lemmas for strong rule inductions\<close>
26847
9254cca608ef added at_set_avoiding lemmas
urbanc
parents: 26806
diff changeset
  2595
(*==========================================================*)
9254cca608ef added at_set_avoiding lemmas
urbanc
parents: 26806
diff changeset
  2596
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  2597
text \<open>
26847
9254cca608ef added at_set_avoiding lemmas
urbanc
parents: 26806
diff changeset
  2598
  For every set of atoms, there is another set of atoms
9254cca608ef added at_set_avoiding lemmas
urbanc
parents: 26806
diff changeset
  2599
  avoiding a finitely supported c and there is a permutation
30990
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  2600
  which 'translates' between both sets.
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  2601
\<close>
32638
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2602
26847
9254cca608ef added at_set_avoiding lemmas
urbanc
parents: 26806
diff changeset
  2603
lemma at_set_avoiding_aux:
9254cca608ef added at_set_avoiding lemmas
urbanc
parents: 26806
diff changeset
  2604
  fixes Xs::"'a set"
9254cca608ef added at_set_avoiding lemmas
urbanc
parents: 26806
diff changeset
  2605
  and   As::"'a set"
9254cca608ef added at_set_avoiding lemmas
urbanc
parents: 26806
diff changeset
  2606
  assumes at: "at TYPE('a)"
9254cca608ef added at_set_avoiding lemmas
urbanc
parents: 26806
diff changeset
  2607
  and     b: "Xs \<subseteq> As"
9254cca608ef added at_set_avoiding lemmas
urbanc
parents: 26806
diff changeset
  2608
  and     c: "finite As"
9254cca608ef added at_set_avoiding lemmas
urbanc
parents: 26806
diff changeset
  2609
  and     d: "finite ((supp c)::'a set)"
32638
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2610
  shows "\<exists>(pi::'a prm). (pi\<bullet>Xs)\<sharp>*c \<and> (pi\<bullet>Xs) \<inter> As = {} \<and> set pi \<subseteq> Xs \<times> (pi\<bullet>Xs)"
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2611
proof -
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2612
  from b c have "finite Xs" by (simp add: finite_subset)
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2613
  then show ?thesis using b 
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2614
  proof (induct)
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2615
    case empty
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2616
    have "({}::'a set)\<sharp>*c" by (simp add: fresh_star_def)
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2617
    moreover
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2618
    have "({}::'a set) \<inter> As = {}" by simp
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2619
    moreover
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2620
    have "set ([]::'a prm) \<subseteq> {} \<times> {}" by simp
46179
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  2621
    ultimately show ?case by (simp add: empty_eqvt)
32638
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2622
  next
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2623
    case (insert x Xs)
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2624
    then have ih: "\<exists>pi. (pi\<bullet>Xs)\<sharp>*c \<and> (pi\<bullet>Xs) \<inter> As = {} \<and> set pi \<subseteq> Xs \<times> (pi\<bullet>Xs)" by simp
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2625
    then obtain pi where a1: "(pi\<bullet>Xs)\<sharp>*c" and a2: "(pi\<bullet>Xs) \<inter> As = {}" and 
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2626
      a4: "set pi \<subseteq> Xs \<times> (pi\<bullet>Xs)" by blast
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2627
    have b: "x\<notin>Xs" by fact
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2628
    have d1: "finite As" by fact
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2629
    have d2: "finite Xs" by fact
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2630
    have d3: "({x} \<union> Xs) \<subseteq> As" using insert(4) by simp
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2631
    from d d1 d2
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2632
    obtain y::"'a" where fr: "y\<sharp>(c,pi\<bullet>Xs,As)" 
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2633
      apply(rule_tac at_exists_fresh[OF at, where x="(c,pi\<bullet>Xs,As)"])
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2634
      apply(auto simp add: supp_prod at_supp[OF at] at_fin_set_supp[OF at]
46179
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  2635
        pt_supp_finite_pi[OF pt_set_inst[OF at_pt_inst[OF at]] at])
32638
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2636
      done
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2637
    have "({y}\<union>(pi\<bullet>Xs))\<sharp>*c" using a1 fr by (simp add: fresh_star_def)
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2638
    moreover
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2639
    have "({y}\<union>(pi\<bullet>Xs))\<inter>As = {}" using a2 d1 fr 
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2640
      by (simp add: fresh_prod at_fin_set_fresh[OF at])
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2641
    moreover
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2642
    have "pi\<bullet>x=x" using a4 b a2 d3 
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2643
      by (rule_tac at_prm_fresh2[OF at]) (auto)
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2644
    then have "set ((pi\<bullet>x,y)#pi) \<subseteq> ({x} \<union> Xs) \<times> ({y}\<union>(pi\<bullet>Xs))" using a4 by auto
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2645
    moreover
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2646
    have "(((pi\<bullet>x,y)#pi)\<bullet>({x} \<union> Xs)) = {y}\<union>(pi\<bullet>Xs)"
26847
9254cca608ef added at_set_avoiding lemmas
urbanc
parents: 26806
diff changeset
  2647
    proof -
32638
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2648
      have eq: "[(pi\<bullet>x,y)]\<bullet>(pi\<bullet>Xs) = (pi\<bullet>Xs)" 
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2649
      proof -
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  2650
        have "(pi\<bullet>x)\<sharp>(pi\<bullet>Xs)" using b d2 
46179
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  2651
          by (simp add: pt_fresh_bij [OF pt_set_inst [OF at_pt_inst [OF at]], OF at]
45961
5cefe17916a6 treatment of type constructor `set`
haftmann
parents: 45694
diff changeset
  2652
            at_fin_set_fresh [OF at])
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  2653
        moreover
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  2654
        have "y\<sharp>(pi\<bullet>Xs)" using fr by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  2655
        ultimately show "[(pi\<bullet>x,y)]\<bullet>(pi\<bullet>Xs) = (pi\<bullet>Xs)" 
46179
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  2656
          by (simp add: pt_fresh_fresh[OF pt_set_inst
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  2657
            [OF at_pt_inst[OF at]], OF at])
32638
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2658
      qed
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2659
      have "(((pi\<bullet>x,y)#pi)\<bullet>({x}\<union>Xs)) = ([(pi\<bullet>x,y)]\<bullet>(pi\<bullet>({x}\<union>Xs)))"
46179
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  2660
        by (simp add: pt2[symmetric, OF pt_set_inst [OF at_pt_inst[OF at]]])
32638
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2661
      also have "\<dots> = {y}\<union>([(pi\<bullet>x,y)]\<bullet>(pi\<bullet>Xs))" 
46179
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  2662
        by (simp only: union_eqvt perm_set_def at_calc[OF at])(auto)
32638
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2663
      finally show "(((pi\<bullet>x,y)#pi)\<bullet>({x} \<union> Xs)) = {y}\<union>(pi\<bullet>Xs)" using eq by simp
26847
9254cca608ef added at_set_avoiding lemmas
urbanc
parents: 26806
diff changeset
  2664
    qed
32638
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2665
    ultimately 
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2666
    show ?case by (rule_tac x="(pi\<bullet>x,y)#pi" in exI) (auto)
26847
9254cca608ef added at_set_avoiding lemmas
urbanc
parents: 26806
diff changeset
  2667
  qed
9254cca608ef added at_set_avoiding lemmas
urbanc
parents: 26806
diff changeset
  2668
qed
9254cca608ef added at_set_avoiding lemmas
urbanc
parents: 26806
diff changeset
  2669
9254cca608ef added at_set_avoiding lemmas
urbanc
parents: 26806
diff changeset
  2670
lemma at_set_avoiding:
9254cca608ef added at_set_avoiding lemmas
urbanc
parents: 26806
diff changeset
  2671
  fixes Xs::"'a set"
9254cca608ef added at_set_avoiding lemmas
urbanc
parents: 26806
diff changeset
  2672
  assumes at: "at TYPE('a)"
9254cca608ef added at_set_avoiding lemmas
urbanc
parents: 26806
diff changeset
  2673
  and     a: "finite Xs"
9254cca608ef added at_set_avoiding lemmas
urbanc
parents: 26806
diff changeset
  2674
  and     b: "finite ((supp c)::'a set)"
32638
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2675
  obtains pi::"'a prm" where "(pi\<bullet>Xs)\<sharp>*c" and "set pi \<subseteq> Xs \<times> (pi\<bullet>Xs)"
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2676
using a b at_set_avoiding_aux[OF at, where Xs="Xs" and As="Xs" and c="c"]
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2677
by (blast)
d9bd7e01a681 tuned some proofs
Christian Urban <urbanc@in.tum.de>
parents: 31936
diff changeset
  2678
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  2679
section \<open>composition instances\<close>
19477
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2680
(* ============================= *)
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2681
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2682
lemma cp_list_inst:
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2683
  assumes c1: "cp TYPE ('a) TYPE('x) TYPE('y)"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2684
  shows "cp TYPE ('a list) TYPE('x) TYPE('y)"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2685
using c1
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2686
apply(simp add: cp_def)
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2687
apply(auto)
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2688
apply(induct_tac x)
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2689
apply(auto)
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2690
done
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2691
46179
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  2692
lemma cp_set_inst:
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  2693
  assumes c1: "cp TYPE ('a) TYPE('x) TYPE('y)"
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  2694
  shows "cp TYPE ('a set) TYPE('x) TYPE('y)"
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  2695
using c1
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  2696
apply(simp add: cp_def)
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  2697
apply(auto)
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  2698
apply(auto simp add: perm_set_def)
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  2699
apply(rule_tac x="pi2\<bullet>xc" in exI)
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  2700
apply(auto)
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  2701
done
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  2702
19477
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2703
lemma cp_option_inst:
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2704
  assumes c1: "cp TYPE ('a) TYPE('x) TYPE('y)"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2705
  shows "cp TYPE ('a option) TYPE('x) TYPE('y)"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2706
using c1
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2707
apply(simp add: cp_def)
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2708
apply(auto)
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2709
apply(case_tac x)
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2710
apply(auto)
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2711
done
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2712
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2713
lemma cp_noption_inst:
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2714
  assumes c1: "cp TYPE ('a) TYPE('x) TYPE('y)"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2715
  shows "cp TYPE ('a noption) TYPE('x) TYPE('y)"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2716
using c1
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2717
apply(simp add: cp_def)
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2718
apply(auto)
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2719
apply(case_tac x)
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2720
apply(auto)
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2721
done
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2722
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2723
lemma cp_unit_inst:
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2724
  shows "cp TYPE (unit) TYPE('x) TYPE('y)"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2725
apply(simp add: cp_def)
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2726
done
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2727
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2728
lemma cp_bool_inst:
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2729
  shows "cp TYPE (bool) TYPE('x) TYPE('y)"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2730
apply(simp add: cp_def)
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2731
apply(rule allI)+
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2732
apply(induct_tac x)
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2733
apply(simp_all)
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2734
done
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2735
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2736
lemma cp_prod_inst:
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2737
  assumes c1: "cp TYPE ('a) TYPE('x) TYPE('y)"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2738
  and     c2: "cp TYPE ('b) TYPE('x) TYPE('y)"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2739
  shows "cp TYPE ('a\<times>'b) TYPE('x) TYPE('y)"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2740
using c1 c2
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2741
apply(simp add: cp_def)
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2742
done
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2743
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2744
lemma cp_fun_inst:
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2745
  assumes c1: "cp TYPE ('a) TYPE('x) TYPE('y)"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2746
  and     c2: "cp TYPE ('b) TYPE('x) TYPE('y)"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2747
  and     pt: "pt TYPE ('y) TYPE('x)"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2748
  and     at: "at TYPE ('x)"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2749
  shows "cp TYPE ('a\<Rightarrow>'b) TYPE('x) TYPE('y)"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2750
using c1 c2
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  2751
apply(auto simp add: cp_def perm_fun_def fun_eq_iff)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  2752
apply(simp add: rev_eqvt[symmetric])
19477
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2753
apply(simp add: pt_rev_pi[OF pt_list_inst[OF pt_prod_inst[OF pt, OF pt]], OF at])
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2754
done
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2755
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2756
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  2757
section \<open>Andy's freshness lemma\<close>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2758
(*================================*)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2759
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2760
lemma freshness_lemma:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2761
  fixes h :: "'x\<Rightarrow>'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2762
  assumes pta: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2763
  and     at:  "at TYPE('x)" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2764
  and     f1:  "finite ((supp h)::'x set)"
22500
8436bfd21bf3 corrected the lemmas min_nat_eqvt and min_int_eqvt
urbanc
parents: 22446
diff changeset
  2765
  and     a: "\<exists>a::'x. a\<sharp>(h,h a)"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2766
  shows  "\<exists>fr::'a. \<forall>a::'x. a\<sharp>h \<longrightarrow> (h a) = fr"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2767
proof -
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2768
  have ptb: "pt TYPE('x) TYPE('x)" by (simp add: at_pt_inst[OF at]) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2769
  have ptc: "pt TYPE('x\<Rightarrow>'a) TYPE('x)" by (simp add: pt_fun_inst[OF ptb, OF pta, OF at]) 
22500
8436bfd21bf3 corrected the lemmas min_nat_eqvt and min_int_eqvt
urbanc
parents: 22446
diff changeset
  2770
  from a obtain a0 where a1: "a0\<sharp>h" and a2: "a0\<sharp>(h a0)" by (force simp add: fresh_prod)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2771
  show ?thesis
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2772
  proof
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2773
    let ?fr = "h (a0::'x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2774
    show "\<forall>(a::'x). (a\<sharp>h \<longrightarrow> ((h a) = ?fr))" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2775
    proof (intro strip)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2776
      fix a
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2777
      assume a3: "(a::'x)\<sharp>h"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2778
      show "h (a::'x) = h a0"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2779
      proof (cases "a=a0")
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  2780
        case True thus "h (a::'x) = h a0" by simp
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2781
      next
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  2782
        case False 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  2783
        assume "a\<noteq>a0"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  2784
        hence c1: "a\<notin>((supp a0)::'x set)" by  (simp add: fresh_def[symmetric] at_fresh[OF at])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  2785
        have c2: "a\<notin>((supp h)::'x set)" using a3 by (simp add: fresh_def)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  2786
        from c1 c2 have c3: "a\<notin>((supp h)\<union>((supp a0)::'x set))" by force
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  2787
        have f2: "finite ((supp a0)::'x set)" by (simp add: at_supp[OF at])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  2788
        from f1 f2 have "((supp (h a0))::'x set)\<subseteq>((supp h)\<union>(supp a0))"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  2789
          by (simp add: pt_supp_fun_subset[OF ptb, OF pta, OF at])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  2790
        hence "a\<notin>((supp (h a0))::'x set)" using c3 by force
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  2791
        hence "a\<sharp>(h a0)" by (simp add: fresh_def) 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  2792
        with a2 have d1: "[(a0,a)]\<bullet>(h a0) = (h a0)" by (rule pt_fresh_fresh[OF pta, OF at])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  2793
        from a1 a3 have d2: "[(a0,a)]\<bullet>h = h" by (rule pt_fresh_fresh[OF ptc, OF at])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  2794
        from d1 have "h a0 = [(a0,a)]\<bullet>(h a0)" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  2795
        also have "\<dots>= ([(a0,a)]\<bullet>h)([(a0,a)]\<bullet>a0)" by (simp add: pt_fun_app_eq[OF ptb, OF at])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  2796
        also have "\<dots> = h ([(a0,a)]\<bullet>a0)" using d2 by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  2797
        also have "\<dots> = h a" by (simp add: at_calc[OF at])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  2798
        finally show "h a = h a0" by simp
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2799
      qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2800
    qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2801
  qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2802
qed
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  2803
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2804
lemma freshness_lemma_unique:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2805
  fixes h :: "'x\<Rightarrow>'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2806
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2807
  and     at: "at TYPE('x)" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2808
  and     f1: "finite ((supp h)::'x set)"
22500
8436bfd21bf3 corrected the lemmas min_nat_eqvt and min_int_eqvt
urbanc
parents: 22446
diff changeset
  2809
  and     a: "\<exists>(a::'x). a\<sharp>(h,h a)"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2810
  shows  "\<exists>!(fr::'a). \<forall>(a::'x). a\<sharp>h \<longrightarrow> (h a) = fr"
18703
13e11abcfc96 fixed one proof that broke because of the changes
urbanc
parents: 18657
diff changeset
  2811
proof (rule ex_ex1I)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2812
  from pt at f1 a show "\<exists>fr::'a. \<forall>a::'x. a\<sharp>h \<longrightarrow> h a = fr" by (simp add: freshness_lemma)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2813
next
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2814
  fix fr1 fr2
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2815
  assume b1: "\<forall>a::'x. a\<sharp>h \<longrightarrow> h a = fr1"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2816
  assume b2: "\<forall>a::'x. a\<sharp>h \<longrightarrow> h a = fr2"
22500
8436bfd21bf3 corrected the lemmas min_nat_eqvt and min_int_eqvt
urbanc
parents: 22446
diff changeset
  2817
  from a obtain a where "(a::'x)\<sharp>h" by (force simp add: fresh_prod) 
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2818
  with b1 b2 have "h a = fr1 \<and> h a = fr2" by force
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2819
  thus "fr1 = fr2" by force
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2820
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2821
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  2822
\<comment> "packaging the freshness lemma into a function"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
  2823
definition fresh_fun :: "('x\<Rightarrow>'a)\<Rightarrow>'a" where
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2824
  "fresh_fun (h) \<equiv> THE fr. (\<forall>(a::'x). a\<sharp>h \<longrightarrow> (h a) = fr)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2825
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2826
lemma fresh_fun_app:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2827
  fixes h :: "'x\<Rightarrow>'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2828
  and   a :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2829
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2830
  and     at: "at TYPE('x)" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2831
  and     f1: "finite ((supp h)::'x set)"
22500
8436bfd21bf3 corrected the lemmas min_nat_eqvt and min_int_eqvt
urbanc
parents: 22446
diff changeset
  2832
  and     a: "\<exists>(a::'x). a\<sharp>(h,h a)"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2833
  and     b: "a\<sharp>h"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2834
  shows "(fresh_fun h) = (h a)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2835
proof (unfold fresh_fun_def, rule the_equality)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2836
  show "\<forall>(a'::'x). a'\<sharp>h \<longrightarrow> h a' = h a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2837
  proof (intro strip)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2838
    fix a'::"'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2839
    assume c: "a'\<sharp>h"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2840
    from pt at f1 a have "\<exists>(fr::'a). \<forall>(a::'x). a\<sharp>h \<longrightarrow> (h a) = fr" by (rule freshness_lemma)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2841
    with b c show "h a' = h a" by force
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2842
  qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2843
next
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2844
  fix fr::"'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2845
  assume "\<forall>a. a\<sharp>h \<longrightarrow> h a = fr"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2846
  with b show "fr = h a" by force
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2847
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2848
22714
ca804eb70d39 added a more usuable lemma for dealing with fresh_fun
urbanc
parents: 22650
diff changeset
  2849
lemma fresh_fun_app':
ca804eb70d39 added a more usuable lemma for dealing with fresh_fun
urbanc
parents: 22650
diff changeset
  2850
  fixes h :: "'x\<Rightarrow>'a"
ca804eb70d39 added a more usuable lemma for dealing with fresh_fun
urbanc
parents: 22650
diff changeset
  2851
  and   a :: "'x"
ca804eb70d39 added a more usuable lemma for dealing with fresh_fun
urbanc
parents: 22650
diff changeset
  2852
  assumes pt: "pt TYPE('a) TYPE('x)"
ca804eb70d39 added a more usuable lemma for dealing with fresh_fun
urbanc
parents: 22650
diff changeset
  2853
  and     at: "at TYPE('x)" 
ca804eb70d39 added a more usuable lemma for dealing with fresh_fun
urbanc
parents: 22650
diff changeset
  2854
  and     f1: "finite ((supp h)::'x set)"
ca804eb70d39 added a more usuable lemma for dealing with fresh_fun
urbanc
parents: 22650
diff changeset
  2855
  and     a: "a\<sharp>h" "a\<sharp>h a"
ca804eb70d39 added a more usuable lemma for dealing with fresh_fun
urbanc
parents: 22650
diff changeset
  2856
  shows "(fresh_fun h) = (h a)"
ca804eb70d39 added a more usuable lemma for dealing with fresh_fun
urbanc
parents: 22650
diff changeset
  2857
apply(rule fresh_fun_app[OF pt, OF at, OF f1])
ca804eb70d39 added a more usuable lemma for dealing with fresh_fun
urbanc
parents: 22650
diff changeset
  2858
apply(auto simp add: fresh_prod intro: a)
ca804eb70d39 added a more usuable lemma for dealing with fresh_fun
urbanc
parents: 22650
diff changeset
  2859
done
ca804eb70d39 added a more usuable lemma for dealing with fresh_fun
urbanc
parents: 22650
diff changeset
  2860
19477
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2861
lemma fresh_fun_equiv_ineq:
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2862
  fixes h :: "'y\<Rightarrow>'a"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2863
  and   pi:: "'x prm"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2864
  assumes pta: "pt TYPE('a) TYPE('x)"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2865
  and     ptb: "pt TYPE('y) TYPE('x)"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2866
  and     ptb':"pt TYPE('a) TYPE('y)"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2867
  and     at:  "at TYPE('x)" 
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2868
  and     at': "at TYPE('y)"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2869
  and     cpa: "cp TYPE('a) TYPE('x) TYPE('y)"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2870
  and     cpb: "cp TYPE('y) TYPE('x) TYPE('y)"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2871
  and     f1: "finite ((supp h)::'y set)"
22500
8436bfd21bf3 corrected the lemmas min_nat_eqvt and min_int_eqvt
urbanc
parents: 22446
diff changeset
  2872
  and     a1: "\<exists>(a::'y). a\<sharp>(h,h a)"
19477
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2873
  shows "pi\<bullet>(fresh_fun h) = fresh_fun(pi\<bullet>h)" (is "?LHS = ?RHS")
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2874
proof -
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2875
  have ptd: "pt TYPE('y) TYPE('y)" by (simp add: at_pt_inst[OF at']) 
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2876
  have ptc: "pt TYPE('y\<Rightarrow>'a) TYPE('x)" by (simp add: pt_fun_inst[OF ptb, OF pta, OF at]) 
23393
31781b2de73d tuned proofs: avoid implicit prems;
wenzelm
parents: 23159
diff changeset
  2877
  have cpc: "cp TYPE('y\<Rightarrow>'a) TYPE ('x) TYPE ('y)" by (rule cp_fun_inst[OF cpb cpa ptb at])
19477
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2878
  have f2: "finite ((supp (pi\<bullet>h))::'y set)"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2879
  proof -
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2880
    from f1 have "finite (pi\<bullet>((supp h)::'y set))"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2881
      by (simp add: pt_set_finite_ineq[OF ptb, OF at])
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2882
    thus ?thesis
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2883
      by (simp add: pt_perm_supp_ineq[OF ptc, OF ptb, OF at, OF cpc])
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2884
  qed
22500
8436bfd21bf3 corrected the lemmas min_nat_eqvt and min_int_eqvt
urbanc
parents: 22446
diff changeset
  2885
  from a1 obtain a' where c0: "a'\<sharp>(h,h a')" by force
8436bfd21bf3 corrected the lemmas min_nat_eqvt and min_int_eqvt
urbanc
parents: 22446
diff changeset
  2886
  hence c1: "a'\<sharp>h" and c2: "a'\<sharp>(h a')" by (simp_all add: fresh_prod)
19477
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2887
  have c3: "(pi\<bullet>a')\<sharp>(pi\<bullet>h)" using c1
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2888
  by (simp add: pt_fresh_bij_ineq[OF ptc, OF ptb, OF at, OF cpc])
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2889
  have c4: "(pi\<bullet>a')\<sharp>(pi\<bullet>h) (pi\<bullet>a')"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2890
  proof -
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2891
    from c2 have "(pi\<bullet>a')\<sharp>(pi\<bullet>(h a'))"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2892
      by (simp add: pt_fresh_bij_ineq[OF pta, OF ptb, OF at,OF cpa])
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2893
    thus ?thesis by (simp add: pt_fun_app_eq[OF ptb, OF at])
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2894
  qed
22500
8436bfd21bf3 corrected the lemmas min_nat_eqvt and min_int_eqvt
urbanc
parents: 22446
diff changeset
  2895
  have a2: "\<exists>(a::'y). a\<sharp>(pi\<bullet>h,(pi\<bullet>h) a)" using c3 c4 by (force simp add: fresh_prod)
19477
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2896
  have d1: "?LHS = pi\<bullet>(h a')" using c1 a1 by (simp add: fresh_fun_app[OF ptb', OF at', OF f1])
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2897
  have d2: "?RHS = (pi\<bullet>h) (pi\<bullet>a')" using c3 a2 
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2898
    by (simp add: fresh_fun_app[OF ptb', OF at', OF f2])
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2899
  show ?thesis using d1 d2 by (simp add: pt_fun_app_eq[OF ptb, OF at])
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2900
qed
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  2901
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2902
lemma fresh_fun_equiv:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2903
  fixes h :: "'x\<Rightarrow>'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2904
  and   pi:: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2905
  assumes pta: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2906
  and     at:  "at TYPE('x)" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2907
  and     f1:  "finite ((supp h)::'x set)"
22500
8436bfd21bf3 corrected the lemmas min_nat_eqvt and min_int_eqvt
urbanc
parents: 22446
diff changeset
  2908
  and     a1: "\<exists>(a::'x). a\<sharp>(h,h a)"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2909
  shows "pi\<bullet>(fresh_fun h) = fresh_fun(pi\<bullet>h)" (is "?LHS = ?RHS")
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2910
proof -
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2911
  have ptb: "pt TYPE('x) TYPE('x)" by (simp add: at_pt_inst[OF at]) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2912
  have ptc: "pt TYPE('x\<Rightarrow>'a) TYPE('x)" by (simp add: pt_fun_inst[OF ptb, OF pta, OF at]) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2913
  have f2: "finite ((supp (pi\<bullet>h))::'x set)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2914
  proof -
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2915
    from f1 have "finite (pi\<bullet>((supp h)::'x set))" by (simp add: pt_set_finite_ineq[OF ptb, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2916
    thus ?thesis by (simp add: pt_perm_supp[OF ptc, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2917
  qed
22500
8436bfd21bf3 corrected the lemmas min_nat_eqvt and min_int_eqvt
urbanc
parents: 22446
diff changeset
  2918
  from a1 obtain a' where c0: "a'\<sharp>(h,h a')" by force
8436bfd21bf3 corrected the lemmas min_nat_eqvt and min_int_eqvt
urbanc
parents: 22446
diff changeset
  2919
  hence c1: "a'\<sharp>h" and c2: "a'\<sharp>(h a')" by (simp_all add: fresh_prod)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2920
  have c3: "(pi\<bullet>a')\<sharp>(pi\<bullet>h)" using c1 by (simp add: pt_fresh_bij[OF ptc, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2921
  have c4: "(pi\<bullet>a')\<sharp>(pi\<bullet>h) (pi\<bullet>a')"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2922
  proof -
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2923
    from c2 have "(pi\<bullet>a')\<sharp>(pi\<bullet>(h a'))" by (simp add: pt_fresh_bij[OF pta, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2924
    thus ?thesis by (simp add: pt_fun_app_eq[OF ptb, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2925
  qed
22500
8436bfd21bf3 corrected the lemmas min_nat_eqvt and min_int_eqvt
urbanc
parents: 22446
diff changeset
  2926
  have a2: "\<exists>(a::'x). a\<sharp>(pi\<bullet>h,(pi\<bullet>h) a)" using c3 c4 by (force simp add: fresh_prod)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2927
  have d1: "?LHS = pi\<bullet>(h a')" using c1 a1 by (simp add: fresh_fun_app[OF pta, OF at, OF f1])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2928
  have d2: "?RHS = (pi\<bullet>h) (pi\<bullet>a')" using c3 a2 by (simp add: fresh_fun_app[OF pta, OF at, OF f2])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2929
  show ?thesis using d1 d2 by (simp add: pt_fun_app_eq[OF ptb, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2930
qed
19216
a45baf1ac094 tuned some of the proofs about fresh_fun
urbanc
parents: 19164
diff changeset
  2931
a45baf1ac094 tuned some of the proofs about fresh_fun
urbanc
parents: 19164
diff changeset
  2932
lemma fresh_fun_supports:
a45baf1ac094 tuned some of the proofs about fresh_fun
urbanc
parents: 19164
diff changeset
  2933
  fixes h :: "'x\<Rightarrow>'a"
a45baf1ac094 tuned some of the proofs about fresh_fun
urbanc
parents: 19164
diff changeset
  2934
  assumes pt: "pt TYPE('a) TYPE('x)"
a45baf1ac094 tuned some of the proofs about fresh_fun
urbanc
parents: 19164
diff changeset
  2935
  and     at: "at TYPE('x)" 
a45baf1ac094 tuned some of the proofs about fresh_fun
urbanc
parents: 19164
diff changeset
  2936
  and     f1: "finite ((supp h)::'x set)"
22500
8436bfd21bf3 corrected the lemmas min_nat_eqvt and min_int_eqvt
urbanc
parents: 22446
diff changeset
  2937
  and     a: "\<exists>(a::'x). a\<sharp>(h,h a)"
19216
a45baf1ac094 tuned some of the proofs about fresh_fun
urbanc
parents: 19164
diff changeset
  2938
  shows "((supp h)::'x set) supports (fresh_fun h)"
22808
a7daa74e2980 eliminated unnamed infixes, tuned syntax;
wenzelm
parents: 22786
diff changeset
  2939
  apply(simp add: supports_def fresh_def[symmetric])
19216
a45baf1ac094 tuned some of the proofs about fresh_fun
urbanc
parents: 19164
diff changeset
  2940
  apply(auto)
a45baf1ac094 tuned some of the proofs about fresh_fun
urbanc
parents: 19164
diff changeset
  2941
  apply(simp add: fresh_fun_equiv[OF pt, OF at, OF f1, OF a])
a45baf1ac094 tuned some of the proofs about fresh_fun
urbanc
parents: 19164
diff changeset
  2942
  apply(simp add: pt_fresh_fresh[OF pt_fun_inst[OF at_pt_inst[OF at], OF pt], OF at, OF at])
a45baf1ac094 tuned some of the proofs about fresh_fun
urbanc
parents: 19164
diff changeset
  2943
  done
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2944
  
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  2945
section \<open>Abstraction function\<close>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2946
(*==============================*)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2947
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2948
lemma pt_abs_fun_inst:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2949
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2950
  and     at: "at TYPE('x)"
18579
002d371401f5 changed the name of the type "nOption" to "noption".
urbanc
parents: 18578
diff changeset
  2951
  shows "pt TYPE('x\<Rightarrow>('a noption)) TYPE('x)"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2952
  by (rule pt_fun_inst[OF at_pt_inst[OF at],OF pt_noption_inst[OF pt],OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2953
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
  2954
definition abs_fun :: "'x\<Rightarrow>'a\<Rightarrow>('x\<Rightarrow>('a noption))" ("[_]._" [100,100] 100) where 
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2955
  "[a].x \<equiv> (\<lambda>b. (if b=a then nSome(x) else (if b\<sharp>x then nSome([(a,b)]\<bullet>x) else nNone)))"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2956
18745
060400dc077c a fixme comments about abs_fun_if, which should be called perm_if
urbanc
parents: 18703
diff changeset
  2957
(* FIXME: should be called perm_if and placed close to the definition of permutations on bools *)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2958
lemma abs_fun_if: 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2959
  fixes pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2960
  and   x  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2961
  and   y  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2962
  and   c  :: "bool"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2963
  shows "pi\<bullet>(if c then x else y) = (if c then (pi\<bullet>x) else (pi\<bullet>y))"   
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2964
  by force
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2965
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2966
lemma abs_fun_pi_ineq:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2967
  fixes a  :: "'y"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2968
  and   x  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2969
  and   pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2970
  assumes pta: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2971
  and     ptb: "pt TYPE('y) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2972
  and     at:  "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2973
  and     cp:  "cp TYPE('a) TYPE('x) TYPE('y)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2974
  shows "pi\<bullet>([a].x) = [(pi\<bullet>a)].(pi\<bullet>x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2975
  apply(simp add: abs_fun_def perm_fun_def abs_fun_if)
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  2976
  apply(simp only: fun_eq_iff)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2977
  apply(rule allI)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2978
  apply(subgoal_tac "(((rev pi)\<bullet>(xa::'y)) = (a::'y)) = (xa = pi\<bullet>a)")(*A*)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2979
  apply(subgoal_tac "(((rev pi)\<bullet>xa)\<sharp>x) = (xa\<sharp>(pi\<bullet>x))")(*B*)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2980
  apply(subgoal_tac "pi\<bullet>([(a,(rev pi)\<bullet>xa)]\<bullet>x) = [(pi\<bullet>a,xa)]\<bullet>(pi\<bullet>x)")(*C*)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2981
  apply(simp)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2982
(*C*)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2983
  apply(simp add: cp1[OF cp])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2984
  apply(simp add: pt_pi_rev[OF ptb, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2985
(*B*)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2986
  apply(simp add: pt_fresh_left_ineq[OF pta, OF ptb, OF at, OF cp])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2987
(*A*)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2988
  apply(rule iffI)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2989
  apply(rule pt_bij2[OF ptb, OF at, THEN sym])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2990
  apply(simp)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2991
  apply(rule pt_bij2[OF ptb, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2992
  apply(simp)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2993
done
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2994
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2995
lemma abs_fun_pi:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2996
  fixes a  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2997
  and   x  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2998
  and   pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  2999
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3000
  and     at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3001
  shows "pi\<bullet>([a].x) = [(pi\<bullet>a)].(pi\<bullet>x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3002
apply(rule abs_fun_pi_ineq)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3003
apply(rule pt)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3004
apply(rule at_pt_inst)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3005
apply(rule at)+
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3006
apply(rule cp_pt_inst)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3007
apply(rule pt)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3008
apply(rule at)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3009
done
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3010
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3011
lemma abs_fun_eq1: 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3012
  fixes x  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3013
  and   y  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3014
  and   a  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3015
  shows "([a].x = [a].y) = (x = y)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3016
apply(auto simp add: abs_fun_def)
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  3017
apply(auto simp add: fun_eq_iff)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3018
apply(drule_tac x="a" in spec)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3019
apply(simp)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3020
done
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3021
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3022
lemma abs_fun_eq2:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3023
  fixes x  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3024
  and   y  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3025
  and   a  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3026
  and   b  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3027
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3028
      and at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3029
      and a1: "a\<noteq>b" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3030
      and a2: "[a].x = [b].y" 
18268
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3031
  shows "x=[(a,b)]\<bullet>y \<and> a\<sharp>y"
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3032
proof -
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  3033
  from a2 have "\<forall>c::'x. ([a].x) c = ([b].y) c" by (force simp add: fun_eq_iff)
18268
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3034
  hence "([a].x) a = ([b].y) a" by simp
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3035
  hence a3: "nSome(x) = ([b].y) a" by (simp add: abs_fun_def)
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3036
  show "x=[(a,b)]\<bullet>y \<and> a\<sharp>y"
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3037
  proof (cases "a\<sharp>y")
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3038
    assume a4: "a\<sharp>y"
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3039
    hence "x=[(b,a)]\<bullet>y" using a3 a1 by (simp add: abs_fun_def)
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3040
    moreover
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3041
    have "[(a,b)]\<bullet>y = [(b,a)]\<bullet>y" by (rule pt3[OF pt], rule at_ds5[OF at])
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3042
    ultimately show ?thesis using a4 by simp
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3043
  next
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3044
    assume "\<not>a\<sharp>y"
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3045
    hence "nSome(x) = nNone" using a1 a3 by (simp add: abs_fun_def)
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3046
    hence False by simp
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3047
    thus ?thesis by simp
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3048
  qed
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3049
qed
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3050
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3051
lemma abs_fun_eq3: 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3052
  fixes x  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3053
  and   y  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3054
  and   a   :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3055
  and   b   :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3056
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3057
      and at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3058
      and a1: "a\<noteq>b" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3059
      and a2: "x=[(a,b)]\<bullet>y" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3060
      and a3: "a\<sharp>y" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3061
  shows "[a].x =[b].y"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3062
proof -
18268
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3063
  show ?thesis 
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  3064
  proof (simp only: abs_fun_def fun_eq_iff, intro strip)
18268
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3065
    fix c::"'x"
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3066
    let ?LHS = "if c=a then nSome(x) else if c\<sharp>x then nSome([(a,c)]\<bullet>x) else nNone"
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3067
    and ?RHS = "if c=b then nSome(y) else if c\<sharp>y then nSome([(b,c)]\<bullet>y) else nNone"
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3068
    show "?LHS=?RHS"
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3069
    proof -
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3070
      have "(c=a) \<or> (c=b) \<or> (c\<noteq>a \<and> c\<noteq>b)" by blast
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  3071
      moreover  \<comment>"case c=a"
18268
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3072
      { have "nSome(x) = nSome([(a,b)]\<bullet>y)" using a2 by simp
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  3073
        also have "\<dots> = nSome([(b,a)]\<bullet>y)" by (simp, rule pt3[OF pt], rule at_ds5[OF at])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  3074
        finally have "nSome(x) = nSome([(b,a)]\<bullet>y)" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  3075
        moreover
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  3076
        assume "c=a"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  3077
        ultimately have "?LHS=?RHS" using a1 a3 by simp
18268
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3078
      }
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  3079
      moreover  \<comment> "case c=b"
18268
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3080
      { have a4: "y=[(a,b)]\<bullet>x" using a2 by (simp only: pt_swap_bij[OF pt, OF at])
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  3081
        hence "a\<sharp>([(a,b)]\<bullet>x)" using a3 by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  3082
        hence "b\<sharp>x" by (simp add: at_calc[OF at] pt_fresh_left[OF pt, OF at])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  3083
        moreover
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  3084
        assume "c=b"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  3085
        ultimately have "?LHS=?RHS" using a1 a4 by simp
18268
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3086
      }
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  3087
      moreover  \<comment> "case c\<noteq>a \<and> c\<noteq>b"
18268
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3088
      { assume a5: "c\<noteq>a \<and> c\<noteq>b"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  3089
        moreover 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  3090
        have "c\<sharp>x = c\<sharp>y" using a2 a5 by (force simp add: at_calc[OF at] pt_fresh_left[OF pt, OF at])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  3091
        moreover 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  3092
        have "c\<sharp>y \<longrightarrow> [(a,c)]\<bullet>x = [(b,c)]\<bullet>y" 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  3093
        proof (intro strip)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  3094
          assume a6: "c\<sharp>y"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  3095
          have "[(a,c),(b,c),(a,c)] \<triangleq> [(a,b)]" using a1 a5 by (force intro: at_ds3[OF at])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  3096
          hence "[(a,c)]\<bullet>([(b,c)]\<bullet>([(a,c)]\<bullet>y)) = [(a,b)]\<bullet>y" 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  3097
            by (simp add: pt2[OF pt, symmetric] pt3[OF pt])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  3098
          hence "[(a,c)]\<bullet>([(b,c)]\<bullet>y) = [(a,b)]\<bullet>y" using a3 a6 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  3099
            by (simp add: pt_fresh_fresh[OF pt, OF at])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  3100
          hence "[(a,c)]\<bullet>([(b,c)]\<bullet>y) = x" using a2 by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  3101
          hence "[(b,c)]\<bullet>y = [(a,c)]\<bullet>x" by (drule_tac pt_bij1[OF pt, OF at], simp)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  3102
          thus "[(a,c)]\<bullet>x = [(b,c)]\<bullet>y" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  3103
        qed
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  3104
        ultimately have "?LHS=?RHS" by simp
18268
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3105
      }
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3106
      ultimately show "?LHS = ?RHS" by blast
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3107
    qed
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3108
  qed
18268
734f23ad5d8f ISAR-fied two proofs about equality for abstraction functions.
urbanc
parents: 18264
diff changeset
  3109
qed
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  3110
        
23158
749b6870b1a1 introduced symmetric variants of the lemmas for alpha-equivalence
urbanc
parents: 23050
diff changeset
  3111
(* alpha equivalence *)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3112
lemma abs_fun_eq: 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3113
  fixes x  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3114
  and   y  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3115
  and   a  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3116
  and   b  :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3117
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3118
      and at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3119
  shows "([a].x = [b].y) = ((a=b \<and> x=y)\<or>(a\<noteq>b \<and> x=[(a,b)]\<bullet>y \<and> a\<sharp>y))"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3120
proof (rule iffI)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3121
  assume b: "[a].x = [b].y"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3122
  show "(a=b \<and> x=y)\<or>(a\<noteq>b \<and> x=[(a,b)]\<bullet>y \<and> a\<sharp>y)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3123
  proof (cases "a=b")
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3124
    case True with b show ?thesis by (simp add: abs_fun_eq1)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3125
  next
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3126
    case False with b show ?thesis by (simp add: abs_fun_eq2[OF pt, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3127
  qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3128
next
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3129
  assume "(a=b \<and> x=y)\<or>(a\<noteq>b \<and> x=[(a,b)]\<bullet>y \<and> a\<sharp>y)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3130
  thus "[a].x = [b].y"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3131
  proof
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3132
    assume "a=b \<and> x=y" thus ?thesis by simp
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3133
  next
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3134
    assume "a\<noteq>b \<and> x=[(a,b)]\<bullet>y \<and> a\<sharp>y" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3135
    thus ?thesis by (simp add: abs_fun_eq3[OF pt, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3136
  qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3137
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3138
23158
749b6870b1a1 introduced symmetric variants of the lemmas for alpha-equivalence
urbanc
parents: 23050
diff changeset
  3139
(* symmetric version of alpha-equivalence *)
19562
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3140
lemma abs_fun_eq': 
23158
749b6870b1a1 introduced symmetric variants of the lemmas for alpha-equivalence
urbanc
parents: 23050
diff changeset
  3141
  fixes x  :: "'a"
749b6870b1a1 introduced symmetric variants of the lemmas for alpha-equivalence
urbanc
parents: 23050
diff changeset
  3142
  and   y  :: "'a"
749b6870b1a1 introduced symmetric variants of the lemmas for alpha-equivalence
urbanc
parents: 23050
diff changeset
  3143
  and   a  :: "'x"
749b6870b1a1 introduced symmetric variants of the lemmas for alpha-equivalence
urbanc
parents: 23050
diff changeset
  3144
  and   b  :: "'x"
749b6870b1a1 introduced symmetric variants of the lemmas for alpha-equivalence
urbanc
parents: 23050
diff changeset
  3145
  assumes pt: "pt TYPE('a) TYPE('x)"
749b6870b1a1 introduced symmetric variants of the lemmas for alpha-equivalence
urbanc
parents: 23050
diff changeset
  3146
      and at: "at TYPE('x)"
23159
792ff2490f91 tuned the proof
urbanc
parents: 23158
diff changeset
  3147
  shows "([a].x = [b].y) = ((a=b \<and> x=y)\<or>(a\<noteq>b \<and> [(b,a)]\<bullet>x=y \<and> b\<sharp>x))"
792ff2490f91 tuned the proof
urbanc
parents: 23158
diff changeset
  3148
by (auto simp add: abs_fun_eq[OF pt, OF at] pt_swap_bij'[OF pt, OF at] 
23158
749b6870b1a1 introduced symmetric variants of the lemmas for alpha-equivalence
urbanc
parents: 23050
diff changeset
  3149
                   pt_fresh_left[OF pt, OF at] 
749b6870b1a1 introduced symmetric variants of the lemmas for alpha-equivalence
urbanc
parents: 23050
diff changeset
  3150
                   at_calc[OF at])
749b6870b1a1 introduced symmetric variants of the lemmas for alpha-equivalence
urbanc
parents: 23050
diff changeset
  3151
749b6870b1a1 introduced symmetric variants of the lemmas for alpha-equivalence
urbanc
parents: 23050
diff changeset
  3152
(* alpha_equivalence with a fresh name *)
749b6870b1a1 introduced symmetric variants of the lemmas for alpha-equivalence
urbanc
parents: 23050
diff changeset
  3153
lemma abs_fun_fresh: 
19562
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3154
  fixes x :: "'a"
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3155
  and   y :: "'a"
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3156
  and   c :: "'x"
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3157
  and   a :: "'x"
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3158
  and   b :: "'x"
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3159
  assumes pt: "pt TYPE('a) TYPE('x)"
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3160
      and at: "at TYPE('x)"
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3161
      and fr: "c\<noteq>a" "c\<noteq>b" "c\<sharp>x" "c\<sharp>y" 
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3162
  shows "([a].x = [b].y) = ([(a,c)]\<bullet>x = [(b,c)]\<bullet>y)"
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3163
proof (rule iffI)
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3164
  assume eq0: "[a].x = [b].y"
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3165
  show "[(a,c)]\<bullet>x = [(b,c)]\<bullet>y"
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3166
  proof (cases "a=b")
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3167
    case True then show ?thesis using eq0 by (simp add: pt_bij[OF pt, OF at] abs_fun_eq[OF pt, OF at])
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3168
  next
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3169
    case False 
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3170
    have ineq: "a\<noteq>b" by fact
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3171
    with eq0 have eq: "x=[(a,b)]\<bullet>y" and fr': "a\<sharp>y" by (simp_all add: abs_fun_eq[OF pt, OF at])
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3172
    from eq have "[(a,c)]\<bullet>x = [(a,c)]\<bullet>[(a,b)]\<bullet>y" by (simp add: pt_bij[OF pt, OF at])
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3173
    also have "\<dots> = ([(a,c)]\<bullet>[(a,b)])\<bullet>([(a,c)]\<bullet>y)" by (rule pt_perm_compose[OF pt, OF at])
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3174
    also have "\<dots> = [(c,b)]\<bullet>y" using ineq fr fr' 
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3175
      by (simp add: pt_fresh_fresh[OF pt, OF at] at_calc[OF at])
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3176
    also have "\<dots> = [(b,c)]\<bullet>y" by (rule pt3[OF pt], rule at_ds5[OF at])
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3177
    finally show ?thesis by simp
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3178
  qed
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3179
next
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3180
  assume eq: "[(a,c)]\<bullet>x = [(b,c)]\<bullet>y"
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3181
  thus "[a].x = [b].y"
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3182
  proof (cases "a=b")
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3183
    case True then show ?thesis using eq by (simp add: pt_bij[OF pt, OF at] abs_fun_eq[OF pt, OF at])
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3184
  next
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3185
    case False
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3186
    have ineq: "a\<noteq>b" by fact
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3187
    from fr have "([(a,c)]\<bullet>c)\<sharp>([(a,c)]\<bullet>x)" by (simp add: pt_fresh_bij[OF pt, OF at])
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3188
    hence "a\<sharp>([(b,c)]\<bullet>y)" using eq fr by (simp add: at_calc[OF at])
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3189
    hence fr0: "a\<sharp>y" using ineq fr by (simp add: pt_fresh_left[OF pt, OF at] at_calc[OF at])
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3190
    from eq have "x = (rev [(a,c)])\<bullet>([(b,c)]\<bullet>y)" by (rule pt_bij1[OF pt, OF at])
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3191
    also have "\<dots> = [(a,c)]\<bullet>([(b,c)]\<bullet>y)" by simp
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3192
    also have "\<dots> = ([(a,c)]\<bullet>[(b,c)])\<bullet>([(a,c)]\<bullet>y)" by (rule pt_perm_compose[OF pt, OF at])
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3193
    also have "\<dots> = [(b,a)]\<bullet>y" using ineq fr fr0  
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3194
      by (simp add: pt_fresh_fresh[OF pt, OF at] at_calc[OF at])
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3195
    also have "\<dots> = [(a,b)]\<bullet>y" by (rule pt3[OF pt], rule at_ds5[OF at])
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3196
    finally show ?thesis using ineq fr0 by (simp add: abs_fun_eq[OF pt, OF at])
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3197
  qed
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3198
qed
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19494
diff changeset
  3199
23158
749b6870b1a1 introduced symmetric variants of the lemmas for alpha-equivalence
urbanc
parents: 23050
diff changeset
  3200
lemma abs_fun_fresh': 
749b6870b1a1 introduced symmetric variants of the lemmas for alpha-equivalence
urbanc
parents: 23050
diff changeset
  3201
  fixes x :: "'a"
749b6870b1a1 introduced symmetric variants of the lemmas for alpha-equivalence
urbanc
parents: 23050
diff changeset
  3202
  and   y :: "'a"
749b6870b1a1 introduced symmetric variants of the lemmas for alpha-equivalence
urbanc
parents: 23050
diff changeset
  3203
  and   c :: "'x"
749b6870b1a1 introduced symmetric variants of the lemmas for alpha-equivalence
urbanc
parents: 23050
diff changeset
  3204
  and   a :: "'x"
749b6870b1a1 introduced symmetric variants of the lemmas for alpha-equivalence
urbanc
parents: 23050
diff changeset
  3205
  and   b :: "'x"
749b6870b1a1 introduced symmetric variants of the lemmas for alpha-equivalence
urbanc
parents: 23050
diff changeset
  3206
  assumes pt: "pt TYPE('a) TYPE('x)"
749b6870b1a1 introduced symmetric variants of the lemmas for alpha-equivalence
urbanc
parents: 23050
diff changeset
  3207
      and at: "at TYPE('x)"
749b6870b1a1 introduced symmetric variants of the lemmas for alpha-equivalence
urbanc
parents: 23050
diff changeset
  3208
      and as: "[a].x = [b].y"
749b6870b1a1 introduced symmetric variants of the lemmas for alpha-equivalence
urbanc
parents: 23050
diff changeset
  3209
      and fr: "c\<noteq>a" "c\<noteq>b" "c\<sharp>x" "c\<sharp>y" 
749b6870b1a1 introduced symmetric variants of the lemmas for alpha-equivalence
urbanc
parents: 23050
diff changeset
  3210
  shows "x = [(a,c)]\<bullet>[(b,c)]\<bullet>y"
749b6870b1a1 introduced symmetric variants of the lemmas for alpha-equivalence
urbanc
parents: 23050
diff changeset
  3211
using as fr
749b6870b1a1 introduced symmetric variants of the lemmas for alpha-equivalence
urbanc
parents: 23050
diff changeset
  3212
apply(drule_tac sym)
749b6870b1a1 introduced symmetric variants of the lemmas for alpha-equivalence
urbanc
parents: 23050
diff changeset
  3213
apply(simp add: abs_fun_fresh[OF pt, OF at] pt_swap_bij[OF pt, OF at])
749b6870b1a1 introduced symmetric variants of the lemmas for alpha-equivalence
urbanc
parents: 23050
diff changeset
  3214
done
749b6870b1a1 introduced symmetric variants of the lemmas for alpha-equivalence
urbanc
parents: 23050
diff changeset
  3215
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3216
lemma abs_fun_supp_approx:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3217
  fixes x :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3218
  and   a :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3219
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3220
  and     at: "at TYPE('x)"
18048
7003308ff73a tuned my last commit
urbanc
parents: 18047
diff changeset
  3221
  shows "((supp ([a].x))::'x set) \<subseteq> (supp (x,a))"
7003308ff73a tuned my last commit
urbanc
parents: 18047
diff changeset
  3222
proof 
7003308ff73a tuned my last commit
urbanc
parents: 18047
diff changeset
  3223
  fix c
7003308ff73a tuned my last commit
urbanc
parents: 18047
diff changeset
  3224
  assume "c\<in>((supp ([a].x))::'x set)"
7003308ff73a tuned my last commit
urbanc
parents: 18047
diff changeset
  3225
  hence "infinite {b. [(c,b)]\<bullet>([a].x) \<noteq> [a].x}" by (simp add: supp_def)
7003308ff73a tuned my last commit
urbanc
parents: 18047
diff changeset
  3226
  hence "infinite {b. [([(c,b)]\<bullet>a)].([(c,b)]\<bullet>x) \<noteq> [a].x}" by (simp add: abs_fun_pi[OF pt, OF at])
7003308ff73a tuned my last commit
urbanc
parents: 18047
diff changeset
  3227
  moreover
7003308ff73a tuned my last commit
urbanc
parents: 18047
diff changeset
  3228
  have "{b. [([(c,b)]\<bullet>a)].([(c,b)]\<bullet>x) \<noteq> [a].x} \<subseteq> {b. ([(c,b)]\<bullet>x,[(c,b)]\<bullet>a) \<noteq> (x, a)}" by force
7003308ff73a tuned my last commit
urbanc
parents: 18047
diff changeset
  3229
  ultimately have "infinite {b. ([(c,b)]\<bullet>x,[(c,b)]\<bullet>a) \<noteq> (x, a)}" by (simp add: infinite_super)
7003308ff73a tuned my last commit
urbanc
parents: 18047
diff changeset
  3230
  thus "c\<in>(supp (x,a))" by (simp add: supp_def)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3231
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3232
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3233
lemma abs_fun_finite_supp:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3234
  fixes x :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3235
  and   a :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3236
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3237
  and     at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3238
  and     f:  "finite ((supp x)::'x set)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3239
  shows "finite ((supp ([a].x))::'x set)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3240
proof -
18048
7003308ff73a tuned my last commit
urbanc
parents: 18047
diff changeset
  3241
  from f have "finite ((supp (x,a))::'x set)" by (simp add: supp_prod at_supp[OF at])
7003308ff73a tuned my last commit
urbanc
parents: 18047
diff changeset
  3242
  moreover
7003308ff73a tuned my last commit
urbanc
parents: 18047
diff changeset
  3243
  have "((supp ([a].x))::'x set) \<subseteq> (supp (x,a))" by (rule abs_fun_supp_approx[OF pt, OF at])
7003308ff73a tuned my last commit
urbanc
parents: 18047
diff changeset
  3244
  ultimately show ?thesis by (simp add: finite_subset)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3245
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3246
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3247
lemma fresh_abs_funI1:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3248
  fixes  x :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3249
  and    a :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3250
  and    b :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3251
  assumes pt:  "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3252
  and     at:   "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3253
  and f:  "finite ((supp x)::'x set)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3254
  and a1: "b\<sharp>x" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3255
  and a2: "a\<noteq>b"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3256
  shows "b\<sharp>([a].x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3257
  proof -
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3258
    have "\<exists>c::'x. c\<sharp>(b,a,x,[a].x)" 
21377
c29146dc14f1 replaced exists_fresh lemma with a version formulated with obtains;
urbanc
parents: 21318
diff changeset
  3259
    proof (rule at_exists_fresh'[OF at], auto simp add: supp_prod at_supp[OF at] f)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3260
      show "finite ((supp ([a].x))::'x set)" using f
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  3261
        by (simp add: abs_fun_finite_supp[OF pt, OF at])        
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3262
    qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3263
    then obtain c where fr1: "c\<noteq>b"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3264
                  and   fr2: "c\<noteq>a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3265
                  and   fr3: "c\<sharp>x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3266
                  and   fr4: "c\<sharp>([a].x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3267
                  by (force simp add: fresh_prod at_fresh[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3268
    have e: "[(c,b)]\<bullet>([a].x) = [a].([(c,b)]\<bullet>x)" using a2 fr1 fr2 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3269
      by (force simp add: abs_fun_pi[OF pt, OF at] at_calc[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3270
    from fr4 have "([(c,b)]\<bullet>c)\<sharp> ([(c,b)]\<bullet>([a].x))"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3271
      by (simp add: pt_fresh_bij[OF pt_abs_fun_inst[OF pt, OF at], OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3272
    hence "b\<sharp>([a].([(c,b)]\<bullet>x))" using fr1 fr2 e  
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3273
      by (simp add: at_calc[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3274
    thus ?thesis using a1 fr3 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3275
      by (simp add: pt_fresh_fresh[OF pt, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3276
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3277
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3278
lemma fresh_abs_funE:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3279
  fixes a :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3280
  and   b :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3281
  and   x :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3282
  assumes pt:  "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3283
  and     at:  "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3284
  and     f:  "finite ((supp x)::'x set)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3285
  and     a1: "b\<sharp>([a].x)" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3286
  and     a2: "b\<noteq>a" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3287
  shows "b\<sharp>x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3288
proof -
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3289
  have "\<exists>c::'x. c\<sharp>(b,a,x,[a].x)"
21377
c29146dc14f1 replaced exists_fresh lemma with a version formulated with obtains;
urbanc
parents: 21318
diff changeset
  3290
  proof (rule at_exists_fresh'[OF at], auto simp add: supp_prod at_supp[OF at] f)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3291
    show "finite ((supp ([a].x))::'x set)" using f
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32638
diff changeset
  3292
      by (simp add: abs_fun_finite_supp[OF pt, OF at])  
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3293
  qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3294
  then obtain c where fr1: "b\<noteq>c"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3295
                and   fr2: "c\<noteq>a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3296
                and   fr3: "c\<sharp>x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3297
                and   fr4: "c\<sharp>([a].x)" by (force simp add: fresh_prod at_fresh[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3298
  have "[a].x = [(b,c)]\<bullet>([a].x)" using a1 fr4 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3299
    by (simp add: pt_fresh_fresh[OF pt_abs_fun_inst[OF pt, OF at], OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3300
  hence "[a].x = [a].([(b,c)]\<bullet>x)" using fr2 a2 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3301
    by (force simp add: abs_fun_pi[OF pt, OF at] at_calc[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3302
  hence b: "([(b,c)]\<bullet>x) = x" by (simp add: abs_fun_eq1)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3303
  from fr3 have "([(b,c)]\<bullet>c)\<sharp>([(b,c)]\<bullet>x)" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3304
    by (simp add: pt_fresh_bij[OF pt, OF at]) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3305
  thus ?thesis using b fr1 by (simp add: at_calc[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3306
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3307
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3308
lemma fresh_abs_funI2:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3309
  fixes a :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3310
  and   x :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3311
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3312
  and     at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3313
  and     f: "finite ((supp x)::'x set)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3314
  shows "a\<sharp>([a].x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3315
proof -
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3316
  have "\<exists>c::'x. c\<sharp>(a,x)"
21377
c29146dc14f1 replaced exists_fresh lemma with a version formulated with obtains;
urbanc
parents: 21318
diff changeset
  3317
    by  (rule at_exists_fresh'[OF at], auto simp add: supp_prod at_supp[OF at] f) 
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3318
  then obtain c where fr1: "a\<noteq>c" and fr1_sym: "c\<noteq>a" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3319
                and   fr2: "c\<sharp>x" by (force simp add: fresh_prod at_fresh[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3320
  have "c\<sharp>([a].x)" using f fr1 fr2 by (simp add: fresh_abs_funI1[OF pt, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3321
  hence "([(c,a)]\<bullet>c)\<sharp>([(c,a)]\<bullet>([a].x))" using fr1  
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3322
    by (simp only: pt_fresh_bij[OF pt_abs_fun_inst[OF pt, OF at], OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3323
  hence a: "a\<sharp>([c].([(c,a)]\<bullet>x))" using fr1_sym 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3324
    by (simp add: abs_fun_pi[OF pt, OF at] at_calc[OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3325
  have "[c].([(c,a)]\<bullet>x) = ([a].x)" using fr1_sym fr2 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3326
    by (simp add: abs_fun_eq[OF pt, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3327
  thus ?thesis using a by simp
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3328
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3329
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3330
lemma fresh_abs_fun_iff: 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3331
  fixes a :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3332
  and   b :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3333
  and   x :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3334
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3335
  and     at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3336
  and     f: "finite ((supp x)::'x set)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3337
  shows "(b\<sharp>([a].x)) = (b=a \<or> b\<sharp>x)" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3338
  by (auto  dest: fresh_abs_funE[OF pt, OF at,OF f] 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3339
           intro: fresh_abs_funI1[OF pt, OF at,OF f] 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3340
                  fresh_abs_funI2[OF pt, OF at,OF f])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3341
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3342
lemma abs_fun_supp: 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3343
  fixes a :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3344
  and   x :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3345
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3346
  and     at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3347
  and     f: "finite ((supp x)::'x set)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3348
  shows "supp ([a].x) = (supp x)-{a}"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3349
 by (force simp add: supp_fresh_iff fresh_abs_fun_iff[OF pt, OF at, OF f])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3350
18048
7003308ff73a tuned my last commit
urbanc
parents: 18047
diff changeset
  3351
(* maybe needs to be better stated as supp intersection supp *)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3352
lemma abs_fun_supp_ineq: 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3353
  fixes a :: "'y"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3354
  and   x :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3355
  assumes pta: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3356
  and     ptb: "pt TYPE('y) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3357
  and     at:  "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3358
  and     cp:  "cp TYPE('a) TYPE('x) TYPE('y)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3359
  and     dj:  "disjoint TYPE('y) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3360
  shows "((supp ([a].x))::'x set) = (supp x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3361
apply(auto simp add: supp_def)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3362
apply(auto simp add: abs_fun_pi_ineq[OF pta, OF ptb, OF at, OF cp])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3363
apply(auto simp add: dj_perm_forget[OF dj])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3364
apply(auto simp add: abs_fun_eq1) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3365
done
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3366
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3367
lemma fresh_abs_fun_iff_ineq: 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3368
  fixes a :: "'y"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3369
  and   b :: "'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3370
  and   x :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3371
  assumes pta: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3372
  and     ptb: "pt TYPE('y) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3373
  and     at:  "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3374
  and     cp:  "cp TYPE('a) TYPE('x) TYPE('y)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3375
  and     dj:  "disjoint TYPE('y) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3376
  shows "b\<sharp>([a].x) = b\<sharp>x" 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3377
  by (simp add: fresh_def abs_fun_supp_ineq[OF pta, OF ptb, OF at, OF cp, OF dj])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3378
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  3379
section \<open>abstraction type for the parsing in nominal datatype\<close>
18048
7003308ff73a tuned my last commit
urbanc
parents: 18047
diff changeset
  3380
(*==============================================================*)
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 23393
diff changeset
  3381
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 23393
diff changeset
  3382
inductive_set ABS_set :: "('x\<Rightarrow>('a noption)) set"
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 23393
diff changeset
  3383
  where
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3384
  ABS_in: "(abs_fun a x)\<in>ABS_set"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3385
45694
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 45625
diff changeset
  3386
definition "ABS = ABS_set"
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 45625
diff changeset
  3387
61260
e6f03fae14d5 explicit indication of overloaded typedefs;
wenzelm
parents: 60585
diff changeset
  3388
typedef ('x, 'a) ABS ("\<guillemotleft>_\<guillemotright>_" [1000,1000] 1000) =
45694
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 45625
diff changeset
  3389
    "ABS::('x\<Rightarrow>('a noption)) set"
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 45625
diff changeset
  3390
  morphisms Rep_ABS Abs_ABS
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 45625
diff changeset
  3391
  unfolding ABS_def
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3392
proof 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3393
  fix x::"'a" and a::"'x"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3394
  show "(abs_fun a x)\<in> ABS_set" by (rule ABS_in)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3395
qed
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3396
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3397
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  3398
section \<open>lemmas for deciding permutation equations\<close>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3399
(*===================================================*)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3400
19477
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  3401
lemma perm_aux_fold:
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  3402
  shows "perm_aux pi x = pi\<bullet>x" by (simp only: perm_aux_def)
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  3403
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  3404
lemma pt_perm_compose_aux:
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  3405
  fixes pi1 :: "'x prm"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  3406
  and   pi2 :: "'x prm"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  3407
  and   x  :: "'a"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  3408
  assumes pt: "pt TYPE('a) TYPE('x)"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  3409
  and     at: "at TYPE('x)"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  3410
  shows "pi2\<bullet>(pi1\<bullet>x) = perm_aux (pi2\<bullet>pi1) (pi2\<bullet>x)" 
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  3411
proof -
23393
31781b2de73d tuned proofs: avoid implicit prems;
wenzelm
parents: 23159
diff changeset
  3412
  have "(pi2@pi1) \<triangleq> ((pi2\<bullet>pi1)@pi2)" by (rule at_ds8[OF at])
19477
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  3413
  hence "(pi2@pi1)\<bullet>x = ((pi2\<bullet>pi1)@pi2)\<bullet>x" by (rule pt3[OF pt])
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  3414
  thus ?thesis by (simp add: pt2[OF pt] perm_aux_def)
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  3415
qed  
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  3416
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  3417
lemma cp1_aux:
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  3418
  fixes pi1::"'x prm"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  3419
  and   pi2::"'y prm"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  3420
  and   x  ::"'a"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  3421
  assumes cp: "cp TYPE ('a) TYPE('x) TYPE('y)"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  3422
  shows "pi1\<bullet>(pi2\<bullet>x) = perm_aux (pi1\<bullet>pi2) (pi1\<bullet>x)"
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  3423
  using cp by (simp add: cp_def perm_aux_def)
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  3424
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3425
lemma perm_eq_app:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3426
  fixes f  :: "'a\<Rightarrow>'b"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3427
  and   x  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3428
  and   pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3429
  assumes pt: "pt TYPE('a) TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3430
  and     at: "at TYPE('x)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3431
  shows "(pi\<bullet>(f x)=y) = ((pi\<bullet>f)(pi\<bullet>x)=y)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3432
  by (simp add: pt_fun_app_eq[OF pt, OF at])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3433
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3434
lemma perm_eq_lam:
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3435
  fixes f  :: "'a\<Rightarrow>'b"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3436
  and   x  :: "'a"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3437
  and   pi :: "'x prm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3438
  shows "((pi\<bullet>(\<lambda>x. f x))=y) = ((\<lambda>x. (pi\<bullet>(f ((rev pi)\<bullet>x))))=y)"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3439
  by (simp add: perm_fun_def)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3440
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  3441
section \<open>test\<close>
19132
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
  3442
lemma at_prm_eq_compose:
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
  3443
  fixes pi1 :: "'x prm"
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
  3444
  and   pi2 :: "'x prm"
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
  3445
  and   pi3 :: "'x prm"
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
  3446
  assumes at: "at TYPE('x)"
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
  3447
  and     a: "pi1 \<triangleq> pi2"
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
  3448
  shows "(pi3\<bullet>pi1) \<triangleq> (pi3\<bullet>pi2)"
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
  3449
proof -
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
  3450
  have pt: "pt TYPE('x) TYPE('x)" by (rule at_pt_inst[OF at])
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
  3451
  have pt_prm: "pt TYPE('x prm) TYPE('x)" 
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
  3452
    by (rule pt_list_inst[OF pt_prod_inst[OF pt, OF pt]])  
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
  3453
  from a show ?thesis
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
  3454
    apply -
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
  3455
    apply(auto simp add: prm_eq_def)
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
  3456
    apply(rule_tac pi="rev pi3" in pt_bij4[OF pt, OF at])
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
  3457
    apply(rule trans)
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
  3458
    apply(rule pt_perm_compose[OF pt, OF at])
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
  3459
    apply(simp add: pt_rev_pi[OF pt_prm, OF at])
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
  3460
    apply(rule sym)
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
  3461
    apply(rule trans)
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
  3462
    apply(rule pt_perm_compose[OF pt, OF at])
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
  3463
    apply(simp add: pt_rev_pi[OF pt_prm, OF at])
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
  3464
    done
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
  3465
qed
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
  3466
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3467
(************************)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3468
(* Various eqvt-lemmas  *)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3469
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3470
lemma Zero_nat_eqvt:
22446
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3471
  shows "pi\<bullet>(0::nat) = 0" 
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3472
by (auto simp add: perm_nat_def)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3473
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3474
lemma One_nat_eqvt:
22446
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3475
  shows "pi\<bullet>(1::nat) = 1"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3476
by (simp add: perm_nat_def)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3477
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3478
lemma Suc_eqvt:
22446
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3479
  shows "pi\<bullet>(Suc x) = Suc (pi\<bullet>x)" 
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3480
by (auto simp add: perm_nat_def)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3481
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3482
lemma numeral_nat_eqvt: 
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46950
diff changeset
  3483
 shows "pi\<bullet>((numeral n)::nat) = numeral n" 
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3484
by (simp add: perm_nat_def perm_int_def)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3485
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3486
lemma max_nat_eqvt:
22446
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3487
  fixes x::"nat"
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3488
  shows "pi\<bullet>(max x y) = max (pi\<bullet>x) (pi\<bullet>y)" 
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3489
by (simp add:perm_nat_def) 
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3490
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3491
lemma min_nat_eqvt:
22446
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3492
  fixes x::"nat"
22500
8436bfd21bf3 corrected the lemmas min_nat_eqvt and min_int_eqvt
urbanc
parents: 22446
diff changeset
  3493
  shows "pi\<bullet>(min x y) = min (pi\<bullet>x) (pi\<bullet>y)" 
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3494
by (simp add:perm_nat_def) 
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3495
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3496
lemma plus_nat_eqvt:
22446
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3497
  fixes x::"nat"
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3498
  shows "pi\<bullet>(x + y) = (pi\<bullet>x) + (pi\<bullet>y)" 
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3499
by (simp add:perm_nat_def) 
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3500
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3501
lemma minus_nat_eqvt:
22446
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3502
  fixes x::"nat"
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3503
  shows "pi\<bullet>(x - y) = (pi\<bullet>x) - (pi\<bullet>y)" 
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3504
by (simp add:perm_nat_def) 
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3505
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3506
lemma mult_nat_eqvt:
22446
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3507
  fixes x::"nat"
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3508
  shows "pi\<bullet>(x * y) = (pi\<bullet>x) * (pi\<bullet>y)" 
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3509
by (simp add:perm_nat_def) 
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3510
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3511
lemma div_nat_eqvt:
22446
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3512
  fixes x::"nat"
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3513
  shows "pi\<bullet>(x div y) = (pi\<bullet>x) div (pi\<bullet>y)" 
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3514
by (simp add:perm_nat_def) 
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3515
22446
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3516
lemma Zero_int_eqvt:
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3517
  shows "pi\<bullet>(0::int) = 0" 
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3518
by (auto simp add: perm_int_def)
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3519
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3520
lemma One_int_eqvt:
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3521
  shows "pi\<bullet>(1::int) = 1"
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3522
by (simp add: perm_int_def)
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3523
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3524
lemma numeral_int_eqvt: 
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46950
diff changeset
  3525
 shows "pi\<bullet>((numeral n)::int) = numeral n" 
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46950
diff changeset
  3526
by (simp add: perm_int_def perm_int_def)
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46950
diff changeset
  3527
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46950
diff changeset
  3528
lemma neg_numeral_int_eqvt:
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 49834
diff changeset
  3529
 shows "pi\<bullet>((- numeral n)::int) = - numeral n"
22446
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3530
by (simp add: perm_int_def perm_int_def)
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3531
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3532
lemma max_int_eqvt:
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3533
  fixes x::"int"
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3534
  shows "pi\<bullet>(max (x::int) y) = max (pi\<bullet>x) (pi\<bullet>y)" 
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3535
by (simp add:perm_int_def) 
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3536
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3537
lemma min_int_eqvt:
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3538
  fixes x::"int"
22500
8436bfd21bf3 corrected the lemmas min_nat_eqvt and min_int_eqvt
urbanc
parents: 22446
diff changeset
  3539
  shows "pi\<bullet>(min x y) = min (pi\<bullet>x) (pi\<bullet>y)" 
22446
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3540
by (simp add:perm_int_def) 
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3541
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3542
lemma plus_int_eqvt:
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3543
  fixes x::"int"
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3544
  shows "pi\<bullet>(x + y) = (pi\<bullet>x) + (pi\<bullet>y)" 
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3545
by (simp add:perm_int_def) 
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3546
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3547
lemma minus_int_eqvt:
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3548
  fixes x::"int"
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3549
  shows "pi\<bullet>(x - y) = (pi\<bullet>x) - (pi\<bullet>y)" 
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3550
by (simp add:perm_int_def) 
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3551
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3552
lemma mult_int_eqvt:
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3553
  fixes x::"int"
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3554
  shows "pi\<bullet>(x * y) = (pi\<bullet>x) * (pi\<bullet>y)" 
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3555
by (simp add:perm_int_def) 
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3556
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3557
lemma div_int_eqvt:
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3558
  fixes x::"int"
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3559
  shows "pi\<bullet>(x div y) = (pi\<bullet>x) div (pi\<bullet>y)" 
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3560
by (simp add:perm_int_def) 
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3561
30990
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  3562
(*******************************************************)
4872eef36167 reorganised the section about fresh_star and added lemma pt_fresh_star_pi
Christian Urban <urbanc@in.tum.de>
parents: 30983
diff changeset
  3563
(* Setup of the theorem attributes eqvt and eqvt_force *)
48891
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 47108
diff changeset
  3564
ML_file "nominal_thmdecls.ML"
22245
1b8f4ef50c48 moved the infrastructure from the nominal_tags file to nominal_thmdecls
urbanc
parents: 22231
diff changeset
  3565
setup "NominalThmDecls.setup"
19132
ff41946e5092 added lemmas
urbanc
parents: 19107
diff changeset
  3566
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3567
lemmas [eqvt] = 
22446
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3568
  (* connectives *)
22732
5bd1a2a94e1b declared lemmas true_eqvt and false_eqvt to be equivariant (suggested by samth at ccs.neu.edu)
urbanc
parents: 22729
diff changeset
  3569
  if_eqvt imp_eqvt disj_eqvt conj_eqvt neg_eqvt 
5bd1a2a94e1b declared lemmas true_eqvt and false_eqvt to be equivariant (suggested by samth at ccs.neu.edu)
urbanc
parents: 22729
diff changeset
  3570
  true_eqvt false_eqvt
59940
087d81f5213e local setup of induction tools, with restricted access to auxiliary consts;
wenzelm
parents: 58372
diff changeset
  3571
  imp_eqvt [folded HOL.induct_implies_def]
22446
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3572
  
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3573
  (* datatypes *)
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3574
  perm_unit.simps
22446
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3575
  perm_list.simps append_eqvt
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3576
  perm_prod.simps
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3577
  fst_eqvt snd_eqvt
22511
ca326e0fb5c5 added the permutation operation on options to the list of equivariance lemmas
urbanc
parents: 22500
diff changeset
  3578
  perm_option.simps
22446
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3579
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3580
  (* nats *)
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3581
  Suc_eqvt Zero_nat_eqvt One_nat_eqvt min_nat_eqvt max_nat_eqvt
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3582
  plus_nat_eqvt minus_nat_eqvt mult_nat_eqvt div_nat_eqvt
22446
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3583
  
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3584
  (* ints *)
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3585
  Zero_int_eqvt One_int_eqvt min_int_eqvt max_int_eqvt
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3586
  plus_int_eqvt minus_int_eqvt mult_int_eqvt div_int_eqvt
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3587
  
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3588
  (* sets *)
46179
47bcf3d5d1f0 Reverted several lemmas involving sets to the state before the
berghofe
parents: 45961
diff changeset
  3589
  union_eqvt empty_eqvt insert_eqvt set_eqvt
22446
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3590
  
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3591
 
22446
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3592
(* the lemmas numeral_nat_eqvt numeral_int_eqvt do not conform with the *)
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3593
(* usual form of an eqvt-lemma, but they are needed for analysing       *)
91951d4177d3 added eqvt-lemmas for integers and eqvt-tagged the lemma append_eqvt
urbanc
parents: 22418
diff changeset
  3594
(* permutations on nats and ints *)
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46950
diff changeset
  3595
lemmas [eqvt_force] = numeral_nat_eqvt numeral_int_eqvt neg_numeral_int_eqvt
22326
a3acee47a883 start adding the attribute eqvt to some lemmas of the nominal library
narboux
parents: 22312
diff changeset
  3596
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3597
(***************************************)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3598
(* setup for the individial atom-kinds *)
18047
3d643b13eb65 simplified the abs_supp_approx proof and tuned some comments in
urbanc
parents: 18012
diff changeset
  3599
(* and nominal datatypes               *)
48891
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 47108
diff changeset
  3600
ML_file "nominal_atoms.ML"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3601
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3602
(************************************************************)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3603
(* various tactics for analysing permutations, supports etc *)
48891
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 47108
diff changeset
  3604
ML_file "nominal_permeq.ML"
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3605
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3606
method_setup perm_simp =
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  3607
  \<open>NominalPermeq.perm_simp_meth\<close>
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  3608
  \<open>simp rules and simprocs for analysing permutations\<close>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3609
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3610
method_setup perm_simp_debug =
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  3611
  \<open>NominalPermeq.perm_simp_meth_debug\<close>
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  3612
  \<open>simp rules and simprocs for analysing permutations including debugging facilities\<close>
19477
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  3613
28322
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28011
diff changeset
  3614
method_setup perm_extend_simp =
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  3615
  \<open>NominalPermeq.perm_extend_simp_meth\<close>
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  3616
  \<open>tactic for deciding equalities involving permutations\<close>
19477
a95176d0f0dd isar-keywords.el
urbanc
parents: 19329
diff changeset
  3617
28322
6f4cf302c798 made the perm_simp tactic to understand options such as (no_asm)
urbanc
parents: 28011
diff changeset
  3618
method_setup perm_extend_simp_debug =
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  3619
  \<open>NominalPermeq.perm_extend_simp_meth_debug\<close>
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  3620
  \<open>tactic for deciding equalities involving permutations including debugging facilities\<close>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3621
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3622
method_setup supports_simp =
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  3623
  \<open>NominalPermeq.supports_meth\<close>
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  3624
  \<open>tactic for deciding whether something supports something else\<close>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3625
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3626
method_setup supports_simp_debug =
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  3627
  \<open>NominalPermeq.supports_meth_debug\<close>
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  3628
  \<open>tactic for deciding whether something supports something else including debugging facilities\<close>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3629
19164
0eccb98b1fdb added initialisation-code for finite_guess
urbanc
parents: 19140
diff changeset
  3630
method_setup finite_guess =
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  3631
  \<open>NominalPermeq.finite_guess_meth\<close>
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  3632
  \<open>tactic for deciding whether something has finite support\<close>
19164
0eccb98b1fdb added initialisation-code for finite_guess
urbanc
parents: 19140
diff changeset
  3633
0eccb98b1fdb added initialisation-code for finite_guess
urbanc
parents: 19140
diff changeset
  3634
method_setup finite_guess_debug =
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  3635
  \<open>NominalPermeq.finite_guess_meth_debug\<close>
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  3636
  \<open>tactic for deciding whether something has finite support including debugging facilities\<close>
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19477
diff changeset
  3637
19638
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
  3638
method_setup fresh_guess =
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  3639
  \<open>NominalPermeq.fresh_guess_meth\<close>
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  3640
  \<open>tactic for deciding whether an atom is fresh for something\<close>
19638
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
  3641
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
  3642
method_setup fresh_guess_debug =
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  3643
  \<open>NominalPermeq.fresh_guess_meth_debug\<close>
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  3644
  \<open>tactic for deciding whether an atom is fresh for something including debugging facilities\<close>
19638
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19634
diff changeset
  3645
22762
f28f62754644 tuned the setup of fresh_fun
urbanc
parents: 22732
diff changeset
  3646
(*****************************************************************)
f28f62754644 tuned the setup of fresh_fun
urbanc
parents: 22732
diff changeset
  3647
(* tactics for generating fresh names and simplifying fresh_funs *)
48891
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 47108
diff changeset
  3648
ML_file "nominal_fresh_fun.ML"
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents: 22715
diff changeset
  3649
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  3650
method_setup generate_fresh = \<open>
56230
3e449273942a more standard method_setup;
wenzelm
parents: 54489
diff changeset
  3651
  Args.type_name {proper = true, strict = true} >>
3e449273942a more standard method_setup;
wenzelm
parents: 54489
diff changeset
  3652
    (fn s => fn ctxt => SIMPLE_METHOD (generate_fresh_tac ctxt s))
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  3653
\<close> "generate a name fresh for all the variables in the goal"
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  3654
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  3655
method_setup fresh_fun_simp = \<open>
56230
3e449273942a more standard method_setup;
wenzelm
parents: 54489
diff changeset
  3656
  Scan.lift (Args.parens (Args.$$$ "no_asm") >> K true || Scan.succeed false) >>
3e449273942a more standard method_setup;
wenzelm
parents: 54489
diff changeset
  3657
    (fn b => fn ctxt => SIMPLE_METHOD' (fresh_fun_tac ctxt b))
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  3658
\<close> "delete one inner occurrence of fresh_fun"
22729
69ef734825c5 add a tactic to generate fresh names
narboux
parents: 22715
diff changeset
  3659
69ef734825c5 add a tactic to generate fresh names
narboux
parents: 22715
diff changeset
  3660
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3661
(************************************************)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3662
(* main file for constructing nominal datatypes *)
27228
4f7976a6ffc3 allE_Nil: only one copy, proven in regular theory source;
wenzelm
parents: 26847
diff changeset
  3663
lemma allE_Nil: assumes "\<forall>x. P x" obtains "P []"
4f7976a6ffc3 allE_Nil: only one copy, proven in regular theory source;
wenzelm
parents: 26847
diff changeset
  3664
  using assms ..
4f7976a6ffc3 allE_Nil: only one copy, proven in regular theory source;
wenzelm
parents: 26847
diff changeset
  3665
48891
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 47108
diff changeset
  3666
ML_file "nominal_datatype.ML"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3667
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3668
(******************************************************)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3669
(* primitive recursive functions on nominal datatypes *)
48891
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 47108
diff changeset
  3670
ML_file "nominal_primrec.ML"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3671
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3672
(****************************************************)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3673
(* inductive definition involving nominal datatypes *)
48891
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 47108
diff changeset
  3674
ML_file "nominal_inductive.ML"
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 47108
diff changeset
  3675
ML_file "nominal_inductive2.ML"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3676
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3677
(*****************************************)
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3678
(* setup for induction principles method *)
48891
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 47108
diff changeset
  3679
ML_file "nominal_induct.ML"
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3680
method_setup nominal_induct =
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  3681
  \<open>NominalInduct.nominal_induct_method\<close>
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62145
diff changeset
  3682
  \<open>nominal induction\<close>
22418
49e2d9744ae1 major update of the nominal package; there is now an infrastructure
urbanc
parents: 22326
diff changeset
  3683
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  3684
end