author | Christian Urban <urbanc@in.tum.de> |
Fri, 26 Nov 2010 10:53:55 +0000 | |
changeset 2583 | cb16f22bc660 |
parent 2476 | 8f8652a8107f |
permissions | -rw-r--r-- |
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:
diff
changeset
|
1 |
(* Title: nominal_dt_rawperm.ML |
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:
diff
changeset
|
2 |
Author: Cezary Kaliszyk |
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:
diff
changeset
|
3 |
Author: Christian Urban |
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:
diff
changeset
|
4 |
|
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:
diff
changeset
|
5 |
Definitions of the raw permutations and |
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:
diff
changeset
|
6 |
proof that the raw datatypes are in the |
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:
diff
changeset
|
7 |
pt-class. |
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:
diff
changeset
|
8 |
*) |
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:
diff
changeset
|
9 |
|
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:
diff
changeset
|
10 |
signature NOMINAL_DT_RAWPERM = |
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:
diff
changeset
|
11 |
sig |
2431
331873ebc5cd
can now deal with type variables in nominal datatype definitions
Christian Urban <urbanc@in.tum.de>
parents:
2401
diff
changeset
|
12 |
val define_raw_perms: string list -> typ list -> (string * sort) list -> term list -> thm -> |
331873ebc5cd
can now deal with type variables in nominal datatype definitions
Christian Urban <urbanc@in.tum.de>
parents:
2401
diff
changeset
|
13 |
local_theory -> (term list * thm list * thm list) * local_theory |
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:
diff
changeset
|
14 |
end |
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:
diff
changeset
|
15 |
|
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:
diff
changeset
|
16 |
|
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:
diff
changeset
|
17 |
structure Nominal_Dt_RawPerm: NOMINAL_DT_RAWPERM = |
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:
diff
changeset
|
18 |
struct |
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:
diff
changeset
|
19 |
|
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:
diff
changeset
|
20 |
|
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:
diff
changeset
|
21 |
(** proves the two pt-type class properties **) |
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:
diff
changeset
|
22 |
|
2401
7645e18e8b19
modified the code for class instantiations (with help from Florian)
Christian Urban <urbanc@in.tum.de>
parents:
2398
diff
changeset
|
23 |
fun prove_permute_zero induct perm_defs perm_fns lthy = |
2476
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
24 |
let |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
25 |
val perm_types = map (body_type o fastype_of) perm_fns |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
26 |
val perm_indnames = Datatype_Prop.make_tnames perm_types |
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:
diff
changeset
|
27 |
|
2476
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
28 |
fun single_goal ((perm_fn, T), x) = |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
29 |
HOLogic.mk_eq (perm_fn $ @{term "0::perm"} $ Free (x, T), Free (x, T)) |
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:
diff
changeset
|
30 |
|
2476
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
31 |
val goals = |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
32 |
HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
33 |
(map single_goal (perm_fns ~~ perm_types ~~ perm_indnames))) |
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:
diff
changeset
|
34 |
|
2476
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
35 |
val simps = HOL_basic_ss addsimps (@{thm permute_zero} :: perm_defs) |
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:
diff
changeset
|
36 |
|
2476
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
37 |
val tac = (Datatype_Aux.indtac induct perm_indnames |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
38 |
THEN_ALL_NEW asm_simp_tac simps) 1 |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
39 |
in |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
40 |
Goal.prove lthy perm_indnames [] goals (K tac) |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
41 |
|> Datatype_Aux.split_conj_thm |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
42 |
end |
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:
diff
changeset
|
43 |
|
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:
diff
changeset
|
44 |
|
2401
7645e18e8b19
modified the code for class instantiations (with help from Florian)
Christian Urban <urbanc@in.tum.de>
parents:
2398
diff
changeset
|
45 |
fun prove_permute_plus induct perm_defs perm_fns lthy = |
2476
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
46 |
let |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
47 |
val p = Free ("p", @{typ perm}) |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
48 |
val q = Free ("q", @{typ perm}) |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
49 |
val perm_types = map (body_type o fastype_of) perm_fns |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
50 |
val perm_indnames = Datatype_Prop.make_tnames perm_types |
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:
diff
changeset
|
51 |
|
2476
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
52 |
fun single_goal ((perm_fn, T), x) = HOLogic.mk_eq |
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:
diff
changeset
|
53 |
(perm_fn $ (mk_plus p q) $ Free (x, T), perm_fn $ p $ (perm_fn $ q $ Free (x, T))) |
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:
diff
changeset
|
54 |
|
2476
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
55 |
val goals = |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
56 |
HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
57 |
(map single_goal (perm_fns ~~ perm_types ~~ perm_indnames))) |
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:
diff
changeset
|
58 |
|
2476
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
59 |
val simps = HOL_basic_ss addsimps (@{thm permute_plus} :: perm_defs) |
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:
diff
changeset
|
60 |
|
2476
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
61 |
val tac = (Datatype_Aux.indtac induct perm_indnames |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
62 |
THEN_ALL_NEW asm_simp_tac simps) 1 |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
63 |
in |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
64 |
Goal.prove lthy ("p" :: "q" :: perm_indnames) [] goals (K tac) |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
65 |
|> Datatype_Aux.split_conj_thm |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
66 |
end |
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:
diff
changeset
|
67 |
|
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:
diff
changeset
|
68 |
|
2398 | 69 |
fun mk_perm_eq ty_perm_assoc cnstr = |
2476
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
70 |
let |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
71 |
fun lookup_perm p (ty, arg) = |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
72 |
case (AList.lookup (op=) ty_perm_assoc ty) of |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
73 |
SOME perm => perm $ p $ arg |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
74 |
| NONE => Const (@{const_name permute}, perm_ty ty) $ p $ arg |
2398 | 75 |
|
2476
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
76 |
val p = Free ("p", @{typ perm}) |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
77 |
val (arg_tys, ty) = |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
78 |
fastype_of cnstr |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
79 |
|> strip_type |
2398 | 80 |
|
2476
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
81 |
val arg_names = Name.variant_list ["p"] (Datatype_Prop.make_tnames arg_tys) |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
82 |
val args = map Free (arg_names ~~ arg_tys) |
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:
diff
changeset
|
83 |
|
2476
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
84 |
val lhs = lookup_perm p (ty, list_comb (cnstr, args)) |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
85 |
val rhs = list_comb (cnstr, map (lookup_perm p) (arg_tys ~~ args)) |
2398 | 86 |
|
2476
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
87 |
val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)) |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
88 |
in |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
89 |
(Attrib.empty_binding, eq) |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
90 |
end |
2398 | 91 |
|
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:
diff
changeset
|
92 |
|
2431
331873ebc5cd
can now deal with type variables in nominal datatype definitions
Christian Urban <urbanc@in.tum.de>
parents:
2401
diff
changeset
|
93 |
fun define_raw_perms full_ty_names tys tvs constrs induct_thm lthy = |
2476
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
94 |
let |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
95 |
val perm_fn_names = full_ty_names |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
96 |
|> map Long_Name.base_name |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
97 |
|> map (prefix "permute_") |
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:
diff
changeset
|
98 |
|
2476
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
99 |
val perm_fn_types = map perm_ty tys |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
100 |
val perm_fn_frees = map Free (perm_fn_names ~~ perm_fn_types) |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
101 |
val perm_fn_binds = map (fn s => (Binding.name s, NONE, NoSyn)) perm_fn_names |
2398 | 102 |
|
2476
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
103 |
val perm_eqs = map (mk_perm_eq (tys ~~ perm_fn_frees)) constrs |
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:
diff
changeset
|
104 |
|
2476
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
105 |
fun tac _ (_, _, simps) = |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
106 |
Class.intro_classes_tac [] THEN ALLGOALS (resolve_tac simps) |
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:
diff
changeset
|
107 |
|
2476
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
108 |
fun morphism phi (fvs, dfs, simps) = |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
109 |
(map (Morphism.term phi) fvs, |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
110 |
map (Morphism.thm phi) dfs, |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
111 |
map (Morphism.thm phi) simps); |
2401
7645e18e8b19
modified the code for class instantiations (with help from Florian)
Christian Urban <urbanc@in.tum.de>
parents:
2398
diff
changeset
|
112 |
|
2476
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
113 |
val ((perm_funs, perm_eq_thms), lthy') = |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
114 |
lthy |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
115 |
|> Local_Theory.exit_global |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
116 |
|> Class.instantiation (full_ty_names, tvs, @{sort pt}) |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
117 |
|> Primrec.add_primrec perm_fn_binds perm_eqs |
2401
7645e18e8b19
modified the code for class instantiations (with help from Florian)
Christian Urban <urbanc@in.tum.de>
parents:
2398
diff
changeset
|
118 |
|
2476
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
119 |
val perm_zero_thms = prove_permute_zero induct_thm perm_eq_thms perm_funs lthy' |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
120 |
val perm_plus_thms = prove_permute_plus induct_thm perm_eq_thms perm_funs lthy' |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
121 |
in |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
122 |
lthy' |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
123 |
|> Class.prove_instantiation_exit_result morphism tac |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
124 |
(perm_funs, perm_eq_thms, perm_zero_thms @ perm_plus_thms) |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
125 |
||> Named_Target.theory_init |
8f8652a8107f
tuned (to conform with indentation policy of Markus)
Christian Urban <urbanc@in.tum.de>
parents:
2431
diff
changeset
|
126 |
end |
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:
diff
changeset
|
127 |
|
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:
diff
changeset
|
128 |
|
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:
diff
changeset
|
129 |
end (* structure *) |
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:
diff
changeset
|
130 |