Nominal-General/nominal_thmdecls.ML
author Christian Urban <urbanc@in.tum.de>
Thu, 13 May 2010 15:12:05 +0100
changeset 2123 2f39ce2aba6c
parent 1953 186d8486dfd5
child 2200 31f1ec832d39
permissions -rw-r--r--
removed internal functions from the signature (they are not needed anymore)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
     1
(*  Title:      nominal_thmdecls.ML
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
     2
    Author:     Christian Urban
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
     3
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
     4
  Infrastructure for the lemma collection "eqvts".
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
     5
1066
96651cddeba9 eqvts and eqvts_raw are separate thm-lists; otherwise permute_eqvt is problematic as it causes looks in eqvts
Christian Urban <urbanc@in.tum.de>
parents: 1059
diff changeset
     6
  Provides the attributes [eqvt] and [eqvt_raw], and the theorem
96651cddeba9 eqvts and eqvts_raw are separate thm-lists; otherwise permute_eqvt is problematic as it causes looks in eqvts
Christian Urban <urbanc@in.tum.de>
parents: 1059
diff changeset
     7
  lists eqvts and eqvts_raw. The first attribute will store the 
96651cddeba9 eqvts and eqvts_raw are separate thm-lists; otherwise permute_eqvt is problematic as it causes looks in eqvts
Christian Urban <urbanc@in.tum.de>
parents: 1059
diff changeset
     8
  theorem in the eqvts list and also in the eqvts_raw list. For 
96651cddeba9 eqvts and eqvts_raw are separate thm-lists; otherwise permute_eqvt is problematic as it causes looks in eqvts
Christian Urban <urbanc@in.tum.de>
parents: 1059
diff changeset
     9
  the latter the theorem is expected to be of the form
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
1811
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
    11
    p o (c x1 x2 ...) = c (p o x1) (p o x2) ...    (1)
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
    12
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
    13
  or
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
    14
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
    15
    c x1 x2 ... ==> c (p o x1) (p o x2) ...        (2)
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    16
1066
96651cddeba9 eqvts and eqvts_raw are separate thm-lists; otherwise permute_eqvt is problematic as it causes looks in eqvts
Christian Urban <urbanc@in.tum.de>
parents: 1059
diff changeset
    17
  and it is stored in the form
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    18
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    19
    p o c == c
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
1066
96651cddeba9 eqvts and eqvts_raw are separate thm-lists; otherwise permute_eqvt is problematic as it causes looks in eqvts
Christian Urban <urbanc@in.tum.de>
parents: 1059
diff changeset
    21
  The [eqvt_raw] attribute just adds the theorem to eqvts_raw.
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    22
1811
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
    23
  TODO: In case of the form in (2) one should also
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
    24
        add the equational form to eqvts
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
*)
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
signature NOMINAL_THMDECLS =
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
sig
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
  val eqvt_add: attribute
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
  val eqvt_del: attribute
1066
96651cddeba9 eqvts and eqvts_raw are separate thm-lists; otherwise permute_eqvt is problematic as it causes looks in eqvts
Christian Urban <urbanc@in.tum.de>
parents: 1059
diff changeset
    31
  val eqvt_raw_add: attribute
96651cddeba9 eqvts and eqvts_raw are separate thm-lists; otherwise permute_eqvt is problematic as it causes looks in eqvts
Christian Urban <urbanc@in.tum.de>
parents: 1059
diff changeset
    32
  val eqvt_raw_del: attribute
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
  val setup: theory -> theory
1066
96651cddeba9 eqvts and eqvts_raw are separate thm-lists; otherwise permute_eqvt is problematic as it causes looks in eqvts
Christian Urban <urbanc@in.tum.de>
parents: 1059
diff changeset
    34
  val get_eqvts_thms: Proof.context -> thm list
96651cddeba9 eqvts and eqvts_raw are separate thm-lists; otherwise permute_eqvt is problematic as it causes looks in eqvts
Christian Urban <urbanc@in.tum.de>
parents: 1059
diff changeset
    35
  val get_eqvts_raw_thms: Proof.context -> thm list
1870
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
    36
  val eqvt_transform: Proof.context -> thm -> thm
1953
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    37
  val is_eqvt: Proof.context -> term -> bool
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
end;
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    40
structure Nominal_ThmDecls: NOMINAL_THMDECLS =
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
struct
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    43
structure EqvtData = Generic_Data
1066
96651cddeba9 eqvts and eqvts_raw are separate thm-lists; otherwise permute_eqvt is problematic as it causes looks in eqvts
Christian Urban <urbanc@in.tum.de>
parents: 1059
diff changeset
    44
( type T = thm Item_Net.T;
96651cddeba9 eqvts and eqvts_raw are separate thm-lists; otherwise permute_eqvt is problematic as it causes looks in eqvts
Christian Urban <urbanc@in.tum.de>
parents: 1059
diff changeset
    45
  val empty = Thm.full_rules;
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    46
  val extend = I;
1953
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    47
  val merge = Item_Net.merge);
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    48
1066
96651cddeba9 eqvts and eqvts_raw are separate thm-lists; otherwise permute_eqvt is problematic as it causes looks in eqvts
Christian Urban <urbanc@in.tum.de>
parents: 1059
diff changeset
    49
structure EqvtRawData = Generic_Data
1953
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    50
( type T = thm Symtab.table;
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    51
  val empty = Symtab.empty;
1066
96651cddeba9 eqvts and eqvts_raw are separate thm-lists; otherwise permute_eqvt is problematic as it causes looks in eqvts
Christian Urban <urbanc@in.tum.de>
parents: 1059
diff changeset
    52
  val extend = I;
1953
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    53
  val merge = Symtab.merge (K true));
1059
090fa3f21380 restored the old behaviour of having an eqvts list; the transformed theorems are stored in eqvts_raw
Christian Urban <urbanc@in.tum.de>
parents: 1039
diff changeset
    54
1066
96651cddeba9 eqvts and eqvts_raw are separate thm-lists; otherwise permute_eqvt is problematic as it causes looks in eqvts
Christian Urban <urbanc@in.tum.de>
parents: 1059
diff changeset
    55
val eqvts = Item_Net.content o EqvtData.get;
1953
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    56
val eqvts_raw = map snd o Symtab.dest o EqvtRawData.get;
1066
96651cddeba9 eqvts and eqvts_raw are separate thm-lists; otherwise permute_eqvt is problematic as it causes looks in eqvts
Christian Urban <urbanc@in.tum.de>
parents: 1059
diff changeset
    57
1811
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
    58
val get_eqvts_thms = eqvts o  Context.Proof;
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
    59
val get_eqvts_raw_thms = eqvts_raw o Context.Proof;
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    61
val add_thm = EqvtData.map o Item_Net.update;
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    62
val del_thm = EqvtData.map o Item_Net.remove;
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    63
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
    64
fun add_raw_thm thm = 
1811
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
    65
  case prop_of thm of
1953
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    66
    Const ("==", _) $ _ $ Const (c, _) => EqvtRawData.map (Symtab.update (c, thm))
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    67
  | _ => raise THM ("Theorem must be a meta-equality where the right-hand side is a constant.", 0, [thm]) 
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
    68
1953
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    69
fun del_raw_thm thm = 
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    70
  case prop_of thm of
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    71
    Const ("==", _) $ _ $ Const (c, _) => EqvtRawData.map (Symtab.delete c)
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    72
  | _ => raise THM ("Theorem must be a meta-equality where the right-hand side is a constant.", 0, [thm]) 
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    73
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    74
fun is_eqvt ctxt trm =
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    75
  case trm of 
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    76
    Const (c, _) => Symtab.defined (EqvtRawData.get (Context.Proof ctxt)) c
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    77
  | _ => raise TERM ("Term must be a constsnt.", [trm])
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    78
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    79
1870
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
    80
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
    81
(** transformation of eqvt lemmas **)
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
    82
1953
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    83
fun get_perms trm =
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    84
  case trm of 
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    85
     Const (@{const_name permute}, _) $ _ $ (Bound _) => 
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    86
       raise TERM ("get_perms called on bound", [trm])
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    87
   | Const (@{const_name permute}, _) $ p $ _ => [p]
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    88
   | t $ u => get_perms t @ get_perms u
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    89
   | Abs (_, _, t) => get_perms t 
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    90
   | _ => []
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    91
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    92
fun put_perm p trm =
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    93
  case trm of 
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    94
     Bound _ => trm
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    95
   | Const _ => trm
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    96
   | t $ u => put_perm p t $ put_perm p u
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    97
   | Abs (x, ty, t) => Abs (x, ty, put_perm p t)
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    98
   | _ => mk_perm p trm
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
    99
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
   100
(* tests whether there is a disagreement between the permutations, 
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
   101
   and that there is at least one permutation *)
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
   102
fun is_bad_list [] = true
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
   103
  | is_bad_list [_] = false
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
   104
  | is_bad_list (p::q::ps) = if p = q then is_bad_list (q::ps) else true
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
   105
1870
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   106
1947
51f411b1197d tuned and cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1870
diff changeset
   107
(* transforms equations into the "p o c == c"-form 
1870
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   108
   from p o (c x1 ...xn) = c (p o x1) ... (p o xn) *)
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   109
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   110
fun eqvt_transform_eq_tac thm = 
1841
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
   111
let
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
   112
  val ss_thms = @{thms permute_minus_cancel permute_prod.simps split_paired_all}
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
   113
in
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
   114
  REPEAT o FIRST' 
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
   115
    [CHANGED o simp_tac (HOL_basic_ss addsimps ss_thms),
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
   116
     rtac (thm RS @{thm trans}),
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
   117
     rtac @{thm trans[OF permute_fun_def]} THEN' rtac @{thm ext}]
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
   118
end
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
1870
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   120
fun eqvt_transform_eq ctxt thm = 
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
   121
let
1947
51f411b1197d tuned and cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1870
diff changeset
   122
  val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of thm))
51f411b1197d tuned and cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1870
diff changeset
   123
    handle TERM _ => error "Equivariance lemma must be an equality."
51f411b1197d tuned and cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1870
diff changeset
   124
  val (p, t) = dest_perm lhs 
51f411b1197d tuned and cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1870
diff changeset
   125
    handle TERM _ => error "Equivariance lemma is not of the form p \<bullet> c...  = c..."
51f411b1197d tuned and cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1870
diff changeset
   126
1953
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
   127
  val ps = get_perms rhs handle TERM _ => []  
1841
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
   128
  val (c, c') = (head_of t, head_of rhs)
1947
51f411b1197d tuned and cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1870
diff changeset
   129
  val msg = "Equivariance lemma is not of the right form "
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
   130
in
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
   131
  if c <> c' 
1947
51f411b1197d tuned and cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1870
diff changeset
   132
    then error (msg ^ "(constants do not agree).")
1953
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
   133
  else if is_bad_list (p :: ps)  
1947
51f411b1197d tuned and cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1870
diff changeset
   134
    then error (msg ^ "(permutations do not agree).") 
1953
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
   135
  else if not (rhs aconv (put_perm p t))
1947
51f411b1197d tuned and cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1870
diff changeset
   136
    then error (msg ^ "(arguments do not agree).")
1841
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
   137
  else if is_Const t 
1870
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   138
    then safe_mk_equiv thm
1841
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
   139
  else 
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
   140
    let 
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
   141
      val goal = HOLogic.mk_Trueprop (HOLogic.mk_eq (mk_perm p c, c))
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
   142
      val ([goal', p'], ctxt') = Variable.import_terms false [goal, p] ctxt
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
   143
    in
1870
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   144
      Goal.prove ctxt [] [] goal' (fn _ => eqvt_transform_eq_tac thm 1)
1841
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
   145
      |> singleton (ProofContext.export ctxt' ctxt)
1870
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   146
      |> safe_mk_equiv
1947
51f411b1197d tuned and cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1870
diff changeset
   147
      |> zero_var_indexes
1841
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
   148
    end
1811
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   149
end
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   150
1947
51f411b1197d tuned and cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1870
diff changeset
   151
(* transforms equations into the "p o c == c"-form 
1870
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   152
   from R x1 ...xn ==> R (p o x1) ... (p o xn) *)
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   153
1947
51f411b1197d tuned and cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1870
diff changeset
   154
fun eqvt_transform_imp_tac ctxt p p' thm = 
1811
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   155
let
1947
51f411b1197d tuned and cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1870
diff changeset
   156
  val thy = ProofContext.theory_of ctxt
1811
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   157
  val cp = Thm.cterm_of thy p
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   158
  val cp' = Thm.cterm_of thy (mk_minus p')
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   159
  val thm' = Drule.cterm_instantiate [(cp, cp')] thm
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   160
  val simp = HOL_basic_ss addsimps @{thms permute_minus_cancel(2)}
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   161
in
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   162
  EVERY' [rtac @{thm iffI}, dtac @{thm permute_boolE}, rtac thm, atac,
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   163
    rtac @{thm permute_boolI}, dtac thm', full_simp_tac simp]
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
   164
end
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
   165
1870
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   166
fun eqvt_transform_imp ctxt thm =
1811
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   167
let
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   168
  val (prem, concl) = pairself HOLogic.dest_Trueprop (Logic.dest_implies (prop_of thm))
1841
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
   169
  val (c, c') = (head_of prem, head_of concl)
1953
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
   170
  val ps = get_perms concl handle TERM _ => []  
1811
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   171
  val p = try hd ps
1947
51f411b1197d tuned and cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1870
diff changeset
   172
  val msg = "Equivariance lemma is not of the right form "
1811
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   173
in
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   174
  if c <> c' 
1947
51f411b1197d tuned and cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1870
diff changeset
   175
    then error (msg ^ "(constants do not agree).")
1811
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   176
  else if is_bad_list ps  
1947
51f411b1197d tuned and cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1870
diff changeset
   177
    then error (msg ^ "(permutations do not agree).") 
1953
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
   178
  else if not (concl aconv (put_perm (the p) prem)) 
1947
51f411b1197d tuned and cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1870
diff changeset
   179
    then error (msg ^ "(arguments do not agree).")
1811
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   180
  else 
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   181
    let
1841
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
   182
      val prem' = mk_perm (the p) prem    
1811
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   183
      val goal = HOLogic.mk_Trueprop (HOLogic.mk_eq (prem', concl))
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   184
      val ([goal', p'], ctxt') = Variable.import_terms false [goal, the p] ctxt
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   185
    in
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   186
      Goal.prove ctxt' [] [] goal'
1947
51f411b1197d tuned and cleaned
Christian Urban <urbanc@in.tum.de>
parents: 1870
diff changeset
   187
        (fn _ => eqvt_transform_imp_tac ctxt' (the p) p' thm 1) 
1811
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   188
      |> singleton (ProofContext.export ctxt' ctxt)
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   189
    end
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   190
end     
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   191
1870
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   192
fun eqvt_transform ctxt thm = 
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   193
 case (prop_of thm) of
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   194
   @{const "Trueprop"} $ (Const (@{const_name "op ="}, _) $ 
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   195
     (Const (@{const_name "permute"}, _) $ _ $ _) $ _) => 
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   196
       eqvt_transform_eq ctxt thm
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   197
 | @{const "==>"} $ (@{const "Trueprop"} $ _) $ (@{const "Trueprop"} $ _) => 
1953
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
   198
     eqvt_transform_imp ctxt thm |> eqvt_transform_eq ctxt
1870
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   199
 | _ => raise error "Only _ = _ and _ ==> _ cases are implemented."
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   200
 
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   201
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   202
(** attributes **)
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   203
1870
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   204
val eqvt_add = Thm.declaration_attribute 
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   205
  (fn thm => fn context =>
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   206
   let
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   207
     val thm' = eqvt_transform (Context.proof_of context) thm
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   208
   in
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   209
     context |> add_thm thm |> add_raw_thm thm'
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   210
   end)
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   211
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   212
val eqvt_del = Thm.declaration_attribute
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   213
  (fn thm => fn context =>
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   214
   let
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   215
     val thm' = eqvt_transform (Context.proof_of context) thm
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   216
   in
1953
186d8486dfd5 rewrote eqvts_raw to be a symtab, that can be looked up
Christian Urban <urbanc@in.tum.de>
parents: 1947
diff changeset
   217
     context |> del_thm thm  |> del_raw_thm thm'
1870
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   218
   end)
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   219
1066
96651cddeba9 eqvts and eqvts_raw are separate thm-lists; otherwise permute_eqvt is problematic as it causes looks in eqvts
Christian Urban <urbanc@in.tum.de>
parents: 1059
diff changeset
   220
val eqvt_raw_add = Thm.declaration_attribute add_raw_thm;
96651cddeba9 eqvts and eqvts_raw are separate thm-lists; otherwise permute_eqvt is problematic as it causes looks in eqvts
Christian Urban <urbanc@in.tum.de>
parents: 1059
diff changeset
   221
val eqvt_raw_del = Thm.declaration_attribute del_raw_thm;
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   222
1870
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   223
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   224
(** setup function **)
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   225
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   226
val setup =
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
   227
  Attrib.setup @{binding "eqvt"} (Attrib.add_del eqvt_add eqvt_del) 
1811
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   228
    (cat_lines ["Declaration of equivariance lemmas - they will automtically be",  
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   229
       "brought into the form p o c == c"]) #>
1066
96651cddeba9 eqvts and eqvts_raw are separate thm-lists; otherwise permute_eqvt is problematic as it causes looks in eqvts
Christian Urban <urbanc@in.tum.de>
parents: 1059
diff changeset
   230
  Attrib.setup @{binding "eqvt_raw"} (Attrib.add_del eqvt_raw_add eqvt_raw_del) 
1811
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   231
    (cat_lines ["Declaration of equivariance lemmas - no",
1059
090fa3f21380 restored the old behaviour of having an eqvts list; the transformed theorems are stored in eqvts_raw
Christian Urban <urbanc@in.tum.de>
parents: 1039
diff changeset
   232
       "transformation is performed"]) #>
090fa3f21380 restored the old behaviour of having an eqvts list; the transformed theorems are stored in eqvts_raw
Christian Urban <urbanc@in.tum.de>
parents: 1039
diff changeset
   233
  PureThy.add_thms_dynamic (@{binding "eqvts"}, eqvts) #>
090fa3f21380 restored the old behaviour of having an eqvts list; the transformed theorems are stored in eqvts_raw
Christian Urban <urbanc@in.tum.de>
parents: 1039
diff changeset
   234
  PureThy.add_thms_dynamic (@{binding "eqvts_raw"}, eqvts_raw);
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   235
1066
96651cddeba9 eqvts and eqvts_raw are separate thm-lists; otherwise permute_eqvt is problematic as it causes looks in eqvts
Christian Urban <urbanc@in.tum.de>
parents: 1059
diff changeset
   236
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   237
end;