Nominal-General/nominal_thmdecls.ML
author Christian Urban <urbanc@in.tum.de>
Sun, 18 Apr 2010 17:57:27 +0200
changeset 1870 55b2da92ff2f
parent 1846 756982b4fe20
child 1947 51f411b1197d
permissions -rw-r--r--
tuned; transformation functions now take a context, a thm and return a thm
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
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
1846
756982b4fe20 temporary fix for CoreHaskell
Christian Urban <urbanc@in.tum.de>
parents: 1841
diff changeset
    38
  (* TEMPORARY FIX *)
756982b4fe20 temporary fix for CoreHaskell
Christian Urban <urbanc@in.tum.de>
parents: 1841
diff changeset
    39
  val add_thm: thm -> Context.generic -> Context.generic
756982b4fe20 temporary fix for CoreHaskell
Christian Urban <urbanc@in.tum.de>
parents: 1841
diff changeset
    40
  val add_raw_thm: thm -> Context.generic -> Context.generic
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
end;
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 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
    44
struct
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
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
    46
fun get_ps trm =
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
    47
  case trm of 
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
    48
     Const (@{const_name permute}, _) $ p $ (Bound _) => raise TERM ("get_ps", [trm])
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
    49
   | Const (@{const_name permute}, _) $ p $ _ => [p]
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
    50
   | t $ u => get_ps t @ get_ps u
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
    51
   | Abs (_, _, t) => get_ps t 
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
    52
   | _ => []
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
    53
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
    54
fun put_p p trm =
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
    55
  case trm of 
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
    56
     Bound _ => trm
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
    57
   | Const _ => trm
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
    58
   | t $ u => put_p p t $ put_p p u
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
    59
   | Abs (x, ty, t) => Abs (x, ty, put_p p t)
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
    60
   | _ => mk_perm p trm
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
    61
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
    62
(* tests whether the lists of ps all agree, and that there is at least one p *)
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
    63
fun is_bad_list [] = true
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
    64
  | is_bad_list [_] = false
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
    65
  | is_bad_list (p::q::ps) = if p = q then is_bad_list (q::ps) else true
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
    66
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
    67
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    68
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
    69
( 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
    70
  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
    71
  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
    72
  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
    73
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
    74
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
    75
( 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
    76
  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
    77
  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
    78
  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
    79
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
    80
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
    81
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
    82
1811
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
    83
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
    84
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
    85
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
    86
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
    87
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
    88
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
    89
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
    90
  case prop_of thm of
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
    91
    Const ("==", _) $ _ $ _ => 
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
    92
      EqvtRawData.map (Item_Net.update (zero_var_indexes 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
    93
  | _ => raise THM ("Theorem must be a meta-equality", 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
    94
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
    95
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
    96
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
    97
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
    98
(** 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
    99
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   100
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   101
(* transforms equations into the "p o c = c"-form 
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   102
   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
   103
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   104
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
   105
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
   106
  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
   107
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
   108
  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
   109
    [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
   110
     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
   111
     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
   112
end
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
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
   114
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
   115
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
   116
  val ((p, t), rhs) = apfst dest_perm 
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
    (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of thm)))
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
    handle TERM _ => error "Eqvt lemma is not of the form p \<bullet> 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
   119
  val ps = get_ps rhs handle TERM _ => []  
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
   120
  val (c, c') = (head_of t, head_of rhs)
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
   121
in
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
   122
  if c <> c' 
1811
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   123
    then error "Eqvt lemma is not of the right form (constants 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
   124
  else if is_bad_list (p::ps)  
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
   125
    then error "Eqvt lemma is not of the right form (permutations do not agree)" 
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
   126
  else if not (rhs aconv (put_p p t))
1811
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   127
    then error "Eqvt lemma is not of the right form (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
   128
  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
   129
    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
   130
  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
   131
    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
   132
      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
   133
      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
   134
    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
   135
      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
   136
      |> 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
   137
      |> safe_mk_equiv
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
   138
    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
   139
end
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   140
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
   141
(* transforms equations into the "p o c = c"-form 
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   142
   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
   143
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
fun eqvt_transform_imp_tac thy 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
   145
let
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   146
  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
   147
  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
   148
  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
   149
  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
   150
in
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   151
  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
   152
    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
   153
end
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
   154
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
   155
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
   156
let
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 thy = ProofContext.theory_of ctxt
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 (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
   159
  val (c, c') = (head_of prem, head_of concl)
fcc660ece040 thmdecls can deal with lemmas like alpha_gen which contain pairs or tuples
Christian Urban <urbanc@in.tum.de>
parents: 1835
diff changeset
   160
  val ps = get_ps 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
   161
  val p = try hd ps
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   162
in
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   163
  if c <> c' 
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   164
    then error "Eqvt lemma is not of the right form (constants do not agree)"
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   165
  else if is_bad_list ps  
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   166
    then error "Eqvt lemma is not of the right form (permutations 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
   167
  else if not (concl aconv (put_p (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
   168
    then error "Eqvt lemma is not of the right form (arguments do not agree)"
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   169
  else 
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   170
    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
   171
      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
   172
      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
   173
      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
   174
    in
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   175
      Goal.prove ctxt' [] [] goal'
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
   176
        (fn _ => eqvt_transform_imp_tac thy (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
   177
      |> 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
   178
      |> eqvt_transform_eq 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
   179
    end
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   180
end     
ae176476b525 implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
Christian Urban <urbanc@in.tum.de>
parents: 1810
diff changeset
   181
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
   182
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
   183
 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
   184
   @{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
   185
     (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
   186
       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
   187
 | @{const "==>"} $ (@{const "Trueprop"} $ _) $ (@{const "Trueprop"} $ _) => 
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   188
     eqvt_transform_imp 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
   189
 | _ => 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
   190
 
55b2da92ff2f tuned; transformation functions now take a context, a thm and return a thm
Christian Urban <urbanc@in.tum.de>
parents: 1846
diff changeset
   191
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
(** attributes **)
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   193
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
   194
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
   195
  (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
   196
   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
   197
     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
   198
   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
   199
     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
   200
   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
   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
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
   203
  (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
   204
   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
   205
     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
   206
   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
   207
     context |> del_thm thm |> del_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
   208
   end)
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   209
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
   210
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
   211
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
   212
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
   213
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
(** 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
   215
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   216
val setup =
1037
2845e736dc1a added a first eqvt_tac which pushes permutations inside terms
Christian Urban <urbanc@in.tum.de>
parents: 947
diff changeset
   217
  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
   218
    (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
   219
       "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
   220
  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
   221
    (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
   222
       "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
   223
  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
   224
  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
   225
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
   226
947
fa810f01f7b5 added an LamEx example together with the new nominal infrastructure
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   227
end;