author | traytel |
Mon, 13 Apr 2015 13:03:41 +0200 | |
changeset 60046 | 894d6d863823 |
parent 59644 | cc78fd8d955d |
child 60825 | bacfb7c45d81 |
permissions | -rw-r--r-- |
54701 | 1 |
(* Title: HOL/Tools/Ctr_Sugar/ctr_sugar_code.ML |
54615
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
2 |
Author: Jasmin Blanchette, TU Muenchen |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
3 |
Author: Dmitriy Traytel, TU Muenchen |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
4 |
Author: Stefan Berghofer, TU Muenchen |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
5 |
Author: Florian Haftmann, TU Muenchen |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
6 |
Copyright 2001-2013 |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
7 |
|
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
8 |
Code generation for freely generated types. |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
9 |
*) |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
10 |
|
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
11 |
signature CTR_SUGAR_CODE = |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
12 |
sig |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
13 |
val add_ctr_code: string -> typ list -> (string * typ) list -> thm list -> thm list -> thm list -> |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
14 |
theory -> theory |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
15 |
end; |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
16 |
|
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
17 |
structure Ctr_Sugar_Code : CTR_SUGAR_CODE = |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
18 |
struct |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
19 |
|
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
20 |
open Ctr_Sugar_Util |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
21 |
|
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
22 |
val eqN = "eq" |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
23 |
val reflN = "refl" |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
24 |
val simpsN = "simps" |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
25 |
|
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
26 |
fun mk_case_certificate thy raw_thms = |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
27 |
let |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
28 |
val thms as thm1 :: _ = raw_thms |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
29 |
|> Conjunction.intr_balanced |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
30 |
|> Thm.unvarify_global |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
31 |
|> Conjunction.elim_balanced (length raw_thms) |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
32 |
|> map Simpdata.mk_meta_eq |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
33 |
|> map Drule.zero_var_indexes; |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
34 |
val params = Term.add_free_names (Thm.prop_of thm1) []; |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
35 |
val rhs = thm1 |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
36 |
|> Thm.prop_of |> Logic.dest_equals |> fst |> Term.strip_comb |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
37 |
||> fst o split_last |> list_comb; |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
38 |
val lhs = Free (singleton (Name.variant_list params) "case", Term.fastype_of rhs); |
59621
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents:
59582
diff
changeset
|
39 |
val assum = Thm.global_cterm_of thy (Logic.mk_equals (lhs, rhs)); |
54615
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
40 |
in |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
41 |
thms |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
42 |
|> Conjunction.intr_balanced |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
54701
diff
changeset
|
43 |
|> rewrite_rule (Proof_Context.init_global thy) [Thm.symmetric (Thm.assume assum)] |
54615
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
44 |
|> Thm.implies_intr assum |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
45 |
|> Thm.generalize ([], params) 0 |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
46 |
|> Axclass.unoverload thy |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
47 |
|> Thm.varifyT_global |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
48 |
end; |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
49 |
|
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
50 |
fun mk_free_ctr_equations fcT ctrs inject_thms distinct_thms thy = |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
51 |
let |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
52 |
fun mk_fcT_eq (t, u) = Const (@{const_name HOL.equal}, fcT --> fcT --> HOLogic.boolT) $ t $ u; |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
53 |
fun true_eq tu = HOLogic.mk_eq (mk_fcT_eq tu, @{term True}); |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
54 |
fun false_eq tu = HOLogic.mk_eq (mk_fcT_eq tu, @{term False}); |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
55 |
|
59582 | 56 |
val monomorphic_prop_of = Thm.prop_of o Thm.unvarify_global o Drule.zero_var_indexes; |
54615
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
57 |
|
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
58 |
fun massage_inject (tp $ (eqv $ (_ $ t $ u) $ rhs)) = tp $ (eqv $ mk_fcT_eq (t, u) $ rhs); |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
59 |
fun massage_distinct (tp $ (_ $ (_ $ t $ u))) = [tp $ false_eq (t, u), tp $ false_eq (u, t)]; |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
60 |
|
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
61 |
val triv_inject_goals = |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
62 |
map_filter (fn c as (_, T) => |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
63 |
if T = fcT then SOME (HOLogic.mk_Trueprop (true_eq (Const c, Const c))) else NONE) |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
64 |
ctrs; |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
65 |
val inject_goals = map (massage_inject o monomorphic_prop_of) inject_thms; |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
66 |
val distinct_goals = maps (massage_distinct o monomorphic_prop_of) distinct_thms; |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
67 |
val refl_goal = HOLogic.mk_Trueprop (true_eq (Free ("x", fcT), Free ("x", fcT))); |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
68 |
|
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
69 |
fun prove goal = |
54895 | 70 |
Goal.prove_sorry_global thy [] [] goal (fn {context = ctxt, ...} => |
60046
894d6d863823
call Goal.prove only once for a quadratic number of theorems
traytel
parents:
59644
diff
changeset
|
71 |
HEADGOAL Goal.conjunction_tac THEN |
54895 | 72 |
ALLGOALS (simp_tac |
73 |
(put_simpset HOL_basic_ss ctxt |
|
60046
894d6d863823
call Goal.prove only once for a quadratic number of theorems
traytel
parents:
59644
diff
changeset
|
74 |
addsimps (map Simpdata.mk_eq (@{thms equal eq_True} @ inject_thms @ distinct_thms))))); |
894d6d863823
call Goal.prove only once for a quadratic number of theorems
traytel
parents:
59644
diff
changeset
|
75 |
|
894d6d863823
call Goal.prove only once for a quadratic number of theorems
traytel
parents:
59644
diff
changeset
|
76 |
fun proves goals = goals |
894d6d863823
call Goal.prove only once for a quadratic number of theorems
traytel
parents:
59644
diff
changeset
|
77 |
|> Logic.mk_conjunction_balanced |
894d6d863823
call Goal.prove only once for a quadratic number of theorems
traytel
parents:
59644
diff
changeset
|
78 |
|> prove |
894d6d863823
call Goal.prove only once for a quadratic number of theorems
traytel
parents:
59644
diff
changeset
|
79 |
|> Conjunction.elim_balanced (length goals) |
894d6d863823
call Goal.prove only once for a quadratic number of theorems
traytel
parents:
59644
diff
changeset
|
80 |
|> map Simpdata.mk_eq; |
54615
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
81 |
in |
60046
894d6d863823
call Goal.prove only once for a quadratic number of theorems
traytel
parents:
59644
diff
changeset
|
82 |
(proves (triv_inject_goals @ inject_goals @ distinct_goals), Simpdata.mk_eq (prove refl_goal)) |
54615
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
83 |
end; |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
84 |
|
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
85 |
fun add_equality fcT fcT_name As ctrs inject_thms distinct_thms = |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
86 |
let |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
87 |
fun add_def lthy = |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
88 |
let |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
89 |
fun mk_side const_name = |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
90 |
Const (const_name, fcT --> fcT --> HOLogic.boolT) $ Free ("x", fcT) $ Free ("y", fcT); |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
91 |
val spec = |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
92 |
mk_Trueprop_eq (mk_side @{const_name HOL.equal}, mk_side @{const_name HOL.eq}) |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
93 |
|> Syntax.check_term lthy; |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
94 |
val ((_, (_, raw_def)), lthy') = |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
95 |
Specification.definition (NONE, (Attrib.empty_binding, spec)) lthy; |
59644 | 96 |
val thy_ctxt = Proof_Context.init_global (Proof_Context.theory_of lthy); (* FIXME? *) |
97 |
val def = singleton (Proof_Context.export lthy' thy_ctxt) raw_def; |
|
54615
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
98 |
in |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
99 |
(def, lthy') |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
100 |
end; |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
101 |
|
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
54895
diff
changeset
|
102 |
fun tac ctxt thms = |
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
54895
diff
changeset
|
103 |
Class.intro_classes_tac ctxt [] THEN ALLGOALS (Proof_Context.fact_tac ctxt thms); |
54615
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
104 |
|
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
105 |
val qualify = |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
106 |
Binding.qualify true (Long_Name.base_name fcT_name) o Binding.qualify true eqN o Binding.name; |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
107 |
in |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
108 |
Class.instantiation ([fcT_name], map dest_TFree As, [HOLogic.class_equal]) |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
109 |
#> add_def |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
54701
diff
changeset
|
110 |
#-> Class.prove_instantiation_exit_result (map o Morphism.thm) tac o single |
54615
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
111 |
#-> fold Code.del_eqn |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
112 |
#> `(mk_free_ctr_equations fcT ctrs inject_thms distinct_thms) |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
113 |
#-> (fn (thms, thm) => Global_Theory.note_thmss Thm.lemmaK |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
114 |
[((qualify reflN, [Code.add_nbe_default_eqn_attribute]), [([thm], [])]), |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
115 |
((qualify simpsN, [Code.add_default_eqn_attribute]), [(rev thms, [])])]) |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
116 |
#> snd |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
117 |
end; |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
118 |
|
54691
506312c293f5
code equations for "local" (co)datatypes available after interpretation of locales with assumptions
traytel
parents:
54615
diff
changeset
|
119 |
fun add_ctr_code fcT_name raw_As raw_ctrs inject_thms distinct_thms case_thms thy = |
54615
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
120 |
let |
54691
506312c293f5
code equations for "local" (co)datatypes available after interpretation of locales with assumptions
traytel
parents:
54615
diff
changeset
|
121 |
val As = map (perhaps (try Logic.unvarifyT_global)) raw_As; |
506312c293f5
code equations for "local" (co)datatypes available after interpretation of locales with assumptions
traytel
parents:
54615
diff
changeset
|
122 |
val ctrs = map (apsnd (perhaps (try Logic.unvarifyT_global))) raw_ctrs; |
54615
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
123 |
val fcT = Type (fcT_name, As); |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
124 |
val unover_ctrs = map (fn ctr as (_, fcT) => (Axclass.unoverload_const thy ctr, fcT)) ctrs; |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
125 |
in |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
126 |
if can (Code.constrset_of_consts thy) unover_ctrs then |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
127 |
thy |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
128 |
|> Code.add_datatype ctrs |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
129 |
|> fold_rev Code.add_default_eqn case_thms |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
130 |
|> Code.add_case (mk_case_certificate thy case_thms) |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
131 |
|> not (Sorts.has_instance (Sign.classes_of thy) fcT_name [HOLogic.class_equal]) |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
132 |
? add_equality fcT fcT_name As ctrs inject_thms distinct_thms |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
133 |
else |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
134 |
thy |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
135 |
end; |
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
136 |
|
62fb5af93fe2
generalized datatype code generation code so that it works with old-style and new-style (co)datatypes (as long as they are not local)
blanchet
parents:
diff
changeset
|
137 |
end; |