author | Christian Urban <christian dot urban at kcl dot ac dot uk> |
Sat, 19 Mar 2016 21:06:48 +0000 | |
branch | Nominal2-Isabelle2016 |
changeset 3243 | c4f31f1564b7 |
parent 3239 | 67370521c09c |
permissions | -rw-r--r-- |
1833
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1 |
(* Title: nominal_eqvt.ML |
1835
636de31888a6
tuned and removed dead code
Christian Urban <urbanc@in.tum.de>
parents:
1833
diff
changeset
|
2 |
Author: Stefan Berghofer (original code) |
1833
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
3 |
Author: Christian Urban |
3214
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
4 |
Author: Tjark Weber |
1833
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
5 |
|
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
6 |
Automatic proofs for equivariance of inductive predicates. |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
7 |
*) |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
8 |
|
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
9 |
signature NOMINAL_EQVT = |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
10 |
sig |
3214
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
11 |
val raw_equivariance: Proof.context -> term list -> thm -> thm list -> thm list |
2107
5686d83db1f9
ingnored parameters in equivariance; added a proper interface to be called from ML
Christian Urban <urbanc@in.tum.de>
parents:
2081
diff
changeset
|
12 |
val equivariance_cmd: string -> Proof.context -> local_theory |
1833
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
13 |
end |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
14 |
|
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
15 |
structure Nominal_Eqvt : NOMINAL_EQVT = |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
16 |
struct |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
17 |
|
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
18 |
open Nominal_Permeq; |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
19 |
open Nominal_ThmDecls; |
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
20 |
|
3218
89158f401b07
updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3214
diff
changeset
|
21 |
fun atomize_conv ctxt = |
2620
81921f8ad245
updated to Isabelle 22 December
Christian Urban <urbanc@in.tum.de>
parents:
2568
diff
changeset
|
22 |
Raw_Simplifier.rewrite_cterm (true, false, false) (K (K NONE)) |
3218
89158f401b07
updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3214
diff
changeset
|
23 |
(put_simpset HOL_basic_ss ctxt addsimps @{thms induct_atomize}) |
3214
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
24 |
|
3218
89158f401b07
updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3214
diff
changeset
|
25 |
fun atomize_intr ctxt = Conv.fconv_rule (Conv.prems_conv ~1 (atomize_conv ctxt)) |
3214
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
26 |
|
1833
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
27 |
fun atomize_induct ctxt = Conv.fconv_rule (Conv.prems_conv ~1 |
3218
89158f401b07
updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3214
diff
changeset
|
28 |
(Conv.params_conv ~1 (K (Conv.prems_conv ~1 (atomize_conv ctxt))) ctxt)) |
1833
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
29 |
|
1835
636de31888a6
tuned and removed dead code
Christian Urban <urbanc@in.tum.de>
parents:
1833
diff
changeset
|
30 |
|
636de31888a6
tuned and removed dead code
Christian Urban <urbanc@in.tum.de>
parents:
1833
diff
changeset
|
31 |
(** equivariance tactics **) |
636de31888a6
tuned and removed dead code
Christian Urban <urbanc@in.tum.de>
parents:
1833
diff
changeset
|
32 |
|
3214
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
33 |
fun eqvt_rel_single_case_tac ctxt pred_names pi intro = |
2477 | 34 |
let |
3239
67370521c09c
updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3218
diff
changeset
|
35 |
val cpi = Thm.cterm_of ctxt pi |
3243
c4f31f1564b7
updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3239
diff
changeset
|
36 |
val pi_intro_rule = Thm.instantiate' [] [NONE, SOME cpi] @{thm permute_boolI} |
3090
19f5e7afad89
fixed problem with equivariance for beta_star
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
37 |
val eqvt_sconfig = eqvt_strict_config addexcls pred_names |
3218
89158f401b07
updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3214
diff
changeset
|
38 |
val simps1 = |
89158f401b07
updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3214
diff
changeset
|
39 |
put_simpset HOL_basic_ss ctxt addsimps @{thms permute_fun_def permute_self split_paired_all} |
89158f401b07
updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3214
diff
changeset
|
40 |
val simps2 = |
89158f401b07
updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3214
diff
changeset
|
41 |
put_simpset HOL_basic_ss ctxt addsimps @{thms permute_bool_def permute_minus_cancel(2)} |
2477 | 42 |
in |
3090
19f5e7afad89
fixed problem with equivariance for beta_star
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
43 |
eqvt_tac ctxt eqvt_sconfig THEN' |
2477 | 44 |
SUBPROOF (fn {prems, context as ctxt, ...} => |
45 |
let |
|
46 |
val prems' = map (transform_prem2 ctxt pred_names) prems |
|
3090
19f5e7afad89
fixed problem with equivariance for beta_star
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
47 |
val prems'' = map (fn thm => eqvt_rule ctxt eqvt_sconfig (thm RS pi_intro_rule)) prems' |
19f5e7afad89
fixed problem with equivariance for beta_star
Christian Urban <urbanc@in.tum.de>
parents:
3045
diff
changeset
|
48 |
val prems''' = map (simplify simps2 o simplify simps1) prems'' |
2477 | 49 |
in |
3243
c4f31f1564b7
updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3239
diff
changeset
|
50 |
HEADGOAL (resolve_tac ctxt [intro] THEN_ALL_NEW |
c4f31f1564b7
updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3239
diff
changeset
|
51 |
resolve_tac ctxt (prems' @ prems'' @ prems''')) |
2477 | 52 |
end) ctxt |
53 |
end |
|
1833
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
54 |
|
1835
636de31888a6
tuned and removed dead code
Christian Urban <urbanc@in.tum.de>
parents:
1833
diff
changeset
|
55 |
fun eqvt_rel_tac ctxt pred_names pi induct intros = |
2477 | 56 |
let |
57 |
val cases = map (eqvt_rel_single_case_tac ctxt pred_names pi) intros |
|
58 |
in |
|
3243
c4f31f1564b7
updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3239
diff
changeset
|
59 |
EVERY' ((DETERM o resolve_tac ctxt [induct]) :: cases) |
2477 | 60 |
end |
1833
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
61 |
|
1835
636de31888a6
tuned and removed dead code
Christian Urban <urbanc@in.tum.de>
parents:
1833
diff
changeset
|
62 |
|
3214
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
63 |
(** equivariance procedure **) |
1835
636de31888a6
tuned and removed dead code
Christian Urban <urbanc@in.tum.de>
parents:
1833
diff
changeset
|
64 |
|
3214
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
65 |
fun prepare_goal ctxt pi pred_with_args = |
2477 | 66 |
let |
3214
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
67 |
val (c, xs) = strip_comb pred_with_args |
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
68 |
fun is_nonfixed_Free (Free (s, _)) = not (Variable.is_fixed ctxt s) |
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
69 |
| is_nonfixed_Free _ = false |
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
70 |
fun mk_perm_nonfixed_Free t = |
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
71 |
if is_nonfixed_Free t then mk_perm pi t else t |
2477 | 72 |
in |
3214
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
73 |
HOLogic.mk_imp (pred_with_args, |
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
74 |
list_comb (c, map mk_perm_nonfixed_Free xs)) |
2477 | 75 |
end |
1833
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
76 |
|
3214
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
77 |
fun name_of (Const (s, _)) = s |
2477 | 78 |
|
3214
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
79 |
fun raw_equivariance ctxt preds raw_induct intrs = |
2477 | 80 |
let |
3214
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
81 |
(* FIXME: polymorphic predicates should either be rejected or |
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
82 |
specialized to arguments of sort pt *) |
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
83 |
|
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
84 |
val is_already_eqvt = filter (is_eqvt ctxt) preds |
2477 | 85 |
val _ = if null is_already_eqvt then () |
3214
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
86 |
else error ("Already equivariant: " ^ commas |
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
87 |
(map (Syntax.string_of_term ctxt) is_already_eqvt)) |
2117
b3a5bda07007
added a test whether some of the constants already equivariant (then the procedure has to fail).
Christian Urban <urbanc@in.tum.de>
parents:
2110
diff
changeset
|
88 |
|
3214
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
89 |
val pred_names = map (name_of o head_of) preds |
2477 | 90 |
val raw_induct' = atomize_induct ctxt raw_induct |
3218
89158f401b07
updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3214
diff
changeset
|
91 |
val intrs' = map (atomize_intr ctxt) intrs |
3214
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
92 |
|
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
93 |
val (([raw_concl], [raw_pi]), ctxt') = |
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
94 |
ctxt |
3239
67370521c09c
updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3218
diff
changeset
|
95 |
|> Variable.import_terms false [Thm.concl_of raw_induct'] |
2477 | 96 |
||>> Variable.variant_fixes ["p"] |
97 |
val pi = Free (raw_pi, @{typ perm}) |
|
3214
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
98 |
|
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
99 |
val preds_with_args = raw_concl |
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
100 |
|> HOLogic.dest_Trueprop |
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
101 |
|> HOLogic.dest_conj |
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
102 |
|> map (fst o HOLogic.dest_imp) |
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
103 |
|
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
104 |
val goal = preds_with_args |
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
105 |
|> map (prepare_goal ctxt pi) |
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
106 |
|> foldr1 HOLogic.mk_conj |
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
107 |
|> HOLogic.mk_Trueprop |
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
108 |
in |
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
109 |
Goal.prove ctxt' [] [] goal |
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
110 |
(fn {context, ...} => eqvt_rel_tac context pred_names pi raw_induct' intrs' 1) |
3239
67370521c09c
updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3218
diff
changeset
|
111 |
|> Old_Datatype_Aux.split_conj_thm |
3045
d0ad264f8c4f
updated to Isabelle 3 Nov; it includes a hack to work around a bug in the localised version of the quotient package
Christian Urban <urbanc@in.tum.de>
parents:
2885
diff
changeset
|
112 |
|> Proof_Context.export ctxt' ctxt |
2477 | 113 |
|> map (fn th => th RS mp) |
114 |
|> map zero_var_indexes |
|
115 |
end |
|
1833
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
116 |
|
2885
1264f2a21ea9
some rudimentary infrastructure for storing data about nominal datatypes
Christian Urban <urbanc@in.tum.de>
parents:
2868
diff
changeset
|
117 |
|
3214
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
118 |
(** stores thm under name.eqvt and adds [eqvt]-attribute **) |
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
119 |
|
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
120 |
fun note_named_thm (name, thm) ctxt = |
2650
e5fa8de0e4bd
derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
parents:
2620
diff
changeset
|
121 |
let |
3214
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
122 |
val thm_name = Binding.qualified_name |
2868
2b8e387d2dfc
got rid of the boolean flag in the raw_equivariance function
Christian Urban <urbanc@in.tum.de>
parents:
2778
diff
changeset
|
123 |
(Long_Name.qualify (Long_Name.base_name name) "eqvt") |
2b8e387d2dfc
got rid of the boolean flag in the raw_equivariance function
Christian Urban <urbanc@in.tum.de>
parents:
2778
diff
changeset
|
124 |
val attr = Attrib.internal (K eqvt_add) |
3214
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
125 |
val ((_, [thm']), ctxt') = Local_Theory.note ((thm_name, [attr]), [thm]) ctxt |
2650
e5fa8de0e4bd
derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
parents:
2620
diff
changeset
|
126 |
in |
2868
2b8e387d2dfc
got rid of the boolean flag in the raw_equivariance function
Christian Urban <urbanc@in.tum.de>
parents:
2778
diff
changeset
|
127 |
(thm', ctxt') |
2650
e5fa8de0e4bd
derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
parents:
2620
diff
changeset
|
128 |
end |
e5fa8de0e4bd
derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
parents:
2620
diff
changeset
|
129 |
|
3214
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
130 |
|
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
131 |
(** equivariance command **) |
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
132 |
|
2107
5686d83db1f9
ingnored parameters in equivariance; added a proper interface to be called from ML
Christian Urban <urbanc@in.tum.de>
parents:
2081
diff
changeset
|
133 |
fun equivariance_cmd pred_name ctxt = |
2477 | 134 |
let |
2868
2b8e387d2dfc
got rid of the boolean flag in the raw_equivariance function
Christian Urban <urbanc@in.tum.de>
parents:
2778
diff
changeset
|
135 |
val ({names, ...}, {preds, raw_induct, intrs, ...}) = |
3214
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
136 |
Inductive.the_inductive ctxt (long_name ctxt pred_name) |
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
137 |
val thms = raw_equivariance ctxt preds raw_induct intrs |
2477 | 138 |
in |
2868
2b8e387d2dfc
got rid of the boolean flag in the raw_equivariance function
Christian Urban <urbanc@in.tum.de>
parents:
2778
diff
changeset
|
139 |
fold_map note_named_thm (names ~~ thms) ctxt |> snd |
2477 | 140 |
end |
1833
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
141 |
|
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
142 |
val _ = |
3239
67370521c09c
updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3218
diff
changeset
|
143 |
Outer_Syntax.local_theory @{command_keyword equivariance} |
3214
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
144 |
"Proves equivariance for inductive predicate involving nominal datatypes." |
13ab4f0a0b0e
Various changes to support Nominal2 commands in local contexts.
webertj
parents:
3193
diff
changeset
|
145 |
(Parse.const >> equivariance_cmd) |
1833
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
146 |
|
2069
2b6ba4d4e19a
Fixes for new isabelle
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2064
diff
changeset
|
147 |
end (* structure *) |