author | Christian Urban <urbanc@in.tum.de> |
Tue, 01 Jun 2010 15:01:05 +0200 | |
changeset 2305 | 93ab397f5980 |
parent 2302 | c6db12ddb60c |
child 2335 | 558c823f96aa |
permissions | -rw-r--r-- |
1159
3c6bee89d826
Ported Stefan's permutation code, still needs some localizing.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
1 |
theory Perm |
2288
3b83960f9544
new fv/fv_bn function (supp breaks now); exported raw perms and raw funs into separate ML-files
Christian Urban <urbanc@in.tum.de>
parents:
2163
diff
changeset
|
2 |
imports |
3b83960f9544
new fv/fv_bn function (supp breaks now); exported raw perms and raw funs into separate ML-files
Christian Urban <urbanc@in.tum.de>
parents:
2163
diff
changeset
|
3 |
"../Nominal-General/Nominal2_Base" |
2296 | 4 |
"../Nominal-General/Nominal2_Atoms" |
2302 | 5 |
"../Nominal-General/Nominal2_Eqvt" |
6 |
"Nominal2_FSet" |
|
7 |
"Abs" |
|
2305
93ab397f5980
smaller code for raw-eqvt proofs
Christian Urban <urbanc@in.tum.de>
parents:
2302
diff
changeset
|
8 |
uses ("nominal_dt_rawperm.ML") |
2296 | 9 |
("nominal_dt_rawfuns.ML") |
2305
93ab397f5980
smaller code for raw-eqvt proofs
Christian Urban <urbanc@in.tum.de>
parents:
2302
diff
changeset
|
10 |
("nominal_dt_alpha.ML") |
1159
3c6bee89d826
Ported Stefan's permutation code, still needs some localizing.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
11 |
begin |
3c6bee89d826
Ported Stefan's permutation code, still needs some localizing.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
12 |
|
2296 | 13 |
use "nominal_dt_rawperm.ML" |
2297 | 14 |
ML {* open Nominal_Dt_RawPerm *} |
2296 | 15 |
|
2297 | 16 |
use "nominal_dt_rawfuns.ML" |
17 |
ML {* open Nominal_Dt_RawFuns *} |
|
2296 | 18 |
|
2297 | 19 |
use "nominal_dt_alpha.ML" |
20 |
ML {* open Nominal_Dt_Alpha *} |
|
2296 | 21 |
|
1910 | 22 |
(* permutations for quotient types *) |
1899
8e0bfb14f6bf
optimised the code of define_raw_perm
Christian Urban <urbanc@in.tum.de>
parents:
1896
diff
changeset
|
23 |
|
1253 | 24 |
ML {* |
1903 | 25 |
fun quotient_lift_consts_export qtys spec ctxt = |
26 |
let |
|
27 |
val (result, ctxt') = fold_map (Quotient_Def.quotient_lift_const qtys) spec ctxt; |
|
28 |
val (ts_loc, defs_loc) = split_list result; |
|
29 |
val morphism = ProofContext.export_morphism ctxt' ctxt; |
|
30 |
val ts = map (Morphism.term morphism) ts_loc |
|
31 |
val defs = Morphism.fact morphism defs_loc |
|
32 |
in |
|
33 |
(ts, defs, ctxt') |
|
34 |
end |
|
35 |
*} |
|
36 |
||
37 |
ML {* |
|
1683
f78c820f67c3
Automatically lift theorems and constants only using the new quotient types. Requires new Isabelle.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1503
diff
changeset
|
38 |
fun define_lifted_perms qtys full_tnames name_term_pairs thms thy = |
1253 | 39 |
let |
40 |
val lthy = |
|
41 |
Theory_Target.instantiation (full_tnames, [], @{sort pt}) thy; |
|
1683
f78c820f67c3
Automatically lift theorems and constants only using the new quotient types. Requires new Isabelle.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1503
diff
changeset
|
42 |
val (_, _, lthy') = quotient_lift_consts_export qtys name_term_pairs lthy; |
f78c820f67c3
Automatically lift theorems and constants only using the new quotient types. Requires new Isabelle.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1503
diff
changeset
|
43 |
val lifted_thms = map (Quotient_Tacs.lifted qtys lthy') thms; |
1253 | 44 |
fun tac _ = |
45 |
Class.intro_classes_tac [] THEN |
|
46 |
(ALLGOALS (resolve_tac lifted_thms)) |
|
47 |
val lthy'' = Class.prove_instantiation_instance tac lthy' |
|
48 |
in |
|
49 |
Local_Theory.exit_global lthy'' |
|
50 |
end |
|
51 |
*} |
|
52 |
||
1342 | 53 |
|
1899
8e0bfb14f6bf
optimised the code of define_raw_perm
Christian Urban <urbanc@in.tum.de>
parents:
1896
diff
changeset
|
54 |
end |