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