| author | haftmann | 
| Thu, 20 Mar 2008 12:04:53 +0100 | |
| changeset 26357 | 19b153ebda0b | 
| parent 26343 | 0dd2eab7b296 | 
| child 26400 | 2f0500e375fe | 
| permissions | -rw-r--r-- | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 1 | (* ID: "$Id$" | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 2 | Authors: Julien Narboux and Christian Urban | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 3 | |
| 22846 | 4 | This file introduces the infrastructure for the lemma | 
| 22562 
80b814fe284b
rename bij and fresh into bijs and freshs and lookup for eqvts lemma using the ml value instead of the name
 narboux parents: 
22541diff
changeset | 5 | declaration "eqvts" "bijs" and "freshs". | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 6 | |
| 22846 | 7 | By attaching [eqvt] [bij] or [fresh] to a lemma, the lemma gets stored | 
| 24571 | 8 | in a data-slot in the context. Possible modifiers | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 9 | are [attribute add] and [attribute del] for adding and deleting, | 
| 22846 | 10 | respectively the lemma from the data-slot. | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 11 | *) | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 12 | |
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 13 | signature NOMINAL_THMDECLS = | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 14 | sig | 
| 22250 
0d7ea7d2bc28
fixed two stupid bugs of SML to do with the value restriction and missing type
 urbanc parents: 
22245diff
changeset | 15 | val print_data: Proof.context -> unit | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 16 | val eqvt_add: attribute | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 17 | val eqvt_del: attribute | 
| 22715 
381e6c45f13b
improved the equivariance lemmas for the quantifiers; had to export the lemma eqvt_force_add and eqvt_force_del in the thmdecls
 urbanc parents: 
22562diff
changeset | 18 | val eqvt_force_add: attribute | 
| 
381e6c45f13b
improved the equivariance lemmas for the quantifiers; had to export the lemma eqvt_force_add and eqvt_force_del in the thmdecls
 urbanc parents: 
22562diff
changeset | 19 | val eqvt_force_del: attribute | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 20 | val setup: theory -> theory | 
| 24571 | 21 | val get_eqvt_thms: Proof.context -> thm list | 
| 22 | val get_fresh_thms: Proof.context -> thm list | |
| 23 | val get_bij_thms: Proof.context -> thm list | |
| 22846 | 24 | |
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 25 | |
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 26 | val NOMINAL_EQVT_DEBUG : bool ref | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 27 | end; | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 28 | |
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 29 | structure NominalThmDecls: NOMINAL_THMDECLS = | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 30 | struct | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 31 | |
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 32 | structure Data = GenericDataFun | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 33 | ( | 
| 22562 
80b814fe284b
rename bij and fresh into bijs and freshs and lookup for eqvts lemma using the ml value instead of the name
 narboux parents: 
22541diff
changeset | 34 |   type T = {eqvts:thm list, freshs:thm list, bijs:thm list};
 | 
| 
80b814fe284b
rename bij and fresh into bijs and freshs and lookup for eqvts lemma using the ml value instead of the name
 narboux parents: 
22541diff
changeset | 35 |   val empty = ({eqvts=[], freshs=[], bijs=[]}:T);
 | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 36 | val extend = I; | 
| 24039 
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
22846diff
changeset | 37 |   fun merge _ (r1:T,r2:T) = {eqvts  = Thm.merge_thms (#eqvts r1, #eqvts r2),
 | 
| 
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
22846diff
changeset | 38 | freshs = Thm.merge_thms (#freshs r1, #freshs r2), | 
| 
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
22846diff
changeset | 39 | bijs = Thm.merge_thms (#bijs r1, #bijs r2)} | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 40 | ); | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 41 | |
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 42 | (* Exception for when a theorem does not conform with form of an equivariance lemma. *) | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 43 | (* There are two forms: one is an implication (for relations) and the other is an *) | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 44 | (* equality (for functions). In the implication-case, say P ==> Q, Q must be equal *) | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 45 | (* to P except that every free variable of Q, say x, is replaced by pi o x. In the *) | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 46 | (* equality case, say lhs = rhs, the lhs must be of the form pi o t and the rhs must *) | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 47 | (* be equal to t except that every free variable, say x, is replaced by pi o x. In *) | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 48 | (* the implicational case it is also checked that the variables and permutation fit *) | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 49 | (* together, i.e. are of the right "pt_class", so that a stronger version of the *) | 
| 24265 
4d5917cc60c4
fix the generation of eqvt lemma of equality form from the imp form when the relation is equality
 narboux parents: 
24039diff
changeset | 50 | (* equality-lemma can be derived. *) | 
| 22418 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22286diff
changeset | 51 | exception EQVT_FORM of string; | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 52 | |
| 22846 | 53 | fun print_data ctxt = | 
| 54 | let | |
| 55 | val data = Data.get (Context.Proof ctxt); | |
| 56 | fun pretty_thms msg thms = | |
| 57 | Pretty.big_list msg (map (ProofContext.pretty_thm ctxt) thms); | |
| 58 | in | |
| 59 | Pretty.writeln (Pretty.chunks | |
| 60 | [pretty_thms "Equivariance lemmas:" (#eqvts data), | |
| 61 | pretty_thms "Freshness lemmas:" (#freshs data), | |
| 62 | pretty_thms "Bijection lemmas:" (#bijs data)]) | |
| 63 | end; | |
| 64 | ||
| 24571 | 65 | val get_eqvt_thms = Context.Proof #> Data.get #> #eqvts; | 
| 66 | val get_fresh_thms = Context.Proof #> Data.get #> #freshs; | |
| 67 | val get_bij_thms = Context.Proof #> Data.get #> #bijs; | |
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 68 | |
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 69 | (* FIXME: should be a function in a library *) | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 70 | fun mk_permT T = HOLogic.listT (HOLogic.mk_prodT (T, T)); | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 71 | |
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 72 | val NOMINAL_EQVT_DEBUG = ref false; | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 73 | |
| 22846 | 74 | fun tactic (msg,tac) = | 
| 75 | if !NOMINAL_EQVT_DEBUG | |
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 76 |     then (EVERY [tac, print_tac ("after "^msg)])
 | 
| 22846 | 77 | else tac | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 78 | |
| 22846 | 79 | fun tactic_eqvt ctx orig_thm pi typi = | 
| 80 | let | |
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 81 | val mypi = Thm.cterm_of ctx (Var (pi,typi)) | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 82 |         val mypifree = Thm.cterm_of ctx (Const ("List.rev",typi --> typi) $ Free (fst pi,typi))
 | 
| 26343 
0dd2eab7b296
simplified get_thm(s): back to plain name argument;
 wenzelm parents: 
26337diff
changeset | 83 | val perm_pi_simp = PureThy.get_thms ctx "perm_pi_simp" | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 84 | in | 
| 22846 | 85 |         EVERY [tactic ("iffI applied",rtac iffI 1),
 | 
| 24571 | 86 | 	       tactic ("remove pi with perm_boolE", (dtac @{thm perm_boolE} 1)),
 | 
| 87 |                tactic ("solve with orig_thm", (etac orig_thm 1)),
 | |
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 88 |                tactic ("applies orig_thm instantiated with rev pi",
 | 
| 22846 | 89 | dtac (Drule.cterm_instantiate [(mypi,mypifree)] orig_thm) 1), | 
| 24265 
4d5917cc60c4
fix the generation of eqvt lemma of equality form from the imp form when the relation is equality
 narboux parents: 
24039diff
changeset | 90 | 	       tactic ("getting rid of the pi on the right",
 | 
| 24571 | 91 |                           (rtac @{thm perm_boolI} 1)),
 | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 92 |                tactic ("getting rid of all remaining perms",
 | 
| 24265 
4d5917cc60c4
fix the generation of eqvt lemma of equality form from the imp form when the relation is equality
 narboux parents: 
24039diff
changeset | 93 | full_simp_tac (HOL_basic_ss addsimps perm_pi_simp) 1)] | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 94 | end; | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 95 | |
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 96 | fun get_derived_thm thy hyp concl orig_thm pi typi = | 
| 22846 | 97 | let | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 98 |        val lhs = (Const("Nominal.perm", typi --> HOLogic.boolT --> HOLogic.boolT) $ Var(pi,typi) $ hyp)
 | 
| 22846 | 99 | val goal_term = Logic.unvarify (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs,concl))) | 
| 100 | val _ = Display.print_cterm (cterm_of thy goal_term) | |
| 24039 
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
22846diff
changeset | 101 | in | 
| 22846 | 102 | Goal.prove_global thy [] [] goal_term (fn _ => (tactic_eqvt thy orig_thm pi typi)) | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 103 | end | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 104 | |
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 105 | (* FIXME: something naughty here, but as long as there is no infrastructure to expose *) | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 106 | (* the eqvt_thm_list to the user, we have to manually update the context and the *) | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 107 | (* thm-list eqvt *) | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 108 | fun update_context flag thms context = | 
| 22846 | 109 | let | 
| 110 | val context' = fold (fn thm => Data.map (flag thm)) thms context | |
| 111 | in | |
| 22541 
c33b542394f3
the name for the collection of equivariance lemmas is now eqvts (changed from eqvt) in order to avoid clashes with eqvt-lemmas generated in nominal_inductive
 urbanc parents: 
22495diff
changeset | 112 |     Context.mapping (snd o PureThy.add_thmss [(("eqvts",(#eqvts (Data.get context'))),[])]) I context'
 | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 113 | end; | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 114 | |
| 22846 | 115 | (* replaces every variable x in t with pi o x *) | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 116 | fun apply_pi trm (pi,typi) = | 
| 22846 | 117 | let | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 118 | fun only_vars t = | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 119 | (case t of | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 120 |           Var (n,ty) => (Const ("Nominal.perm",typi --> ty --> ty) $ (Var (pi,typi)) $ (Var (n,ty)))
 | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 121 | | _ => t) | 
| 22846 | 122 | in | 
| 123 | map_aterms only_vars trm | |
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 124 | end; | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 125 | |
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 126 | (* returns *the* pi which is in front of all variables, provided there *) | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 127 | (* exists such a pi; otherwise raises EQVT_FORM *) | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 128 | fun get_pi t thy = | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 129 | let fun get_pi_aux s = | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 130 | (case s of | 
| 22846 | 131 |           (Const ("Nominal.perm",typrm) $
 | 
| 132 |              (Var (pi,typi as Type("List.list",[Type ("*",[Type (tyatm,[]),_])]))) $
 | |
| 133 | (Var (n,ty))) => | |
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 134 | let | 
| 24039 
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
22846diff
changeset | 135 | (* FIXME: this should be an operation the library *) | 
| 22846 | 136 | val class_name = (NameSpace.map_base (fn s => "pt_"^s) tyatm) | 
| 24039 
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
22846diff
changeset | 137 | in | 
| 24271 | 138 | if (Sign.of_sort thy (ty,[class_name])) | 
| 22846 | 139 | then [(pi,typi)] | 
| 140 | else raise | |
| 141 |                 EQVT_FORM ("Could not find any permutation or an argument is not an instance of "^class_name)
 | |
| 142 | end | |
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 143 | | Abs (_,_,t1) => get_pi_aux t1 | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 144 | | (t1 $ t2) => get_pi_aux t1 @ get_pi_aux t2 | 
| 22846 | 145 | | _ => []) | 
| 146 | in | |
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 147 | (* collect first all pi's in front of variables in t and then use distinct *) | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 148 | (* to ensure that all pi's must have been the same, i.e. distinct returns *) | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 149 | (* a singleton-list *) | 
| 22846 | 150 | (case (distinct (op =) (get_pi_aux t)) of | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 151 | [(pi,typi)] => (pi,typi) | 
| 22418 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22286diff
changeset | 152 | | _ => raise EQVT_FORM "All permutation should be the same") | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 153 | end; | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 154 | |
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 155 | (* Either adds a theorem (orig_thm) to or deletes one from the equivaraince *) | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 156 | (* lemma list depending on flag. To be added the lemma has to satisfy a *) | 
| 22846 | 157 | (* certain form. *) | 
| 24265 
4d5917cc60c4
fix the generation of eqvt lemma of equality form from the imp form when the relation is equality
 narboux parents: 
24039diff
changeset | 158 | |
| 
4d5917cc60c4
fix the generation of eqvt lemma of equality form from the imp form when the relation is equality
 narboux parents: 
24039diff
changeset | 159 | fun eqvt_add_del_aux flag orig_thm context = | 
| 22846 | 160 | let | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 161 | val thy = Context.theory_of context | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 162 | val thms_to_be_added = (case (prop_of orig_thm) of | 
| 22846 | 163 | (* case: eqvt-lemma is of the implicational form *) | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 164 |         (Const("==>", _) $ (Const ("Trueprop",_) $ hyp) $ (Const ("Trueprop",_) $ concl)) =>
 | 
| 22846 | 165 | let | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 166 | val (pi,typi) = get_pi concl thy | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 167 | in | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 168 | if (apply_pi hyp (pi,typi) = concl) | 
| 22846 | 169 | then | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 170 |                (warning ("equivariance lemma of the relational form");
 | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 171 | [orig_thm, get_derived_thm thy hyp concl orig_thm pi typi]) | 
| 22418 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22286diff
changeset | 172 | else raise EQVT_FORM "Type Implication" | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 173 | end | 
| 22846 | 174 | (* case: eqvt-lemma is of the equational form *) | 
| 175 |       | (Const ("Trueprop", _) $ (Const ("op =", _) $
 | |
| 22418 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22286diff
changeset | 176 |             (Const ("Nominal.perm",typrm) $ Var (pi,typi) $ lhs) $ rhs)) =>
 | 
| 24039 
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
22846diff
changeset | 177 | (if (apply_pi lhs (pi,typi)) = rhs | 
| 22437 
b3526cd2a336
removes some unused code that used to try to derive a simpler version of the eqvt lemmas
 narboux parents: 
22418diff
changeset | 178 | then [orig_thm] | 
| 22418 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22286diff
changeset | 179 | else raise EQVT_FORM "Type Equality") | 
| 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22286diff
changeset | 180 | | _ => raise EQVT_FORM "Type unknown") | 
| 22846 | 181 | in | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 182 | update_context flag thms_to_be_added context | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 183 | end | 
| 22846 | 184 | handle EQVT_FORM s => | 
| 22418 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22286diff
changeset | 185 |       error (string_of_thm orig_thm ^ " does not comply with the form of an equivariance lemma ("^s^").")
 | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 186 | |
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 187 | (* in cases of bij- and freshness, we just add the lemmas to the *) | 
| 22846 | 188 | (* data-slot *) | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 189 | |
| 22846 | 190 | fun simple_add_del_aux record_access name flag th context = | 
| 191 | let | |
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 192 | val context' = Data.map (flag th) context | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 193 | in | 
| 22846 | 194 | Context.mapping (snd o PureThy.add_thmss [((name,(record_access (Data.get context'))),[])]) I context' | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 195 | end | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 196 | |
| 22562 
80b814fe284b
rename bij and fresh into bijs and freshs and lookup for eqvts lemma using the ml value instead of the name
 narboux parents: 
22541diff
changeset | 197 | fun bij_add_del_aux f = simple_add_del_aux #bijs "bijs" f | 
| 
80b814fe284b
rename bij and fresh into bijs and freshs and lookup for eqvts lemma using the ml value instead of the name
 narboux parents: 
22541diff
changeset | 198 | fun fresh_add_del_aux f = simple_add_del_aux #freshs "freshs" f | 
| 22541 
c33b542394f3
the name for the collection of equivariance lemmas is now eqvts (changed from eqvt) in order to avoid clashes with eqvt-lemmas generated in nominal_inductive
 urbanc parents: 
22495diff
changeset | 199 | fun eqvt_force_add_del_aux f = simple_add_del_aux #eqvts "eqvts" f | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 200 | |
| 22562 
80b814fe284b
rename bij and fresh into bijs and freshs and lookup for eqvts lemma using the ml value instead of the name
 narboux parents: 
22541diff
changeset | 201 | fun eqvt_map f th (r:Data.T)  = {eqvts = (f th (#eqvts r)), freshs = #freshs r, bijs = #bijs r};
 | 
| 
80b814fe284b
rename bij and fresh into bijs and freshs and lookup for eqvts lemma using the ml value instead of the name
 narboux parents: 
22541diff
changeset | 202 | fun fresh_map f th (r:Data.T) = {eqvts = #eqvts r, freshs = (f th (#freshs r)), bijs = #bijs r};
 | 
| 
80b814fe284b
rename bij and fresh into bijs and freshs and lookup for eqvts lemma using the ml value instead of the name
 narboux parents: 
22541diff
changeset | 203 | fun bij_map f th (r:Data.T)   = {eqvts = #eqvts r, freshs = #freshs r, bijs = (f th (#bijs r))};
 | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 204 | |
| 24039 
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
22846diff
changeset | 205 | val eqvt_add = Thm.declaration_attribute (eqvt_add_del_aux (eqvt_map Thm.add_thm)); | 
| 
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
22846diff
changeset | 206 | val eqvt_del = Thm.declaration_attribute (eqvt_add_del_aux (eqvt_map Thm.del_thm)); | 
| 
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
22846diff
changeset | 207 | val eqvt_force_add = Thm.declaration_attribute (eqvt_force_add_del_aux (eqvt_map Thm.add_thm)); | 
| 
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
22846diff
changeset | 208 | val eqvt_force_del = Thm.declaration_attribute (eqvt_force_add_del_aux (eqvt_map Thm.del_thm)); | 
| 22715 
381e6c45f13b
improved the equivariance lemmas for the quantifiers; had to export the lemma eqvt_force_add and eqvt_force_del in the thmdecls
 urbanc parents: 
22562diff
changeset | 209 | |
| 24039 
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
22846diff
changeset | 210 | val bij_add = Thm.declaration_attribute (bij_add_del_aux (bij_map Thm.add_thm)); | 
| 
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
22846diff
changeset | 211 | val bij_del = Thm.declaration_attribute (bij_add_del_aux (bij_map Thm.del_thm)); | 
| 
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
22846diff
changeset | 212 | val fresh_add = Thm.declaration_attribute (fresh_add_del_aux (fresh_map Thm.add_thm)); | 
| 
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
22846diff
changeset | 213 | val fresh_del = Thm.declaration_attribute (fresh_add_del_aux (fresh_map Thm.del_thm)); | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 214 | |
| 22715 
381e6c45f13b
improved the equivariance lemmas for the quantifiers; had to export the lemma eqvt_force_add and eqvt_force_del in the thmdecls
 urbanc parents: 
22562diff
changeset | 215 | |
| 22418 
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
 urbanc parents: 
22286diff
changeset | 216 | |
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 217 | val setup = | 
| 22846 | 218 | Attrib.add_attributes | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 219 |      [("eqvt",  Attrib.add_del_args eqvt_add eqvt_del,   "equivariance theorem declaration"),
 | 
| 22715 
381e6c45f13b
improved the equivariance lemmas for the quantifiers; had to export the lemma eqvt_force_add and eqvt_force_del in the thmdecls
 urbanc parents: 
22562diff
changeset | 220 |       ("eqvt_force", Attrib.add_del_args eqvt_force_add eqvt_force_del,
 | 
| 24039 
273698405054
renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
22846diff
changeset | 221 | "equivariance theorem declaration (without checking the form of the lemma)"), | 
| 22245 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 222 |       ("fresh", Attrib.add_del_args fresh_add fresh_del, "freshness theorem declaration"),
 | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 223 |       ("bij",   Attrib.add_del_args bij_add bij_del,     "bijection theorem declaration")];
 | 
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 224 | |
| 
1b8f4ef50c48
moved the infrastructure from the nominal_tags file to nominal_thmdecls
 urbanc parents: diff
changeset | 225 | end; |