Quot/Nominal/nominal_thmdecls.ML
author Christian Urban <urbanc@in.tum.de>
Thu, 04 Feb 2010 14:55:21 +0100
changeset 1059 090fa3f21380
parent 1039 0d832c36b1bb
child 1066 96651cddeba9
permissions -rw-r--r--
restored the old behaviour of having an eqvts list; the transformed theorems are stored in eqvts_raw
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
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
     6
  Provides the attributes [eqvt] and [eqvt_force], and the theorem
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
     7
  lists eqvts and eqvts_raw. The first attribute will store the theorem
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
     8
  in the eqvts list and also in the eqvts_raw list. For the latter 
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
     9
  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
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    11
    p o (c x1 x2 ...) = c (p o x1) (p o x2) ...
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    12
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
    13
  and 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
    14
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    15
    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
    16
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
    17
  The [eqvt_force] attribute just adds the theorem to eqvts.
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    18
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
    19
  TODO:
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
    20
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
    21
   - deal with eqvt-lemmas of the form 
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    22
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    23
       c x1 x2 ... ==> c (p o x1) (p o x2) ..
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
*)
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
signature NOMINAL_THMDECLS =
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
sig
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
  val eqvt_add: attribute
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
  val eqvt_del: attribute
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
  val eqvt_force_add: attribute
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
  val eqvt_force_del: attribute
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
  val setup: theory -> theory
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
  val get_eqvt_thms: Proof.context -> thm list
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
    34
  val get_eqvt_raw_thms: Proof.context -> thm list
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
end;
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    38
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
    39
struct
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    41
structure EqvtData = Generic_Data
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
(
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
    43
  type T = (thm * thm option) Item_Net.T;
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
    44
  val empty = Item_Net.init (eq_fst Thm.eq_thm_prop) (single o Thm.full_prop_of o fst)
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    45
  val extend = I;
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    46
  val merge = Item_Net.merge;
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    47
);
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    48
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
    49
val eqvts = (map fst) o Item_Net.content o EqvtData.get;
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
    50
val eqvts_raw = (map_filter snd) o Item_Net.content o EqvtData.get;
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
    51
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
    52
val get_eqvt_thms = eqvts o Context.Proof; 
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
    53
val get_eqvt_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
    54
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    55
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
    56
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
    57
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    58
val add_force_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
    59
val del_force_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
    60
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    61
fun dest_perm (Const (@{const_name "permute"}, _) $ p $ t) = (p, t)
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    62
  | dest_perm t = raise TERM("dest_perm", [t])
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    64
fun mk_perm p trm =
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
let
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    66
  val ty = fastype_of trm
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
in
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    68
  Const (@{const_name "permute"}, @{typ "perm"} --> ty --> ty) $ p $ trm
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    69
end
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    71
fun eqvt_transform_tac thm = REPEAT o FIRST' 
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    72
  [CHANGED o simp_tac (HOL_basic_ss addsimps @{thms permute_minus_cancel}),
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    73
   rtac (thm RS @{thm trans}),
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    74
   rtac @{thm trans[OF permute_fun_def]} THEN' rtac @{thm ext}]
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    76
(* transform equations into the required form *)
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    77
fun transform_eq ctxt thm lhs rhs = 
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    78
let
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    79
  val (p, t) = dest_perm lhs
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    80
  val (c, args) = strip_comb t
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    81
  val (c', args') = strip_comb rhs 
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    82
  val eargs = map Envir.eta_contract args 
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    83
  val eargs' = map Envir.eta_contract args'
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    84
  val p_str = fst (fst (dest_Var p))
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    85
  val goal = HOLogic.mk_Trueprop (HOLogic.mk_eq (mk_perm p c, c))
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    86
in
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    87
  if c <> c' 
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    88
    then error "eqvt lemma is not of the right form (constants do not agree)"
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    89
  else if eargs' <> map (mk_perm p) eargs 
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    90
    then error "eqvt lemma is not of the right form (arguments do not agree)"
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    91
  else if args = [] 
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    92
    then thm
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    93
  else Goal.prove ctxt [p_str] [] goal
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    94
    (fn _ => eqvt_transform_tac thm 1)
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    95
end
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    96
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    97
fun transform addel_fn thm context = 
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    98
let
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    99
  val ctxt = Context.proof_of context
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
   100
  val trm = HOLogic.dest_Trueprop (prop_of thm)
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
   101
in
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
   102
  case trm of
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
   103
    Const (@{const_name "op ="}, _) $ lhs $ rhs =>
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
   104
      let 
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
   105
        val thm2 = transform_eq ctxt thm lhs rhs RS @{thm eq_reflection}
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
   106
      in
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
   107
        addel_fn (thm, SOME thm2) context
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
   108
      end
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
   109
  | _ => raise (error "only (op=) case implemented yet")
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
   110
end 
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
   112
val eqvt_add = Thm.declaration_attribute (transform add_thm);
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
   113
val eqvt_del = Thm.declaration_attribute (transform del_thm);
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   114
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
   115
val eqvt_force_add = Thm.declaration_attribute (add_force_thm o rpair NONE);
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
   116
val eqvt_force_del = Thm.declaration_attribute (del_force_thm o rpair NONE);
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
val setup =
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
   119
  Attrib.setup @{binding "eqvt"} (Attrib.add_del eqvt_add eqvt_del) 
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
   120
    (cat_lines ["declaration of equivariance lemmas - they will automtically be",  
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
   121
       "brought into the form p o c = c"]) #>
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
   122
  Attrib.setup @{binding "eqvt_force"} (Attrib.add_del eqvt_force_add eqvt_force_del) 
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
   123
    (cat_lines ["declaration of equivariance lemmas - no",
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
   124
       "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
   125
  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
   126
  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
   127
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
end;