src/HOL/Tools/Ctr_Sugar/ctr_sugar_code.ML
author wenzelm
Fri, 09 Jun 2017 21:57:30 +0200
changeset 66055 07175635f78c
parent 63352 4eaf35781b23
child 66251 cd935b7cb3fb
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
54701
4ed7454aebde tuning -- moved ML files to subdirectory
blanchet
parents: 54691
diff changeset
     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
63073
413184c7a2a2 clarified context, notably for internal use of Simplifier;
wenzelm
parents: 63064
diff changeset
    26
fun mk_case_certificate ctxt raw_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
    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
63073
413184c7a2a2 clarified context, notably for internal use of Simplifier;
wenzelm
parents: 63064
diff changeset
    30
      |> Thm.unvarify_global (Proof_Context.theory_of ctxt)
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
    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);
63073
413184c7a2a2 clarified context, notably for internal use of Simplifier;
wenzelm
parents: 63064
diff changeset
    39
    val assum = Thm.cterm_of ctxt (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
63073
413184c7a2a2 clarified context, notably for internal use of Simplifier;
wenzelm
parents: 63064
diff changeset
    43
    |> rewrite_rule ctxt [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
63073
413184c7a2a2 clarified context, notably for internal use of Simplifier;
wenzelm
parents: 63064
diff changeset
    46
    |> Axclass.unoverload ctxt
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
    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
60825
bacfb7c45d81 more explicit context;
wenzelm
parents: 60046
diff changeset
    56
    val monomorphic_prop_of = Thm.prop_of o Thm.unvarify_global thy 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
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
    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
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
    72
        ALLGOALS (simp_tac
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
    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
61116
6189d179c2b5 close derivation *before* splitting conjuncts, like Goal.prove_common (see also 757cad5a3fe9) -- potential improvement of performance;
wenzelm
parents: 61065
diff changeset
    79
      |> Thm.close_derivation
60046
894d6d863823 call Goal.prove only once for a quadratic number of theorems
traytel
parents: 59644
diff changeset
    80
      |> Conjunction.elim_balanced (length goals)
894d6d863823 call Goal.prove only once for a quadratic number of theorems
traytel
parents: 59644
diff changeset
    81
      |> 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
    82
  in
60046
894d6d863823 call Goal.prove only once for a quadratic number of theorems
traytel
parents: 59644
diff changeset
    83
    (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
    84
  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
    85
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
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
    87
  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
    88
    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
    89
      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
    90
        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
    91
          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
    92
        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
    93
          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
    94
          |> 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
    95
        val ((_, (_, raw_def)), lthy') =
63352
4eaf35781b23 tuned signature;
wenzelm
parents: 63239
diff changeset
    96
          Specification.definition NONE [] [] (Binding.empty_atts, spec) lthy;
61065
ca4ebc63d8ac clarified context;
wenzelm
parents: 60825
diff changeset
    97
        val thy_ctxt = Proof_Context.init_global (Proof_Context.theory_of lthy');
59644
wenzelm
parents: 59621
diff changeset
    98
        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
    99
      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
   100
        (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
   101
      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
   102
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 54895
diff changeset
   103
    fun tac ctxt thms =
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 54895
diff changeset
   104
      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
   105
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
    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
   107
      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
   108
  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
   109
    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
   110
    #> add_def
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 54701
diff changeset
   111
    #-> 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
   112
    #-> 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
   113
    #> `(mk_free_ctr_equations fcT ctrs inject_thms distinct_thms)
61336
fa4ebbd350ae just one theorem kind, which is legacy anyway;
wenzelm
parents: 61116
diff changeset
   114
    #-> (fn (thms, thm) => Global_Theory.note_thmss Thm.theoremK
63239
d562c9948dee explicit tagging of code equations de-baroquifies interface
haftmann
parents: 63180
diff changeset
   115
      [((qualify reflN, [Code.add_default_eqn_attribute Code.Nbe]), [([thm], [])]),
d562c9948dee explicit tagging of code equations de-baroquifies interface
haftmann
parents: 63180
diff changeset
   116
        ((qualify simpsN, [Code.add_default_eqn_attribute Code.Equation]), [(rev 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
   117
    #> 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
   118
  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
   119
54691
506312c293f5 code equations for "local" (co)datatypes available after interpretation of locales with assumptions
traytel
parents: 54615
diff changeset
   120
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
   121
  let
54691
506312c293f5 code equations for "local" (co)datatypes available after interpretation of locales with assumptions
traytel
parents: 54615
diff changeset
   122
    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
   123
    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
   124
    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
   125
    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
   126
  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
   127
    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
   128
      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
   129
      |> Code.add_datatype ctrs
63239
d562c9948dee explicit tagging of code equations de-baroquifies interface
haftmann
parents: 63180
diff changeset
   130
      |> fold_rev (Code.add_eqn (Code.Equation, true)) case_thms
63073
413184c7a2a2 clarified context, notably for internal use of Simplifier;
wenzelm
parents: 63064
diff changeset
   131
      |> Code.add_case (mk_case_certificate (Proof_Context.init_global thy) case_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
   132
      |> 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
   133
        ? 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
   134
    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
   135
      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
   136
  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
   137
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
   138
end;