Nominal/nominal_thmdecls.ML
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Wed, 10 Mar 2010 09:58:43 +0100
changeset 1387 9d70a0733786
parent 1258 7d8949da7d99
permissions -rw-r--r--
rhs of alpha_bn, and template for the arguments.
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
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
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
    13
  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
    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
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
  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
    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
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
    30
  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
    31
  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
    32
  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
    33
  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
    34
  val get_eqvts_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
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
    41
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    42
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
    43
( 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
    44
  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
    45
  val extend = I;
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
    46
  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
    47
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
    48
structure EqvtRawData = Generic_Data
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
( 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
    50
  val empty = Thm.full_rules;
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
    51
  val extend = I;
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 merge = Item_Net.merge );
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
    53
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
    54
val eqvts = Item_Net.content o EqvtData.get;
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_raw = Item_Net.content o EqvtRawData.get;
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
    56
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
val get_eqvts_thms = eqvts o Context.Proof; 
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
    58
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
    59
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    60
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
    61
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
    62
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
    63
val add_raw_thm = EqvtRawData.map o Item_Net.update;
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
    64
val del_raw_thm = EqvtRawData.map o Item_Net.remove;
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    65
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    66
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
    67
  | 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
    68
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    69
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
    70
let
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    71
  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
    72
in
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    73
  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
    74
end
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
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
    77
  [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
    78
   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
    79
   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
    80
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    81
(* 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
    82
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
    83
let
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    84
  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
    85
  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
    86
  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
    87
  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
    88
  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
    89
  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
    90
  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
    91
in
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    92
  if c <> c' 
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    93
    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
    94
  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
    95
    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
    96
  else if args = [] 
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    97
    then thm
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    98
  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
    99
    (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
   100
end
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
   101
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
   102
fun transform addel_fun thm context = 
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
   103
let
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
   104
  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
   105
  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
   106
in
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
   107
  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
   108
    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
   109
      let 
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
   110
        val thm' = transform_eq ctxt thm lhs rhs RS @{thm eq_reflection}
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
   111
      in
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
   112
        addel_fun thm' context
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
   113
      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
   114
  | _ => 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
   115
end 
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
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
   117
val eqvt_add = Thm.declaration_attribute (fn thm => (add_thm thm) o (transform add_raw_thm 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
   118
val eqvt_del = Thm.declaration_attribute (fn thm => (del_thm thm) o (transform del_raw_thm thm));
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
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
   120
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
   121
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
   122
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
val setup =
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
   124
  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
   125
    (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
   126
       "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
   127
  Attrib.setup @{binding "eqvt_raw"} (Attrib.add_del eqvt_raw_add eqvt_raw_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
   128
    (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
   129
       "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
   130
  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
   131
  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
   132
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
   133
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   134
end;