author | Christian Urban <urbanc@in.tum.de> |
Mon, 24 May 2010 21:11:33 +0100 | |
changeset 2298 | aead2aad845c |
parent 2297 | 9ca7b249760e |
child 2300 | 9fb315392493 |
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" |
2297 | 5 |
"Nominal2_FSet" "Abs" |
2296 | 6 |
uses ("nominal_dt_rawperm.ML") |
7 |
("nominal_dt_rawfuns.ML") |
|
2297 | 8 |
("nominal_dt_alpha.ML") |
1159
3c6bee89d826
Ported Stefan's permutation code, still needs some localizing.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
9 |
begin |
3c6bee89d826
Ported Stefan's permutation code, still needs some localizing.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
10 |
|
2296 | 11 |
use "nominal_dt_rawperm.ML" |
2297 | 12 |
ML {* open Nominal_Dt_RawPerm *} |
2296 | 13 |
|
2297 | 14 |
use "nominal_dt_rawfuns.ML" |
15 |
ML {* open Nominal_Dt_RawFuns *} |
|
2296 | 16 |
|
2297 | 17 |
use "nominal_dt_alpha.ML" |
18 |
ML {* open Nominal_Dt_Alpha *} |
|
2296 | 19 |
|
1910 | 20 |
(* permutations for quotient types *) |
1899
8e0bfb14f6bf
optimised the code of define_raw_perm
Christian Urban <urbanc@in.tum.de>
parents:
1896
diff
changeset
|
21 |
|
1253 | 22 |
ML {* |
1903 | 23 |
fun quotient_lift_consts_export qtys spec ctxt = |
24 |
let |
|
25 |
val (result, ctxt') = fold_map (Quotient_Def.quotient_lift_const qtys) spec ctxt; |
|
26 |
val (ts_loc, defs_loc) = split_list result; |
|
27 |
val morphism = ProofContext.export_morphism ctxt' ctxt; |
|
28 |
val ts = map (Morphism.term morphism) ts_loc |
|
29 |
val defs = Morphism.fact morphism defs_loc |
|
30 |
in |
|
31 |
(ts, defs, ctxt') |
|
32 |
end |
|
33 |
*} |
|
34 |
||
35 |
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
|
36 |
fun define_lifted_perms qtys full_tnames name_term_pairs thms thy = |
1253 | 37 |
let |
38 |
val lthy = |
|
39 |
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
|
40 |
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
|
41 |
val lifted_thms = map (Quotient_Tacs.lifted qtys lthy') thms; |
1253 | 42 |
fun tac _ = |
43 |
Class.intro_classes_tac [] THEN |
|
44 |
(ALLGOALS (resolve_tac lifted_thms)) |
|
45 |
val lthy'' = Class.prove_instantiation_instance tac lthy' |
|
46 |
in |
|
47 |
Local_Theory.exit_global lthy'' |
|
48 |
end |
|
49 |
*} |
|
50 |
||
1342 | 51 |
ML {* |
52 |
fun neq_to_rel r neq = |
|
53 |
let |
|
54 |
val neq = HOLogic.dest_Trueprop (prop_of neq) |
|
55 |
val eq = HOLogic.dest_not neq |
|
56 |
val (lhs, rhs) = HOLogic.dest_eq eq |
|
57 |
val rel = r $ lhs $ rhs |
|
58 |
val nrel = HOLogic.mk_not rel |
|
59 |
in |
|
60 |
HOLogic.mk_Trueprop nrel |
|
61 |
end |
|
62 |
*} |
|
63 |
||
64 |
ML {* |
|
65 |
fun neq_to_rel_tac cases distinct = |
|
66 |
rtac notI THEN' eresolve_tac cases THEN_ALL_NEW asm_full_simp_tac (HOL_ss addsimps distinct) |
|
67 |
*} |
|
68 |
||
69 |
ML {* |
|
70 |
fun distinct_rel ctxt cases (dists, rel) = |
|
71 |
let |
|
72 |
val ((_, thms), ctxt') = Variable.import false dists ctxt |
|
73 |
val terms = map (neq_to_rel rel) thms |
|
74 |
val nrels = map (fn t => Goal.prove ctxt' [] [] t (fn _ => neq_to_rel_tac cases dists 1)) terms |
|
75 |
in |
|
76 |
Variable.export ctxt' ctxt nrels |
|
77 |
end |
|
78 |
*} |
|
79 |
||
1899
8e0bfb14f6bf
optimised the code of define_raw_perm
Christian Urban <urbanc@in.tum.de>
parents:
1896
diff
changeset
|
80 |
end |