Nominal-General/Nominal2_Eqvt.thy
author Christian Urban <urbanc@in.tum.de>
Wed, 28 Apr 2010 08:22:20 +0200
changeset 1971 8daf6ff5e11a
parent 1962 84a13d1e2511
child 1995 652f310f0dba
permissions -rw-r--r--
simpliied and moved the remaining lemmas about the atom-function to Nominal2_Base
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
     1
(*  Title:      Nominal2_Eqvt
1810
894930834ca8 fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1803
diff changeset
     2
    Author:     Brian Huffman, 
894930834ca8 fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1803
diff changeset
     3
    Author:     Christian Urban
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
     4
1869
Christian Urban <urbanc@in.tum.de>
parents: 1867
diff changeset
     5
    Equivariance, supp and freshness lemmas for various operators 
Christian Urban <urbanc@in.tum.de>
parents: 1867
diff changeset
     6
    (contains many, but not all such lemmas).
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
     7
*)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
     8
theory Nominal2_Eqvt
1933
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1872
diff changeset
     9
imports Nominal2_Base 
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    10
uses ("nominal_thmdecls.ML")
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    11
     ("nominal_permeq.ML")
1833
2050b5723c04 added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents: 1827
diff changeset
    12
     ("nominal_eqvt.ML")
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    13
begin
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    14
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    15
section {* Logical Operators *}
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    16
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    17
lemma eq_eqvt:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    18
  shows "p \<bullet> (x = y) \<longleftrightarrow> (p \<bullet> x) = (p \<bullet> y)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    19
  unfolding permute_eq_iff permute_bool_def ..
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    20
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    21
lemma if_eqvt:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    22
  shows "p \<bullet> (if b then x else y) = (if p \<bullet> b then p \<bullet> x else p \<bullet> y)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    23
  by (simp add: permute_fun_def permute_bool_def)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    24
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    25
lemma True_eqvt:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    26
  shows "p \<bullet> True = True"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    27
  unfolding permute_bool_def ..
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    28
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    29
lemma False_eqvt:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    30
  shows "p \<bullet> False = False"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    31
  unfolding permute_bool_def ..
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    32
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    33
lemma imp_eqvt:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    34
  shows "p \<bullet> (A \<longrightarrow> B) = ((p \<bullet> A) \<longrightarrow> (p \<bullet> B))"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    35
  by (simp add: permute_bool_def)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    36
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    37
lemma conj_eqvt:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    38
  shows "p \<bullet> (A \<and> B) = ((p \<bullet> A) \<and> (p \<bullet> B))"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    39
  by (simp add: permute_bool_def)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    40
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    41
lemma disj_eqvt:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    42
  shows "p \<bullet> (A \<or> B) = ((p \<bullet> A) \<or> (p \<bullet> B))"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    43
  by (simp add: permute_bool_def)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    44
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    45
lemma Not_eqvt:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    46
  shows "p \<bullet> (\<not> A) = (\<not> (p \<bullet> A))"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    47
  by (simp add: permute_bool_def)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    48
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    49
lemma all_eqvt:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    50
  shows "p \<bullet> (\<forall>x. P x) = (\<forall>x. (p \<bullet> P) x)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    51
  unfolding permute_fun_def permute_bool_def
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    52
  by (auto, drule_tac x="p \<bullet> x" in spec, simp)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    53
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    54
lemma all_eqvt2:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    55
  shows "p \<bullet> (\<forall>x. P x) = (\<forall>x. p \<bullet> P (- p \<bullet> x))"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    56
  unfolding permute_fun_def permute_bool_def
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    57
  by (auto, drule_tac x="p \<bullet> x" in spec, simp)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    58
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    59
lemma ex_eqvt:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    60
  shows "p \<bullet> (\<exists>x. P x) = (\<exists>x. (p \<bullet> P) x)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    61
  unfolding permute_fun_def permute_bool_def
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    62
  by (auto, rule_tac x="p \<bullet> x" in exI, simp)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    63
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    64
lemma ex_eqvt2:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    65
  shows "p \<bullet> (\<exists>x. P x) = (\<exists>x. p \<bullet> P (- p \<bullet> x))"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    66
  unfolding permute_fun_def permute_bool_def
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    67
  by (auto, rule_tac x="p \<bullet> x" in exI, simp)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    68
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    69
lemma ex1_eqvt:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    70
  shows "p \<bullet> (\<exists>!x. P x) = (\<exists>!x. (p \<bullet> P) x)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    71
  unfolding Ex1_def 
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    72
  by (simp add: ex_eqvt permute_fun_def conj_eqvt all_eqvt imp_eqvt eq_eqvt)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    73
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    74
lemma ex1_eqvt2:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    75
  shows "p \<bullet> (\<exists>!x. P x) = (\<exists>!x. p \<bullet> P (- p \<bullet> x))"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    76
  unfolding Ex1_def ex_eqvt2 conj_eqvt all_eqvt2 imp_eqvt eq_eqvt
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    77
  by simp
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    78
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    79
lemma the_eqvt:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    80
  assumes unique: "\<exists>!x. P x"
1087
bb7f4457091a moved some lemmas to Nominal; updated all files
Christian Urban <urbanc@in.tum.de>
parents: 1079
diff changeset
    81
  shows "(p \<bullet> (THE x. P x)) = (THE x. p \<bullet> P (- p \<bullet> x))"
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    82
  apply(rule the1_equality [symmetric])
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    83
  apply(simp add: ex1_eqvt2[symmetric])
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    84
  apply(simp add: permute_bool_def unique)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    85
  apply(simp add: permute_bool_def)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    86
  apply(rule theI'[OF unique])
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    87
  done
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    88
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    89
section {* Set Operations *}
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    90
1087
bb7f4457091a moved some lemmas to Nominal; updated all files
Christian Urban <urbanc@in.tum.de>
parents: 1079
diff changeset
    91
lemma mem_permute_iff:
bb7f4457091a moved some lemmas to Nominal; updated all files
Christian Urban <urbanc@in.tum.de>
parents: 1079
diff changeset
    92
  shows "(p \<bullet> x) \<in> (p \<bullet> X) \<longleftrightarrow> x \<in> X"
bb7f4457091a moved some lemmas to Nominal; updated all files
Christian Urban <urbanc@in.tum.de>
parents: 1079
diff changeset
    93
unfolding mem_def permute_fun_def permute_bool_def
bb7f4457091a moved some lemmas to Nominal; updated all files
Christian Urban <urbanc@in.tum.de>
parents: 1079
diff changeset
    94
by simp
bb7f4457091a moved some lemmas to Nominal; updated all files
Christian Urban <urbanc@in.tum.de>
parents: 1079
diff changeset
    95
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    96
lemma mem_eqvt:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    97
  shows "p \<bullet> (x \<in> A) \<longleftrightarrow> (p \<bullet> x) \<in> (p \<bullet> A)"
1087
bb7f4457091a moved some lemmas to Nominal; updated all files
Christian Urban <urbanc@in.tum.de>
parents: 1079
diff changeset
    98
  unfolding mem_permute_iff permute_bool_def by simp
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    99
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   100
lemma not_mem_eqvt:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   101
  shows "p \<bullet> (x \<notin> A) \<longleftrightarrow> (p \<bullet> x) \<notin> (p \<bullet> A)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   102
  unfolding mem_def permute_fun_def by (simp add: Not_eqvt)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   103
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   104
lemma Collect_eqvt:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   105
  shows "p \<bullet> {x. P x} = {x. (p \<bullet> P) x}"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   106
  unfolding Collect_def permute_fun_def ..
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   107
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   108
lemma Collect_eqvt2:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   109
  shows "p \<bullet> {x. P x} = {x. p \<bullet> (P (-p \<bullet> x))}"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   110
  unfolding Collect_def permute_fun_def ..
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   111
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   112
lemma empty_eqvt:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   113
  shows "p \<bullet> {} = {}"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   114
  unfolding empty_def Collect_eqvt2 False_eqvt ..
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   115
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   116
lemma supp_set_empty:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   117
  shows "supp {} = {}"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   118
  by (simp add: supp_def empty_eqvt)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   119
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   120
lemma fresh_set_empty:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   121
  shows "a \<sharp> {}"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   122
  by (simp add: fresh_def supp_set_empty)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   123
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   124
lemma UNIV_eqvt:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   125
  shows "p \<bullet> UNIV = UNIV"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   126
  unfolding UNIV_def Collect_eqvt2 True_eqvt ..
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   127
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   128
lemma union_eqvt:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   129
  shows "p \<bullet> (A \<union> B) = (p \<bullet> A) \<union> (p \<bullet> B)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   130
  unfolding Un_def Collect_eqvt2 disj_eqvt mem_eqvt by simp
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   131
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   132
lemma inter_eqvt:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   133
  shows "p \<bullet> (A \<inter> B) = (p \<bullet> A) \<inter> (p \<bullet> B)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   134
  unfolding Int_def Collect_eqvt2 conj_eqvt mem_eqvt by simp
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   135
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   136
lemma Diff_eqvt:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   137
  fixes A B :: "'a::pt set"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   138
  shows "p \<bullet> (A - B) = p \<bullet> A - p \<bullet> B"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   139
  unfolding set_diff_eq Collect_eqvt2 conj_eqvt Not_eqvt mem_eqvt by simp
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   140
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   141
lemma Compl_eqvt:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   142
  fixes A :: "'a::pt set"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   143
  shows "p \<bullet> (- A) = - (p \<bullet> A)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   144
  unfolding Compl_eq_Diff_UNIV Diff_eqvt UNIV_eqvt ..
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   145
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   146
lemma insert_eqvt:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   147
  shows "p \<bullet> (insert x A) = insert (p \<bullet> x) (p \<bullet> A)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   148
  unfolding permute_set_eq_image image_insert ..
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   149
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   150
lemma vimage_eqvt:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   151
  shows "p \<bullet> (f -` A) = (p \<bullet> f) -` (p \<bullet> A)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   152
  unfolding vimage_def permute_fun_def [where f=f]
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   153
  unfolding Collect_eqvt2 mem_eqvt ..
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   154
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   155
lemma finite_permute_iff:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   156
  shows "finite (p \<bullet> A) \<longleftrightarrow> finite A"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   157
  unfolding permute_set_eq_vimage
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   158
  using bij_permute by (rule finite_vimage_iff)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   159
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   160
lemma finite_eqvt:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   161
  shows "p \<bullet> finite A = finite (p \<bullet> A)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   162
  unfolding finite_permute_iff permute_bool_def ..
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   163
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   164
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   165
section {* List Operations *}
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   166
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   167
lemma append_eqvt:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   168
  shows "p \<bullet> (xs @ ys) = (p \<bullet> xs) @ (p \<bullet> ys)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   169
  by (induct xs) auto
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   170
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   171
lemma supp_append:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   172
  shows "supp (xs @ ys) = supp xs \<union> supp ys"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   173
  by (induct xs) (auto simp add: supp_Nil supp_Cons)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   174
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   175
lemma fresh_append:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   176
  shows "a \<sharp> (xs @ ys) \<longleftrightarrow> a \<sharp> xs \<and> a \<sharp> ys"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   177
  by (induct xs) (simp_all add: fresh_Nil fresh_Cons)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   178
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   179
lemma rev_eqvt:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   180
  shows "p \<bullet> (rev xs) = rev (p \<bullet> xs)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   181
  by (induct xs) (simp_all add: append_eqvt)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   182
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   183
lemma supp_rev:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   184
  shows "supp (rev xs) = supp xs"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   185
  by (induct xs) (auto simp add: supp_append supp_Cons supp_Nil)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   186
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   187
lemma fresh_rev:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   188
  shows "a \<sharp> rev xs \<longleftrightarrow> a \<sharp> xs"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   189
  by (induct xs) (auto simp add: fresh_append fresh_Cons fresh_Nil)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   190
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   191
lemma set_eqvt:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   192
  shows "p \<bullet> (set xs) = set (p \<bullet> xs)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   193
  by (induct xs) (simp_all add: empty_eqvt insert_eqvt)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   194
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   195
(* needs finite support premise
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   196
lemma supp_set:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   197
  fixes x :: "'a::pt"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   198
  shows "supp (set xs) = supp xs"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   199
*)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   200
1815
4135198bbb8a moved equivariance of map into Nominal2_Eqvt file
Christian Urban <urbanc@in.tum.de>
parents: 1811
diff changeset
   201
lemma map_eqvt: 
4135198bbb8a moved equivariance of map into Nominal2_Eqvt file
Christian Urban <urbanc@in.tum.de>
parents: 1811
diff changeset
   202
  shows "p \<bullet> (map f xs) = map (p \<bullet> f) (p \<bullet> xs)"
4135198bbb8a moved equivariance of map into Nominal2_Eqvt file
Christian Urban <urbanc@in.tum.de>
parents: 1811
diff changeset
   203
  by (induct xs) (simp_all, simp only: permute_fun_app_eq)
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   204
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   205
section {* Product Operations *}
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   206
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   207
lemma fst_eqvt:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   208
  "p \<bullet> (fst x) = fst (p \<bullet> x)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   209
 by (cases x) simp
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   210
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   211
lemma snd_eqvt:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   212
  "p \<bullet> (snd x) = snd (p \<bullet> x)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   213
 by (cases x) simp
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   214
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   215
section {* Units *}
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   216
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   217
lemma supp_unit:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   218
  shows "supp () = {}"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   219
  by (simp add: supp_def)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   220
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   221
lemma fresh_unit:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   222
  shows "a \<sharp> ()"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   223
  by (simp add: fresh_def supp_unit)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   224
1800
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   225
lemma permute_eqvt_raw:
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   226
  shows "p \<bullet> permute = permute"
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   227
apply(simp add: expand_fun_eq permute_fun_def)
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   228
apply(subst permute_eqvt)
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   229
apply(simp)
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   230
done
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   231
1953
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
   232
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   233
section {* Equivariance automation *}
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   234
1869
Christian Urban <urbanc@in.tum.de>
parents: 1867
diff changeset
   235
text {* Setup of the theorem attributes @{text eqvt} and @{text eqvt_raw} *}
Christian Urban <urbanc@in.tum.de>
parents: 1867
diff changeset
   236
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   237
use "nominal_thmdecls.ML"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   238
setup "Nominal_ThmDecls.setup"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   239
1953
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
   240
ML {* Thm.full_rules *}
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   241
lemmas [eqvt] = 
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   242
  (* connectives *)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   243
  eq_eqvt if_eqvt imp_eqvt disj_eqvt conj_eqvt Not_eqvt 
1087
bb7f4457091a moved some lemmas to Nominal; updated all files
Christian Urban <urbanc@in.tum.de>
parents: 1079
diff changeset
   244
  True_eqvt False_eqvt ex_eqvt all_eqvt ex1_eqvt
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   245
  imp_eqvt [folded induct_implies_def]
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   246
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   247
  (* nominal *)
1962
84a13d1e2511 moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
Christian Urban <urbanc@in.tum.de>
parents: 1953
diff changeset
   248
  supp_eqvt fresh_eqvt add_perm_eqvt atom_eqvt
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   249
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   250
  (* datatypes *)
1087
bb7f4457091a moved some lemmas to Nominal; updated all files
Christian Urban <urbanc@in.tum.de>
parents: 1079
diff changeset
   251
  permute_prod.simps append_eqvt rev_eqvt set_eqvt
1815
4135198bbb8a moved equivariance of map into Nominal2_Eqvt file
Christian Urban <urbanc@in.tum.de>
parents: 1811
diff changeset
   252
  map_eqvt fst_eqvt snd_eqvt Pair_eqvt permute_list.simps
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   253
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   254
  (* sets *)
1087
bb7f4457091a moved some lemmas to Nominal; updated all files
Christian Urban <urbanc@in.tum.de>
parents: 1079
diff changeset
   255
  empty_eqvt UNIV_eqvt union_eqvt inter_eqvt mem_eqvt
1315
43d6e3730353 Add image_eqvt and atom_eqvt to eqvt bases.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1258
diff changeset
   256
  Diff_eqvt Compl_eqvt insert_eqvt Collect_eqvt image_eqvt
43d6e3730353 Add image_eqvt and atom_eqvt to eqvt bases.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1258
diff changeset
   257
1934
8f6e68dc6cbc tuned parser
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   258
 
1800
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   259
lemmas [eqvt_raw] =
1811
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   260
  permute_eqvt_raw[THEN eq_reflection] (* the normal version of this lemma loops *)
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   261
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   262
text {* helper lemmas for the eqvt_tac *}
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   263
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   264
definition
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   265
  "unpermute p = permute (- p)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   266
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   267
lemma eqvt_apply:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   268
  fixes f :: "'a::pt \<Rightarrow> 'b::pt" 
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   269
  and x :: "'a::pt"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   270
  shows "p \<bullet> (f x) \<equiv> (p \<bullet> f) (p \<bullet> x)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   271
  unfolding permute_fun_def by simp
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   272
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   273
lemma eqvt_lambda:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   274
  fixes f :: "'a::pt \<Rightarrow> 'b::pt"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   275
  shows "p \<bullet> (\<lambda>x. f x) \<equiv> (\<lambda>x. p \<bullet> (f (unpermute p x)))"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   276
  unfolding permute_fun_def unpermute_def by simp
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   277
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   278
lemma eqvt_bound:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   279
  shows "p \<bullet> unpermute p x \<equiv> x"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   280
  unfolding unpermute_def by simp
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   281
1835
636de31888a6 tuned and removed dead code
Christian Urban <urbanc@in.tum.de>
parents: 1833
diff changeset
   282
(* provides perm_simp methods *)
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   283
use "nominal_permeq.ML"
1800
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   284
setup Nominal_Permeq.setup
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   285
1800
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   286
method_setup perm_simp =
1947
51f411b1197d tuned and cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1934
diff changeset
   287
 {* Nominal_Permeq.args_parser >> Nominal_Permeq.perm_simp_meth *}
51f411b1197d tuned and cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1934
diff changeset
   288
 {* pushes permutations inside. *}
1800
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   289
1801
6d2a39db3862 added more robust tracing infrastructure; a strict version of the eqvt_tac raises an error if not all permutations cannot be analysed
Christian Urban <urbanc@in.tum.de>
parents: 1800
diff changeset
   290
method_setup perm_strict_simp =
1947
51f411b1197d tuned and cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1934
diff changeset
   291
 {* Nominal_Permeq.args_parser >> Nominal_Permeq.perm_strict_simp_meth *}
51f411b1197d tuned and cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1934
diff changeset
   292
 {* pushes permutations inside, raises an error if it cannot solve all permutations. *}
1801
6d2a39db3862 added more robust tracing infrastructure; a strict version of the eqvt_tac raises an error if not all permutations cannot be analysed
Christian Urban <urbanc@in.tum.de>
parents: 1800
diff changeset
   293
1800
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   294
declare [[trace_eqvt = true]]
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   295
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   296
lemma 
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   297
  fixes B::"'a::pt"
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   298
  shows "p \<bullet> (B = C)"
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   299
apply(perm_simp)
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   300
oops
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   301
1800
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   302
lemma 
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   303
  fixes B::"bool"
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   304
  shows "p \<bullet> (B = C)"
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   305
apply(perm_simp)
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   306
oops
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   307
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   308
lemma 
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   309
  fixes B::"bool"
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   310
  shows "p \<bullet> (A \<longrightarrow> B = C)"
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   311
apply (perm_simp) 
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   312
oops
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   313
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   314
lemma 
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   315
  shows "p \<bullet> (\<lambda>(x::'a::pt). A \<longrightarrow> (B::'a \<Rightarrow> bool) x = C) = foo"
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   316
apply(perm_simp)
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   317
oops
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   318
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   319
lemma 
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   320
  shows "p \<bullet> (\<lambda>B::bool. A \<longrightarrow> (B = C)) = foo"
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   321
apply (perm_simp)
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   322
oops
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   323
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   324
lemma 
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   325
  shows "p \<bullet> (\<lambda>x y. \<exists>z. x = z \<and> x = y \<longrightarrow> z \<noteq> x) = foo"
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   326
apply (perm_simp)
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   327
oops
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   328
1800
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   329
lemma 
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   330
  shows "p \<bullet> (\<lambda>f x. f (g (f x))) = foo"
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   331
apply (perm_simp)
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   332
oops
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   333
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   334
lemma 
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   335
  fixes p q::"perm"
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   336
  and   x::"'a::pt"
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   337
  shows "p \<bullet> (q \<bullet> x) = foo"
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   338
apply(perm_simp)
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   339
oops
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   340
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   341
lemma 
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   342
  fixes p q r::"perm"
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   343
  and   x::"'a::pt"
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   344
  shows "p \<bullet> (q \<bullet> r \<bullet> x) = foo"
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   345
apply(perm_simp)
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   346
oops
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   347
1800
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   348
lemma 
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   349
  fixes p r::"perm"
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   350
  shows "p \<bullet> (\<lambda>q::perm. q \<bullet> (r \<bullet> x)) = foo"
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   351
apply (perm_simp)
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   352
oops
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   353
1800
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   354
lemma 
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   355
  fixes C D::"bool"
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   356
  shows "B (p \<bullet> (C = D))"
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   357
apply(perm_simp)
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   358
oops
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   359
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   360
declare [[trace_eqvt = false]]
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   361
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   362
text {* Problem: there is no raw eqvt-rule for The *}
78fdc6b36a1c changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   363
lemma "p \<bullet> (THE x. P x) = foo"
1867
f4477d3fe520 preliminary parser for perm_simp metod
Christian Urban <urbanc@in.tum.de>
parents: 1866
diff changeset
   364
apply(perm_strict_simp exclude: The)
f4477d3fe520 preliminary parser for perm_simp metod
Christian Urban <urbanc@in.tum.de>
parents: 1866
diff changeset
   365
apply(perm_simp exclude: The)
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   366
oops
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   367
1953
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
   368
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
   369
thm eqvts
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
   370
thm eqvts_raw
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
   371
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
   372
ML {*
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
   373
Nominal_ThmDecls.is_eqvt @{context} @{term "supp"}
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
   374
*}
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
   375
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
   376
1835
636de31888a6 tuned and removed dead code
Christian Urban <urbanc@in.tum.de>
parents: 1833
diff changeset
   377
(* automatic equivariance procedure for 
636de31888a6 tuned and removed dead code
Christian Urban <urbanc@in.tum.de>
parents: 1833
diff changeset
   378
   inductive definitions *)
636de31888a6 tuned and removed dead code
Christian Urban <urbanc@in.tum.de>
parents: 1833
diff changeset
   379
use "nominal_eqvt.ML"
1833
2050b5723c04 added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents: 1827
diff changeset
   380
1315
43d6e3730353 Add image_eqvt and atom_eqvt to eqvt bases.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1258
diff changeset
   381
end