src/HOL/BNF/Tools/bnf_fp_rec_sugar.ML
author blanchet
Sun, 20 Oct 2013 22:39:40 +0200
changeset 54172 9c276e656712
parent 54170 402fcacb5c88
child 54173 8a2a5fa8c591
permissions -rw-r--r--
reintroduced parts of 07a8145aaeba that a22ded8a7f7d wrongly took out
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
     1
(*  Title:      HOL/BNF/Tools/bnf_fp_rec_sugar.ML
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
     2
    Author:     Lorenz Panny, TU Muenchen
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
     3
    Copyright   2013
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
     4
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
     5
Recursor and corecursor sugar.
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
     6
*)
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
     7
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
     8
signature BNF_FP_REC_SUGAR =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
     9
sig
54013
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
    10
  val add_primrec: (binding * typ option * mixfix) list ->
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
    11
    (Attrib.binding * term) list -> local_theory -> (term list * thm list list) * local_theory
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    12
  val add_primrec_cmd: (binding * string option * mixfix) list ->
54013
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
    13
    (Attrib.binding * string) list -> local_theory -> (term list * thm list list) * local_theory
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
    14
  val add_primrec_global: (binding * typ option * mixfix) list ->
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
    15
    (Attrib.binding * term) list -> theory -> (term list * thm list list) * theory
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
    16
  val add_primrec_overloaded: (string * (string * typ) * bool) list ->
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
    17
    (binding * typ option * mixfix) list ->
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
    18
    (Attrib.binding * term) list -> theory -> (term list * thm list list) * theory
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
    19
  val add_primrec_simple: ((binding * typ) * mixfix) list -> term list ->
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
    20
    local_theory -> (string list * (term list * (int list list * thm list list))) * local_theory
53753
ae7f50e70c09 renamed "primcorec" to "primcorecursive", to open the door to a 'theory -> theory' command called "primcorec" (cf. "fun" vs. "function")
blanchet
parents: 53744
diff changeset
    21
  val add_primcorecursive_cmd: bool ->
53831
80423b9080cf support "of" syntax to disambiguate selector equations
panny
parents: 53830
diff changeset
    22
    (binding * string option * mixfix) list * ((Attrib.binding * string) * string option) list ->
80423b9080cf support "of" syntax to disambiguate selector equations
panny
parents: 53830
diff changeset
    23
    Proof.context -> Proof.state
53822
6304b12c7627 add "primcorec" command (cf. ae7f50e70c09)
panny
parents: 53811
diff changeset
    24
  val add_primcorec_cmd: bool ->
53831
80423b9080cf support "of" syntax to disambiguate selector equations
panny
parents: 53830
diff changeset
    25
    (binding * string option * mixfix) list * ((Attrib.binding * string) * string option) list ->
80423b9080cf support "of" syntax to disambiguate selector equations
panny
parents: 53830
diff changeset
    26
    local_theory -> local_theory
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    27
end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    28
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    29
structure BNF_FP_Rec_Sugar : BNF_FP_REC_SUGAR =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    30
struct
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    31
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    32
open BNF_Util
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    33
open BNF_FP_Util
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    34
open BNF_FP_Rec_Sugar_Util
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    35
open BNF_FP_Rec_Sugar_Tactics
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    36
54065
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
    37
val codeN = "code";
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
    38
val ctrN = "ctr";
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
    39
val discN = "disc";
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
    40
val selN = "sel";
53791
0ddf045113c9 tuned code
blanchet
parents: 53753
diff changeset
    41
54145
297d1c603999 make sure that registered code equations are actually equations
blanchet
parents: 54133
diff changeset
    42
val nitpicksimp_attrs = @{attributes [nitpick_simp]};
53794
af7d1533a25b undid copy-paste
blanchet
parents: 53793
diff changeset
    43
val simp_attrs = @{attributes [simp]};
54145
297d1c603999 make sure that registered code equations are actually equations
blanchet
parents: 54133
diff changeset
    44
val code_nitpicksimp_attrs = Code.add_default_eqn_attrib :: nitpicksimp_attrs;
297d1c603999 make sure that registered code equations are actually equations
blanchet
parents: 54133
diff changeset
    45
val code_nitpicksimp_simp_attrs = Code.add_default_eqn_attrib :: nitpicksimp_attrs @ simp_attrs;
53794
af7d1533a25b undid copy-paste
blanchet
parents: 53793
diff changeset
    46
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    47
exception Primrec_Error of string * term list;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    48
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    49
fun primrec_error str = raise Primrec_Error (str, []);
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    50
fun primrec_error_eqn str eqn = raise Primrec_Error (str, [eqn]);
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    51
fun primrec_error_eqns str eqns = raise Primrec_Error (str, eqns);
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    52
53358
b46e6cd75dc6 improved interfaces
panny
parents: 53357
diff changeset
    53
fun finds eq = fold_map (fn x => List.partition (curry eq x) #>> pair x);
b46e6cd75dc6 improved interfaces
panny
parents: 53357
diff changeset
    54
53357
46b0c7a08af7 simplified rewriting of map arguments
panny
parents: 53354
diff changeset
    55
val free_name = try (fn Free (v, _) => v);
46b0c7a08af7 simplified rewriting of map arguments
panny
parents: 53354
diff changeset
    56
val const_name = try (fn Const (v, _) => v);
53358
b46e6cd75dc6 improved interfaces
panny
parents: 53357
diff changeset
    57
val undef_const = Const (@{const_name undefined}, dummyT);
53357
46b0c7a08af7 simplified rewriting of map arguments
panny
parents: 53354
diff changeset
    58
53358
b46e6cd75dc6 improved interfaces
panny
parents: 53357
diff changeset
    59
fun permute_args n t = list_comb (t, map Bound (0 :: (n downto 1)))
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
    60
  |> fold (K (Term.abs (Name.uu, dummyT))) (0 upto n);
53401
2101a97e6220 various refactoring;
panny
parents: 53360
diff changeset
    61
val abs_tuple = HOLogic.tupled_lambda o HOLogic.mk_tuple;
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
    62
fun drop_All t = subst_bounds (strip_qnt_vars @{const_name all} t |> map Free |> rev,
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
    63
  strip_qnt_body @{const_name all} t)
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
    64
fun abstract vs =
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
    65
  let fun a n (t $ u) = a n t $ a n u
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
    66
        | a n (Abs (v, T, b)) = Abs (v, T, a (n + 1) b)
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
    67
        | a n t = let val idx = find_index (equal t) vs in
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
    68
            if idx < 0 then t else Bound (n + idx) end
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
    69
  in a 0 end;
53735
99331dac1e1c simplified code; eliminated some dummyTs
panny
parents: 53734
diff changeset
    70
fun mk_prod1 Ts (t, u) = HOLogic.pair_const (fastype_of1 (Ts, t)) (fastype_of1 (Ts, u)) $ t $ u;
99331dac1e1c simplified code; eliminated some dummyTs
panny
parents: 53734
diff changeset
    71
fun mk_tuple1 Ts = the_default HOLogic.unit o try (foldr1 (mk_prod1 Ts));
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    72
53794
af7d1533a25b undid copy-paste
blanchet
parents: 53793
diff changeset
    73
fun get_indices fixes t = map (fst #>> Binding.name_of #> Free) fixes
af7d1533a25b undid copy-paste
blanchet
parents: 53793
diff changeset
    74
  |> map_index (fn (i, v) => if exists_subterm (equal v) t then SOME i else NONE)
af7d1533a25b undid copy-paste
blanchet
parents: 53793
diff changeset
    75
  |> map_filter I;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    76
53310
8af01463b2d3 moved keywords down the hierarchy
blanchet
parents: 53304
diff changeset
    77
8af01463b2d3 moved keywords down the hierarchy
blanchet
parents: 53304
diff changeset
    78
(* Primrec *)
8af01463b2d3 moved keywords down the hierarchy
blanchet
parents: 53304
diff changeset
    79
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    80
type eqn_data = {
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    81
  fun_name: string,
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    82
  rec_type: typ,
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    83
  ctr: term,
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    84
  ctr_args: term list,
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    85
  left_args: term list,
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    86
  right_args: term list,
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    87
  res_type: typ,
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    88
  rhs_term: term,
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    89
  user_eqn: term
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    90
};
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    91
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    92
fun dissect_eqn lthy fun_names eqn' =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    93
  let
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
    94
    val eqn = drop_All eqn' |> HOLogic.dest_Trueprop
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
    95
      handle TERM _ =>
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
    96
        primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn';
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    97
    val (lhs, rhs) = HOLogic.dest_eq eqn
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    98
        handle TERM _ =>
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    99
          primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn';
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   100
    val (fun_name, args) = strip_comb lhs
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   101
      |>> (fn x => if is_Free x then fst (dest_Free x)
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   102
          else primrec_error_eqn "malformed function equation (does not start with free)" eqn);
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   103
    val (left_args, rest) = take_prefix is_Free args;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   104
    val (nonfrees, right_args) = take_suffix is_Free rest;
53830
ed2eb7df2aac don't note more induction principles than there are functions + tuning
blanchet
parents: 53822
diff changeset
   105
    val num_nonfrees = length nonfrees;
ed2eb7df2aac don't note more induction principles than there are functions + tuning
blanchet
parents: 53822
diff changeset
   106
    val _ = num_nonfrees = 1 orelse if num_nonfrees = 0 then
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   107
      primrec_error_eqn "constructor pattern missing in left-hand side" eqn else
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   108
      primrec_error_eqn "more than one non-variable argument in left-hand side" eqn;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   109
    val _ = member (op =) fun_names fun_name orelse
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   110
      primrec_error_eqn "malformed function equation (does not start with function name)" eqn
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   111
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   112
    val (ctr, ctr_args) = strip_comb (the_single nonfrees);
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   113
    val _ = try (num_binder_types o fastype_of) ctr = SOME (length ctr_args) orelse
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   114
      primrec_error_eqn "partially applied constructor in pattern" eqn;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   115
    val _ = let val d = duplicates (op =) (left_args @ ctr_args @ right_args) in null d orelse
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   116
      primrec_error_eqn ("duplicate variable \"" ^ Syntax.string_of_term lthy (hd d) ^
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   117
        "\" in left-hand side") eqn end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   118
    val _ = forall is_Free ctr_args orelse
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   119
      primrec_error_eqn "non-primitive pattern in left-hand side" eqn;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   120
    val _ =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   121
      let val b = fold_aterms (fn x as Free (v, _) =>
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   122
        if (not (member (op =) (left_args @ ctr_args @ right_args) x) andalso
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   123
        not (member (op =) fun_names v) andalso
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   124
        not (Variable.is_fixed lthy v)) then cons x else I | _ => I) rhs []
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   125
      in
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   126
        null b orelse
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   127
        primrec_error_eqn ("extra variable(s) in right-hand side: " ^
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   128
          commas (map (Syntax.string_of_term lthy) b)) eqn
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   129
      end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   130
  in
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   131
    {fun_name = fun_name,
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   132
     rec_type = body_type (type_of ctr),
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   133
     ctr = ctr,
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   134
     ctr_args = ctr_args,
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   135
     left_args = left_args,
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   136
     right_args = right_args,
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   137
     res_type = map fastype_of (left_args @ right_args) ---> fastype_of rhs,
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   138
     rhs_term = rhs,
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   139
     user_eqn = eqn'}
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   140
  end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   141
53401
2101a97e6220 various refactoring;
panny
parents: 53360
diff changeset
   142
fun rewrite_map_arg get_ctr_pos rec_type res_type =
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   143
  let
53357
46b0c7a08af7 simplified rewriting of map arguments
panny
parents: 53354
diff changeset
   144
    val pT = HOLogic.mk_prodT (rec_type, res_type);
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   145
53357
46b0c7a08af7 simplified rewriting of map arguments
panny
parents: 53354
diff changeset
   146
    val maybe_suc = Option.map (fn x => x + 1);
46b0c7a08af7 simplified rewriting of map arguments
panny
parents: 53354
diff changeset
   147
    fun subst d (t as Bound d') = t |> d = SOME d' ? curry (op $) (fst_const pT)
46b0c7a08af7 simplified rewriting of map arguments
panny
parents: 53354
diff changeset
   148
      | subst d (Abs (v, T, b)) = Abs (v, if d = SOME ~1 then pT else T, subst (maybe_suc d) b)
46b0c7a08af7 simplified rewriting of map arguments
panny
parents: 53354
diff changeset
   149
      | subst d t =
53358
b46e6cd75dc6 improved interfaces
panny
parents: 53357
diff changeset
   150
        let
b46e6cd75dc6 improved interfaces
panny
parents: 53357
diff changeset
   151
          val (u, vs) = strip_comb t;
53401
2101a97e6220 various refactoring;
panny
parents: 53360
diff changeset
   152
          val ctr_pos = try (get_ctr_pos o the) (free_name u) |> the_default ~1;
53358
b46e6cd75dc6 improved interfaces
panny
parents: 53357
diff changeset
   153
        in
53401
2101a97e6220 various refactoring;
panny
parents: 53360
diff changeset
   154
          if ctr_pos >= 0 then
53357
46b0c7a08af7 simplified rewriting of map arguments
panny
parents: 53354
diff changeset
   155
            if d = SOME ~1 andalso length vs = ctr_pos then
46b0c7a08af7 simplified rewriting of map arguments
panny
parents: 53354
diff changeset
   156
              list_comb (permute_args ctr_pos (snd_const pT), vs)
46b0c7a08af7 simplified rewriting of map arguments
panny
parents: 53354
diff changeset
   157
            else if length vs > ctr_pos andalso is_some d
46b0c7a08af7 simplified rewriting of map arguments
panny
parents: 53354
diff changeset
   158
                andalso d = try (fn Bound n => n) (nth vs ctr_pos) then
46b0c7a08af7 simplified rewriting of map arguments
panny
parents: 53354
diff changeset
   159
              list_comb (snd_const pT $ nth vs ctr_pos, map (subst d) (nth_drop ctr_pos vs))
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   160
            else
53357
46b0c7a08af7 simplified rewriting of map arguments
panny
parents: 53354
diff changeset
   161
              primrec_error_eqn ("recursive call not directly applied to constructor argument") t
46b0c7a08af7 simplified rewriting of map arguments
panny
parents: 53354
diff changeset
   162
          else if d = SOME ~1 andalso const_name u = SOME @{const_name comp} then
46b0c7a08af7 simplified rewriting of map arguments
panny
parents: 53354
diff changeset
   163
            list_comb (map_types (K dummyT) u, map2 subst [NONE, d] vs)
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   164
          else
53357
46b0c7a08af7 simplified rewriting of map arguments
panny
parents: 53354
diff changeset
   165
            list_comb (u, map (subst (d |> d = SOME ~1 ? K NONE)) vs)
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   166
        end
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   167
  in
53357
46b0c7a08af7 simplified rewriting of map arguments
panny
parents: 53354
diff changeset
   168
    subst (SOME ~1)
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   169
  end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   170
54170
402fcacb5c88 gracefully handle case where a nested recursive variable appears directly on the rhs of a primcorec equation
blanchet
parents: 54169
diff changeset
   171
fun subst_rec_calls lthy get_ctr_pos has_call ctr_args mutual_calls nested_calls =
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   172
  let
54169
blanchet
parents: 54168
diff changeset
   173
    fun try_nested_rec bound_Ts y t =
blanchet
parents: 54168
diff changeset
   174
      AList.lookup (op =) nested_calls y
blanchet
parents: 54168
diff changeset
   175
      |> Option.map (fn y' =>
blanchet
parents: 54168
diff changeset
   176
        massage_nested_rec_call lthy has_call (rewrite_map_arg get_ctr_pos) bound_Ts y y' t);
blanchet
parents: 54168
diff changeset
   177
54166
blanchet
parents: 54165
diff changeset
   178
    fun subst bound_Ts (t as g' $ y) =
54170
402fcacb5c88 gracefully handle case where a nested recursive variable appears directly on the rhs of a primcorec equation
blanchet
parents: 54169
diff changeset
   179
        let
402fcacb5c88 gracefully handle case where a nested recursive variable appears directly on the rhs of a primcorec equation
blanchet
parents: 54169
diff changeset
   180
          fun subst_rec () = subst bound_Ts g' $ subst bound_Ts y;
402fcacb5c88 gracefully handle case where a nested recursive variable appears directly on the rhs of a primcorec equation
blanchet
parents: 54169
diff changeset
   181
          val y_head = head_of y;
402fcacb5c88 gracefully handle case where a nested recursive variable appears directly on the rhs of a primcorec equation
blanchet
parents: 54169
diff changeset
   182
        in
54159
eb5d58c99049 set stage for more flexible 'primrec' syntax for recursion through functions
blanchet
parents: 54157
diff changeset
   183
          if not (member (op =) ctr_args y_head) then
54170
402fcacb5c88 gracefully handle case where a nested recursive variable appears directly on the rhs of a primcorec equation
blanchet
parents: 54169
diff changeset
   184
            subst_rec ()
54159
eb5d58c99049 set stage for more flexible 'primrec' syntax for recursion through functions
blanchet
parents: 54157
diff changeset
   185
          else
54169
blanchet
parents: 54168
diff changeset
   186
            (case try_nested_rec bound_Ts y_head t of
blanchet
parents: 54168
diff changeset
   187
              SOME t' => t'
54168
blanchet
parents: 54167
diff changeset
   188
            | NONE =>
blanchet
parents: 54167
diff changeset
   189
              let val (g, g_args) = strip_comb g' in
blanchet
parents: 54167
diff changeset
   190
                (case try (get_ctr_pos o the) (free_name g) of
blanchet
parents: 54167
diff changeset
   191
                  SOME ctr_pos =>
blanchet
parents: 54167
diff changeset
   192
                  (length g_args >= ctr_pos orelse
blanchet
parents: 54167
diff changeset
   193
                   primrec_error_eqn "too few arguments in recursive call" t;
blanchet
parents: 54167
diff changeset
   194
                   (case AList.lookup (op =) mutual_calls y of
blanchet
parents: 54167
diff changeset
   195
                     SOME y' => list_comb (y', g_args)
54170
402fcacb5c88 gracefully handle case where a nested recursive variable appears directly on the rhs of a primcorec equation
blanchet
parents: 54169
diff changeset
   196
                   | NONE => subst_rec ()))
402fcacb5c88 gracefully handle case where a nested recursive variable appears directly on the rhs of a primcorec equation
blanchet
parents: 54169
diff changeset
   197
                | NONE => subst_rec ())
54168
blanchet
parents: 54167
diff changeset
   198
              end)
54159
eb5d58c99049 set stage for more flexible 'primrec' syntax for recursion through functions
blanchet
parents: 54157
diff changeset
   199
        end
54166
blanchet
parents: 54165
diff changeset
   200
      | subst bound_Ts (Abs (v, T, b)) = Abs (v, T, subst (T :: bound_Ts) b)
53350
17632ef6cfe8 simplified recursive calls' replacement
panny
parents: 53341
diff changeset
   201
      | subst _ t = t
54170
402fcacb5c88 gracefully handle case where a nested recursive variable appears directly on the rhs of a primcorec equation
blanchet
parents: 54169
diff changeset
   202
402fcacb5c88 gracefully handle case where a nested recursive variable appears directly on the rhs of a primcorec equation
blanchet
parents: 54169
diff changeset
   203
    fun subst' t =
402fcacb5c88 gracefully handle case where a nested recursive variable appears directly on the rhs of a primcorec equation
blanchet
parents: 54169
diff changeset
   204
      if has_call t then
402fcacb5c88 gracefully handle case where a nested recursive variable appears directly on the rhs of a primcorec equation
blanchet
parents: 54169
diff changeset
   205
        (* FIXME detect this case earlier? *)
402fcacb5c88 gracefully handle case where a nested recursive variable appears directly on the rhs of a primcorec equation
blanchet
parents: 54169
diff changeset
   206
        primrec_error_eqn "recursive call not directly applied to constructor argument" t
402fcacb5c88 gracefully handle case where a nested recursive variable appears directly on the rhs of a primcorec equation
blanchet
parents: 54169
diff changeset
   207
      else
402fcacb5c88 gracefully handle case where a nested recursive variable appears directly on the rhs of a primcorec equation
blanchet
parents: 54169
diff changeset
   208
        try_nested_rec [] (head_of t) t |> the_default t
53350
17632ef6cfe8 simplified recursive calls' replacement
panny
parents: 53341
diff changeset
   209
  in
54170
402fcacb5c88 gracefully handle case where a nested recursive variable appears directly on the rhs of a primcorec equation
blanchet
parents: 54169
diff changeset
   210
    subst' o subst []
53350
17632ef6cfe8 simplified recursive calls' replacement
panny
parents: 53341
diff changeset
   211
  end;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   212
54003
c4343c31f86d made SML/NJ happy
blanchet
parents: 54002
diff changeset
   213
fun build_rec_arg lthy (funs_data : eqn_data list list) has_call (ctr_spec : rec_ctr_spec)
c4343c31f86d made SML/NJ happy
blanchet
parents: 54002
diff changeset
   214
    (maybe_eqn_data : eqn_data option) =
53358
b46e6cd75dc6 improved interfaces
panny
parents: 53357
diff changeset
   215
  if is_none maybe_eqn_data then undef_const else
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   216
    let
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   217
      val eqn_data = the maybe_eqn_data;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   218
      val t = #rhs_term eqn_data;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   219
      val ctr_args = #ctr_args eqn_data;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   220
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   221
      val calls = #calls ctr_spec;
54102
82ada0a92dde tuned names
blanchet
parents: 54101
diff changeset
   222
      val n_args = fold (curry (op +) o (fn Mutual_Rec _ => 2 | _ => 1)) calls 0;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   223
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   224
      val no_calls' = tag_list 0 calls
54102
82ada0a92dde tuned names
blanchet
parents: 54101
diff changeset
   225
        |> map_filter (try (apsnd (fn No_Rec n => n | Mutual_Rec (n, _) => n)));
82ada0a92dde tuned names
blanchet
parents: 54101
diff changeset
   226
      val mutual_calls' = tag_list 0 calls
82ada0a92dde tuned names
blanchet
parents: 54101
diff changeset
   227
        |> map_filter (try (apsnd (fn Mutual_Rec (_, n) => n)));
82ada0a92dde tuned names
blanchet
parents: 54101
diff changeset
   228
      val nested_calls' = tag_list 0 calls
82ada0a92dde tuned names
blanchet
parents: 54101
diff changeset
   229
        |> map_filter (try (apsnd (fn Nested_Rec n => n)));
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   230
54102
82ada0a92dde tuned names
blanchet
parents: 54101
diff changeset
   231
      fun make_mutual_type _ = dummyT; (* FIXME? *)
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   232
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   233
      val rec_res_type_list = map (fn (x :: _) => (#rec_type x, #res_type x)) funs_data;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   234
54102
82ada0a92dde tuned names
blanchet
parents: 54101
diff changeset
   235
      fun make_nested_type (Type (Tname, Ts)) = Type (Tname, Ts |> map (fn T =>
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   236
        let val maybe_res_type = AList.lookup (op =) rec_res_type_list T in
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   237
          if is_some maybe_res_type
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   238
          then HOLogic.mk_prodT (T, the maybe_res_type)
54102
82ada0a92dde tuned names
blanchet
parents: 54101
diff changeset
   239
          else make_nested_type T end))
82ada0a92dde tuned names
blanchet
parents: 54101
diff changeset
   240
        | make_nested_type T = T;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   241
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   242
      val args = replicate n_args ("", dummyT)
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   243
        |> Term.rename_wrt_term t
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   244
        |> map Free
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   245
        |> fold (fn (ctr_arg_idx, arg_idx) =>
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   246
            nth_map arg_idx (K (nth ctr_args ctr_arg_idx)))
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   247
          no_calls'
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   248
        |> fold (fn (ctr_arg_idx, arg_idx) =>
54102
82ada0a92dde tuned names
blanchet
parents: 54101
diff changeset
   249
            nth_map arg_idx (K (nth ctr_args ctr_arg_idx |> map_types make_mutual_type)))
82ada0a92dde tuned names
blanchet
parents: 54101
diff changeset
   250
          mutual_calls'
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   251
        |> fold (fn (ctr_arg_idx, arg_idx) =>
54102
82ada0a92dde tuned names
blanchet
parents: 54101
diff changeset
   252
            nth_map arg_idx (K (nth ctr_args ctr_arg_idx |> map_types make_nested_type)))
82ada0a92dde tuned names
blanchet
parents: 54101
diff changeset
   253
          nested_calls';
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   254
53401
2101a97e6220 various refactoring;
panny
parents: 53360
diff changeset
   255
      val fun_name_ctr_pos_list =
2101a97e6220 various refactoring;
panny
parents: 53360
diff changeset
   256
        map (fn (x :: _) => (#fun_name x, length (#left_args x))) funs_data;
2101a97e6220 various refactoring;
panny
parents: 53360
diff changeset
   257
      val get_ctr_pos = try (the o AList.lookup (op =) fun_name_ctr_pos_list) #> the_default ~1;
54102
82ada0a92dde tuned names
blanchet
parents: 54101
diff changeset
   258
      val mutual_calls = map (apfst (nth ctr_args) o apsnd (nth args)) mutual_calls';
82ada0a92dde tuned names
blanchet
parents: 54101
diff changeset
   259
      val nested_calls = map (apfst (nth ctr_args) o apsnd (nth args)) nested_calls';
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   260
53401
2101a97e6220 various refactoring;
panny
parents: 53360
diff changeset
   261
      val abstractions = args @ #left_args eqn_data @ #right_args eqn_data;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   262
    in
53350
17632ef6cfe8 simplified recursive calls' replacement
panny
parents: 53341
diff changeset
   263
      t
54102
82ada0a92dde tuned names
blanchet
parents: 54101
diff changeset
   264
      |> subst_rec_calls lthy get_ctr_pos has_call ctr_args mutual_calls nested_calls
53401
2101a97e6220 various refactoring;
panny
parents: 53360
diff changeset
   265
      |> fold_rev lambda abstractions
53350
17632ef6cfe8 simplified recursive calls' replacement
panny
parents: 53341
diff changeset
   266
    end;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   267
54003
c4343c31f86d made SML/NJ happy
blanchet
parents: 54002
diff changeset
   268
fun build_defs lthy bs mxs (funs_data : eqn_data list list) (rec_specs : rec_spec list) has_call =
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   269
  let
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   270
    val n_funs = length funs_data;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   271
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   272
    val ctr_spec_eqn_data_list' =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   273
      (take n_funs rec_specs |> map #ctr_specs) ~~ funs_data
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   274
      |> maps (uncurry (finds (fn (x, y) => #ctr x = #ctr y))
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   275
          ##> (fn x => null x orelse
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   276
            primrec_error_eqns "excess equations in definition" (map #rhs_term x)) #> fst);
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   277
    val _ = ctr_spec_eqn_data_list' |> map (fn (_, x) => length x <= 1 orelse
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   278
      primrec_error_eqns ("multiple equations for constructor") (map #user_eqn x));
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   279
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   280
    val ctr_spec_eqn_data_list =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   281
      ctr_spec_eqn_data_list' @ (drop n_funs rec_specs |> maps #ctr_specs |> map (rpair []));
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   282
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   283
    val recs = take n_funs rec_specs |> map #recx;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   284
    val rec_args = ctr_spec_eqn_data_list
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   285
      |> sort ((op <) o pairself (#offset o fst) |> make_ord)
53358
b46e6cd75dc6 improved interfaces
panny
parents: 53357
diff changeset
   286
      |> map (uncurry (build_rec_arg lthy funs_data has_call) o apsnd (try the_single));
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   287
    val ctr_poss = map (fn x =>
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   288
      if length (distinct ((op =) o pairself (length o #left_args)) x) <> 1 then
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   289
        primrec_error ("inconstant constructor pattern position for function " ^
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   290
          quote (#fun_name (hd x)))
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   291
      else
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   292
        hd x |> #left_args |> length) funs_data;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   293
  in
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   294
    (recs, ctr_poss)
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   295
    |-> map2 (fn recx => fn ctr_pos => list_comb (recx, rec_args) |> permute_args ctr_pos)
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   296
    |> Syntax.check_terms lthy
54155
b964fad0cbbd conceal more ugly constructions
blanchet
parents: 54154
diff changeset
   297
    |> map3 (fn b => fn mx => fn t => ((b, mx), ((Binding.conceal (Thm.def_binding b), []), t)))
b964fad0cbbd conceal more ugly constructions
blanchet
parents: 54154
diff changeset
   298
      bs mxs
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   299
  end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   300
54001
65fc58793ed5 made SML/NJ happier
blanchet
parents: 53925
diff changeset
   301
fun find_rec_calls has_call (eqn_data : eqn_data) =
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   302
  let
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   303
    fun find (Abs (_, _, b)) ctr_arg = find b ctr_arg
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   304
      | find (t as _ $ _) ctr_arg =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   305
        let
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   306
          val (f', args') = strip_comb t;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   307
          val n = find_index (equal ctr_arg) args';
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   308
        in
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   309
          if n < 0 then
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   310
            find f' ctr_arg @ maps (fn x => find x ctr_arg) args'
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   311
          else
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   312
            let val (f, args) = chop n args' |>> curry list_comb f' in
53358
b46e6cd75dc6 improved interfaces
panny
parents: 53357
diff changeset
   313
              if has_call f then
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   314
                f :: maps (fn x => find x ctr_arg) args
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   315
              else
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   316
                find f ctr_arg @ maps (fn x => find x ctr_arg) args
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   317
            end
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   318
        end
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   319
      | find _ _ = [];
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   320
  in
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   321
    map (find (#rhs_term eqn_data)) (#ctr_args eqn_data)
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   322
    |> (fn [] => NONE | callss => SOME (#ctr eqn_data, callss))
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   323
  end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   324
54013
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   325
fun prepare_primrec fixes specs lthy =
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   326
  let
53352
43a1cc050943 honor mixfix specifications
traytel
parents: 53341
diff changeset
   327
    val (bs, mxs) = map_split (apfst fst) fixes;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   328
    val fun_names = map Binding.name_of bs;
54013
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   329
    val eqns_data = map (dissect_eqn lthy fun_names) specs;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   330
    val funs_data = eqns_data
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   331
      |> partition_eq ((op =) o pairself #fun_name)
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   332
      |> finds (fn (x, y) => x = #fun_name (hd y)) fun_names |> fst
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   333
      |> map (fn (x, y) => the_single y handle List.Empty =>
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   334
          primrec_error ("missing equations for function " ^ quote x));
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   335
53358
b46e6cd75dc6 improved interfaces
panny
parents: 53357
diff changeset
   336
    val has_call = exists_subterm (map (fst #>> Binding.name_of #> Free) fixes |> member (op =));
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   337
    val arg_Ts = map (#rec_type o hd) funs_data;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   338
    val res_Ts = map (#res_type o hd) funs_data;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   339
    val callssss = funs_data
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   340
      |> map (partition_eq ((op =) o pairself #ctr))
53358
b46e6cd75dc6 improved interfaces
panny
parents: 53357
diff changeset
   341
      |> map (maps (map_filter (find_rec_calls has_call)));
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   342
53830
ed2eb7df2aac don't note more induction principles than there are functions + tuning
blanchet
parents: 53822
diff changeset
   343
    val ((n2m, rec_specs, _, induct_thm, induct_thms), lthy') =
53794
af7d1533a25b undid copy-paste
blanchet
parents: 53793
diff changeset
   344
      rec_specs_of bs arg_Ts res_Ts (get_indices fixes) callssss lthy;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   345
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   346
    val actual_nn = length funs_data;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   347
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   348
    val _ = let val ctrs = (maps (map #ctr o #ctr_specs) rec_specs) in
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   349
      map (fn {ctr, user_eqn, ...} => member (op =) ctrs ctr orelse
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   350
        primrec_error_eqn ("argument " ^ quote (Syntax.string_of_term lthy' ctr) ^
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   351
          " is not a constructor in left-hand side") user_eqn) eqns_data end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   352
53358
b46e6cd75dc6 improved interfaces
panny
parents: 53357
diff changeset
   353
    val defs = build_defs lthy' bs mxs funs_data rec_specs has_call;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   354
54013
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   355
    fun prove lthy def_thms' ({ctr_specs, nested_map_idents, nested_map_comps, ...} : rec_spec)
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   356
        (fun_data : eqn_data list) =
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   357
      let
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   358
        val def_thms = map (snd o snd) def_thms';
54013
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   359
        val simp_thmss = finds (fn (x, y) => #ctr x = #ctr y) fun_data ctr_specs
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   360
          |> fst
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   361
          |> map_filter (try (fn (x, [y]) =>
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   362
            (#user_eqn x, length (#left_args x) + length (#right_args x), #rec_thm y)))
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   363
          |> map (fn (user_eqn, num_extra_args, rec_thm) =>
53329
c31c0c311cf0 more canonical naming
blanchet
parents: 53310
diff changeset
   364
            mk_primrec_tac lthy num_extra_args nested_map_idents nested_map_comps def_thms rec_thm
54013
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   365
            |> K |> Goal.prove lthy [] [] user_eqn);
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   366
        val poss = find_indices (fn (x, y) => #ctr x = #ctr y) fun_data eqns_data;
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   367
      in
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   368
        (poss, simp_thmss)
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   369
      end;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   370
54013
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   371
    val notes =
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   372
      (if n2m then map2 (fn name => fn thm =>
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   373
        (name, inductN, [thm], [])) fun_names (take actual_nn induct_thms) else [])
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   374
      |> map (fn (prefix, thmN, thms, attrs) =>
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   375
        ((Binding.qualify true prefix (Binding.name thmN), attrs), [(thms, [])]));
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   376
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   377
    val common_name = mk_common_name fun_names;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   378
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   379
    val common_notes =
54013
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   380
      (if n2m then [(inductN, [induct_thm], [])] else [])
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   381
      |> map (fn (thmN, thms, attrs) =>
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   382
        ((Binding.qualify true common_name (Binding.name thmN), attrs), [(thms, [])]));
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   383
  in
54013
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   384
    (((fun_names, defs),
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   385
      fn lthy => fn defs =>
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   386
        split_list (map2 (prove lthy defs) (take actual_nn rec_specs) funs_data)),
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   387
      lthy' |> Local_Theory.notes (notes @ common_notes) |> snd)
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   388
  end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   389
54013
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   390
(* primrec definition *)
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   391
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   392
fun add_primrec_simple fixes ts lthy =
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   393
  let
54013
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   394
    val (((names, defs), prove), lthy) = prepare_primrec fixes ts lthy
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   395
      handle ERROR str => primrec_error str;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   396
  in
54013
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   397
    lthy
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   398
    |> fold_map Local_Theory.define defs
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   399
    |-> (fn defs => `(fn lthy => (names, (map fst defs, prove lthy defs))))
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   400
  end
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   401
  handle Primrec_Error (str, eqns) =>
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   402
    if null eqns
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   403
    then error ("primrec_new error:\n  " ^ str)
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   404
    else error ("primrec_new error:\n  " ^ str ^ "\nin\n  " ^
53822
6304b12c7627 add "primcorec" command (cf. ae7f50e70c09)
panny
parents: 53811
diff changeset
   405
      space_implode "\n  " (map (quote o Syntax.string_of_term lthy) eqns));
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   406
54013
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   407
local
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   408
54028
4d087a8950f3 made SML/NJ happy
traytel
parents: 54013
diff changeset
   409
fun gen_primrec prep_spec (raw_fixes : (binding * 'a option * mixfix) list) raw_spec lthy =
54013
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   410
  let
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   411
    val d = duplicates (op =) (map (Binding.name_of o #1) raw_fixes)
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   412
    val _ = null d orelse primrec_error ("duplicate function name(s): " ^ commas d);
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   413
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   414
    val (fixes, specs) = fst (prep_spec raw_fixes raw_spec lthy);
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   415
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   416
    val mk_notes =
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   417
      flat ooo map3 (fn poss => fn prefix => fn thms =>
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   418
        let
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   419
          val (bs, attrss) = map_split (fst o nth specs) poss;
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   420
          val notes =
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   421
            map3 (fn b => fn attrs => fn thm =>
54145
297d1c603999 make sure that registered code equations are actually equations
blanchet
parents: 54133
diff changeset
   422
              ((Binding.qualify false prefix b, code_nitpicksimp_simp_attrs @ attrs), [([thm], [])]))
54013
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   423
            bs attrss thms;
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   424
        in
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   425
          ((Binding.qualify true prefix (Binding.name simpsN), []), [(thms, [])]) :: notes
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   426
        end);
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   427
  in
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   428
    lthy
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   429
    |> add_primrec_simple fixes (map snd specs)
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   430
    |-> (fn (names, (ts, (posss, simpss))) =>
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   431
      Spec_Rules.add Spec_Rules.Equational (ts, flat simpss)
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   432
      #> Local_Theory.notes (mk_notes posss names simpss)
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   433
      #>> pair ts o map snd)
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   434
  end;
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   435
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   436
in
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   437
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   438
val add_primrec = gen_primrec Specification.check_spec;
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   439
val add_primrec_cmd = gen_primrec Specification.read_spec;
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   440
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   441
end;
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   442
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   443
fun add_primrec_global fixes specs thy =
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   444
  let
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   445
    val lthy = Named_Target.theory_init thy;
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   446
    val ((ts, simps), lthy') = add_primrec fixes specs lthy;
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   447
    val simps' = burrow (Proof_Context.export lthy' lthy) simps;
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   448
  in ((ts, simps'), Local_Theory.exit_global lthy') end;
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   449
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   450
fun add_primrec_overloaded ops fixes specs thy =
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   451
  let
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   452
    val lthy = Overloading.overloading ops thy;
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   453
    val ((ts, simps), lthy') = add_primrec fixes specs lthy;
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   454
    val simps' = burrow (Proof_Context.export lthy' lthy) simps;
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   455
  in ((ts, simps'), Local_Theory.exit_global lthy') end;
38c0bbb8348b improved backwards compatiblity of primrec_new (Isabelle/ML interface, attributes, etc.)
traytel
parents: 54003
diff changeset
   456
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   457
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   458
53310
8af01463b2d3 moved keywords down the hierarchy
blanchet
parents: 53304
diff changeset
   459
(* Primcorec *)
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   460
54153
67487a607ce2 avoid 'co_' prefix with underscore meaning 'co', since it is our only possible identifier representation of '(co)'
blanchet
parents: 54145
diff changeset
   461
type coeqn_data_disc = {
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   462
  fun_name: string,
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   463
  fun_T: typ,
53401
2101a97e6220 various refactoring;
panny
parents: 53360
diff changeset
   464
  fun_args: term list,
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   465
  ctr: term,
53341
63015d035301 handle selector formulae with no corecursive calls
panny
parents: 53335
diff changeset
   466
  ctr_no: int, (*###*)
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   467
  disc: term,
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   468
  prems: term list,
53822
6304b12c7627 add "primcorec" command (cf. ae7f50e70c09)
panny
parents: 53811
diff changeset
   469
  auto_gen: bool,
54097
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   470
  maybe_ctr_rhs: term option,
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   471
  maybe_code_rhs: term option,
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   472
  user_eqn: term
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   473
};
54001
65fc58793ed5 made SML/NJ happier
blanchet
parents: 53925
diff changeset
   474
54153
67487a607ce2 avoid 'co_' prefix with underscore meaning 'co', since it is our only possible identifier representation of '(co)'
blanchet
parents: 54145
diff changeset
   475
type coeqn_data_sel = {
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   476
  fun_name: string,
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   477
  fun_T: typ,
53401
2101a97e6220 various refactoring;
panny
parents: 53360
diff changeset
   478
  fun_args: term list,
53341
63015d035301 handle selector formulae with no corecursive calls
panny
parents: 53335
diff changeset
   479
  ctr: term,
63015d035301 handle selector formulae with no corecursive calls
panny
parents: 53335
diff changeset
   480
  sel: term,
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   481
  rhs_term: term,
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   482
  user_eqn: term
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   483
};
54001
65fc58793ed5 made SML/NJ happier
blanchet
parents: 53925
diff changeset
   484
54153
67487a607ce2 avoid 'co_' prefix with underscore meaning 'co', since it is our only possible identifier representation of '(co)'
blanchet
parents: 54145
diff changeset
   485
datatype coeqn_data =
67487a607ce2 avoid 'co_' prefix with underscore meaning 'co', since it is our only possible identifier representation of '(co)'
blanchet
parents: 54145
diff changeset
   486
  Disc of coeqn_data_disc |
67487a607ce2 avoid 'co_' prefix with underscore meaning 'co', since it is our only possible identifier representation of '(co)'
blanchet
parents: 54145
diff changeset
   487
  Sel of coeqn_data_sel;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   488
54160
a179353111db generate callssss
panny
parents: 54157
diff changeset
   489
fun dissect_coeqn_disc seq fun_names (basic_ctr_specss : basic_corec_ctr_spec list list)
a179353111db generate callssss
panny
parents: 54157
diff changeset
   490
    maybe_ctr_rhs maybe_code_rhs prems' concl matchedsss =
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   491
  let
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   492
    fun find_subterm p = let (* FIXME \<exists>? *)
53401
2101a97e6220 various refactoring;
panny
parents: 53360
diff changeset
   493
      fun f (t as u $ v) = if p t then SOME t else merge_options (f u, f v)
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   494
        | f t = if p t then SOME t else NONE
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   495
      in f end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   496
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   497
    val applied_fun = concl
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   498
      |> find_subterm (member ((op =) o apsnd SOME) fun_names o try (fst o dest_Free o head_of))
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   499
      |> the
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   500
      handle Option.Option => primrec_error_eqn "malformed discriminator equation" concl;
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   501
    val ((fun_name, fun_T), fun_args) = strip_comb applied_fun |>> dest_Free;
54160
a179353111db generate callssss
panny
parents: 54157
diff changeset
   502
    val basic_ctr_specs = the (AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name);
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   503
54160
a179353111db generate callssss
panny
parents: 54157
diff changeset
   504
    val discs = map #disc basic_ctr_specs;
a179353111db generate callssss
panny
parents: 54157
diff changeset
   505
    val ctrs = map #ctr basic_ctr_specs;
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   506
    val not_disc = head_of concl = @{term Not};
53401
2101a97e6220 various refactoring;
panny
parents: 53360
diff changeset
   507
    val _ = not_disc andalso length ctrs <> 2 andalso
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   508
      primrec_error_eqn "\<not>ed discriminator for a type with \<noteq> 2 constructors" concl;
54160
a179353111db generate callssss
panny
parents: 54157
diff changeset
   509
    val disc' = find_subterm (member (op =) discs o head_of) concl;
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   510
    val eq_ctr0 = concl |> perhaps (try (HOLogic.dest_not)) |> try (HOLogic.dest_eq #> snd)
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   511
        |> (fn SOME t => let val n = find_index (equal t) ctrs in
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   512
          if n >= 0 then SOME n else NONE end | _ => NONE);
54160
a179353111db generate callssss
panny
parents: 54157
diff changeset
   513
    val _ = is_some disc' orelse is_some eq_ctr0 orelse
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   514
      primrec_error_eqn "no discriminator in equation" concl;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   515
    val ctr_no' =
54160
a179353111db generate callssss
panny
parents: 54157
diff changeset
   516
      if is_none disc' then the eq_ctr0 else find_index (equal (head_of (the disc'))) discs;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   517
    val ctr_no = if not_disc then 1 - ctr_no' else ctr_no';
54160
a179353111db generate callssss
panny
parents: 54157
diff changeset
   518
    val {ctr, disc, ...} = nth basic_ctr_specs ctr_no;
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   519
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   520
    val catch_all = try (fst o dest_Free o the_single) prems' = SOME Name.uu_;
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   521
    val matchedss = AList.lookup (op =) matchedsss fun_name |> the_default [];
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   522
    val prems = map (abstract (List.rev fun_args)) prems';
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   523
    val real_prems =
54067
7be49e2bfccc rationalized negation code
blanchet
parents: 54065
diff changeset
   524
      (if catch_all orelse seq then maps s_not_conj matchedss else []) @
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   525
      (if catch_all then [] else prems);
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   526
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   527
    val matchedsss' = AList.delete (op =) fun_name matchedsss
54065
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
   528
      |> cons (fun_name, if seq then matchedss @ [prems] else matchedss @ [real_prems]);
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   529
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   530
    val user_eqn =
54097
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   531
      (real_prems, concl)
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   532
      |>> map HOLogic.mk_Trueprop ||> HOLogic.mk_Trueprop o abstract (List.rev fun_args)
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   533
      |> curry Logic.list_all (map dest_Free fun_args) o Logic.list_implies;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   534
  in
53341
63015d035301 handle selector formulae with no corecursive calls
panny
parents: 53335
diff changeset
   535
    (Disc {
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   536
      fun_name = fun_name,
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   537
      fun_T = fun_T,
53401
2101a97e6220 various refactoring;
panny
parents: 53360
diff changeset
   538
      fun_args = fun_args,
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   539
      ctr = ctr,
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   540
      ctr_no = ctr_no,
54160
a179353111db generate callssss
panny
parents: 54157
diff changeset
   541
      disc = disc,
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   542
      prems = real_prems,
53822
6304b12c7627 add "primcorec" command (cf. ae7f50e70c09)
panny
parents: 53811
diff changeset
   543
      auto_gen = catch_all,
54097
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   544
      maybe_ctr_rhs = maybe_ctr_rhs,
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   545
      maybe_code_rhs = maybe_code_rhs,
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   546
      user_eqn = user_eqn
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   547
    }, matchedsss')
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   548
  end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   549
54160
a179353111db generate callssss
panny
parents: 54157
diff changeset
   550
fun dissect_coeqn_sel fun_names (basic_ctr_specss : basic_corec_ctr_spec list list) eqn' of_spec
a179353111db generate callssss
panny
parents: 54157
diff changeset
   551
    eqn =
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   552
  let
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   553
    val (lhs, rhs) = HOLogic.dest_eq eqn
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   554
      handle TERM _ =>
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   555
        primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   556
    val sel = head_of lhs;
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   557
    val ((fun_name, fun_T), fun_args) = dest_comb lhs |> snd |> strip_comb |> apfst dest_Free
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   558
      handle TERM _ =>
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   559
        primrec_error_eqn "malformed selector argument in left-hand side" eqn;
54160
a179353111db generate callssss
panny
parents: 54157
diff changeset
   560
    val basic_ctr_specs = the (AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name)
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   561
      handle Option.Option => primrec_error_eqn "malformed selector argument in left-hand side" eqn;
54160
a179353111db generate callssss
panny
parents: 54157
diff changeset
   562
    val {ctr, ...} =
53831
80423b9080cf support "of" syntax to disambiguate selector equations
panny
parents: 53830
diff changeset
   563
      if is_some of_spec
54160
a179353111db generate callssss
panny
parents: 54157
diff changeset
   564
      then the (find_first (equal (the of_spec) o #ctr) basic_ctr_specs)
a179353111db generate callssss
panny
parents: 54157
diff changeset
   565
      else filter (exists (equal sel) o #sels) basic_ctr_specs |> the_single
53831
80423b9080cf support "of" syntax to disambiguate selector equations
panny
parents: 53830
diff changeset
   566
        handle List.Empty => primrec_error_eqn "ambiguous selector - use \"of\"" eqn;
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   567
    val user_eqn = drop_All eqn';
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   568
  in
53341
63015d035301 handle selector formulae with no corecursive calls
panny
parents: 53335
diff changeset
   569
    Sel {
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   570
      fun_name = fun_name,
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   571
      fun_T = fun_T,
53401
2101a97e6220 various refactoring;
panny
parents: 53360
diff changeset
   572
      fun_args = fun_args,
54160
a179353111db generate callssss
panny
parents: 54157
diff changeset
   573
      ctr = ctr,
53341
63015d035301 handle selector formulae with no corecursive calls
panny
parents: 53335
diff changeset
   574
      sel = sel,
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   575
      rhs_term = rhs,
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   576
      user_eqn = user_eqn
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   577
    }
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   578
  end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   579
54160
a179353111db generate callssss
panny
parents: 54157
diff changeset
   580
fun dissect_coeqn_ctr seq fun_names (basic_ctr_specss : basic_corec_ctr_spec list list) eqn'
a179353111db generate callssss
panny
parents: 54157
diff changeset
   581
    maybe_code_rhs prems concl matchedsss =
53910
2c5055a3583d strengthen tactic
blanchet
parents: 53909
diff changeset
   582
  let
54065
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
   583
    val (lhs, rhs) = HOLogic.dest_eq concl;
54097
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   584
    val (fun_name, fun_args) = strip_comb lhs |>> fst o dest_Free;
54160
a179353111db generate callssss
panny
parents: 54157
diff changeset
   585
    val basic_ctr_specs = the (AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name);
54074
43cdae9524bf allow 'let's around constructors in constructor view
blanchet
parents: 54068
diff changeset
   586
    val (ctr, ctr_args) = strip_comb (unfold_let rhs);
54160
a179353111db generate callssss
panny
parents: 54157
diff changeset
   587
    val {disc, sels, ...} = the (find_first (equal ctr o #ctr) basic_ctr_specs)
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   588
      handle Option.Option => primrec_error_eqn "not a constructor" ctr;
53341
63015d035301 handle selector formulae with no corecursive calls
panny
parents: 53335
diff changeset
   589
54065
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
   590
    val disc_concl = betapply (disc, lhs);
54160
a179353111db generate callssss
panny
parents: 54157
diff changeset
   591
    val (maybe_eqn_data_disc, matchedsss') = if length basic_ctr_specs = 1
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   592
      then (NONE, matchedsss)
54160
a179353111db generate callssss
panny
parents: 54157
diff changeset
   593
      else apfst SOME (dissect_coeqn_disc seq fun_names basic_ctr_specss
54097
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   594
          (SOME (abstract (List.rev fun_args) rhs)) maybe_code_rhs prems disc_concl matchedsss);
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   595
54159
eb5d58c99049 set stage for more flexible 'primrec' syntax for recursion through functions
blanchet
parents: 54157
diff changeset
   596
    val sel_concls = sels ~~ ctr_args
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   597
      |> map (fn (sel, ctr_arg) => HOLogic.mk_eq (betapply (sel, lhs), ctr_arg));
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   598
53856
54c8dee1295a commented out debugging output in "primcorec"
blanchet
parents: 53835
diff changeset
   599
(*
54065
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
   600
val _ = tracing ("reduced\n    " ^ Syntax.string_of_term @{context} concl ^ "\nto\n    \<cdot> " ^
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
   601
 (is_some maybe_eqn_data_disc ? K (Syntax.string_of_term @{context} disc_concl ^ "\n    \<cdot> ")) "" ^
54097
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   602
 space_implode "\n    \<cdot> " (map (Syntax.string_of_term @{context}) sel_concls) ^
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   603
 "\nfor premise(s)\n    \<cdot> " ^
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   604
 space_implode "\n    \<cdot> " (map (Syntax.string_of_term @{context}) prems));
53856
54c8dee1295a commented out debugging output in "primcorec"
blanchet
parents: 53835
diff changeset
   605
*)
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   606
54160
a179353111db generate callssss
panny
parents: 54157
diff changeset
   607
    val eqns_data_sel =
a179353111db generate callssss
panny
parents: 54157
diff changeset
   608
      map (dissect_coeqn_sel fun_names basic_ctr_specss eqn' (SOME ctr)) sel_concls;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   609
  in
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   610
    (the_list maybe_eqn_data_disc @ eqns_data_sel, matchedsss')
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   611
  end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   612
54160
a179353111db generate callssss
panny
parents: 54157
diff changeset
   613
fun dissect_coeqn_code lthy has_call fun_names basic_ctr_specss eqn' concl matchedsss =
54065
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
   614
  let
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
   615
    val (lhs, (rhs', rhs)) = HOLogic.dest_eq concl ||> `(expand_corec_code_rhs lthy has_call []);
54097
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   616
    val (fun_name, fun_args) = strip_comb lhs |>> fst o dest_Free;
54160
a179353111db generate callssss
panny
parents: 54157
diff changeset
   617
    val basic_ctr_specs = the (AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name);
54065
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
   618
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
   619
    val cond_ctrs = fold_rev_corec_code_rhs lthy (fn cs => fn ctr => fn _ =>
54160
a179353111db generate callssss
panny
parents: 54157
diff changeset
   620
        if member ((op =) o apsnd #ctr) basic_ctr_specs ctr
54065
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
   621
        then cons (ctr, cs)
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
   622
        else primrec_error_eqn "not a constructor" ctr) [] rhs' []
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
   623
      |> AList.group (op =);
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
   624
54068
447354985f6a generate optimized DNF formula
blanchet
parents: 54067
diff changeset
   625
    val ctr_premss = (case cond_ctrs of [_] => [[]] | _ => map (s_dnf o snd) cond_ctrs);
54065
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
   626
    val ctr_concls = cond_ctrs |> map (fn (ctr, _) =>
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
   627
        binder_types (fastype_of ctr)
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
   628
        |> map_index (fn (n, T) => massage_corec_code_rhs lthy (fn _ => fn ctr' => fn args =>
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
   629
          if ctr' = ctr then nth args n else Const (@{const_name undefined}, T)) [] rhs')
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
   630
        |> curry list_comb ctr
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
   631
        |> curry HOLogic.mk_eq lhs);
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
   632
  in
54160
a179353111db generate callssss
panny
parents: 54157
diff changeset
   633
    fold_map2 (dissect_coeqn_ctr false fun_names basic_ctr_specss eqn'
54097
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   634
        (SOME (abstract (List.rev fun_args) rhs)))
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   635
      ctr_premss ctr_concls matchedsss
54065
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
   636
  end;
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
   637
54160
a179353111db generate callssss
panny
parents: 54157
diff changeset
   638
fun dissect_coeqn lthy seq has_call fun_names (basic_ctr_specss : basic_corec_ctr_spec list list)
a179353111db generate callssss
panny
parents: 54157
diff changeset
   639
    eqn' of_spec matchedsss =
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   640
  let
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   641
    val eqn = drop_All eqn'
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   642
      handle TERM _ => primrec_error_eqn "malformed function equation" eqn';
54065
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
   643
    val (prems, concl) = Logic.strip_horn eqn
53341
63015d035301 handle selector formulae with no corecursive calls
panny
parents: 53335
diff changeset
   644
      |> apfst (map HOLogic.dest_Trueprop) o apsnd HOLogic.dest_Trueprop;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   645
54065
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
   646
    val head = concl
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   647
      |> perhaps (try HOLogic.dest_not) |> perhaps (try (fst o HOLogic.dest_eq))
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   648
      |> head_of;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   649
54065
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
   650
    val maybe_rhs = concl |> perhaps (try (HOLogic.dest_not)) |> try (snd o HOLogic.dest_eq);
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   651
54160
a179353111db generate callssss
panny
parents: 54157
diff changeset
   652
    val discs = maps (map #disc) basic_ctr_specss;
a179353111db generate callssss
panny
parents: 54157
diff changeset
   653
    val sels = maps (maps #sels) basic_ctr_specss;
a179353111db generate callssss
panny
parents: 54157
diff changeset
   654
    val ctrs = maps (map #ctr) basic_ctr_specss;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   655
  in
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   656
    if member (op =) discs head orelse
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   657
      is_some maybe_rhs andalso
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   658
        member (op =) (filter (null o binder_types o fastype_of) ctrs) (the maybe_rhs) then
54160
a179353111db generate callssss
panny
parents: 54157
diff changeset
   659
      dissect_coeqn_disc seq fun_names basic_ctr_specss NONE NONE prems concl matchedsss
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   660
      |>> single
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   661
    else if member (op =) sels head then
54160
a179353111db generate callssss
panny
parents: 54157
diff changeset
   662
      ([dissect_coeqn_sel fun_names basic_ctr_specss eqn' of_spec concl], matchedsss)
54065
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
   663
    else if is_Free head andalso member (op =) fun_names (fst (dest_Free head)) andalso
54074
43cdae9524bf allow 'let's around constructors in constructor view
blanchet
parents: 54068
diff changeset
   664
      member (op =) ctrs (head_of (unfold_let (the maybe_rhs))) then
54160
a179353111db generate callssss
panny
parents: 54157
diff changeset
   665
      dissect_coeqn_ctr seq fun_names basic_ctr_specss eqn' NONE prems concl matchedsss
54065
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
   666
    else if is_Free head andalso member (op =) fun_names (fst (dest_Free head)) andalso
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
   667
      null prems then
54160
a179353111db generate callssss
panny
parents: 54157
diff changeset
   668
      dissect_coeqn_code lthy has_call fun_names basic_ctr_specss eqn' concl matchedsss
54065
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
   669
      |>> flat
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   670
    else
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   671
      primrec_error_eqn "malformed function equation" eqn
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   672
  end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   673
54002
01c8f9d3b084 made SML/NJ happy
blanchet
parents: 54001
diff changeset
   674
fun build_corec_arg_disc (ctr_specs : corec_ctr_spec list)
54153
67487a607ce2 avoid 'co_' prefix with underscore meaning 'co', since it is our only possible identifier representation of '(co)'
blanchet
parents: 54145
diff changeset
   675
    ({fun_args, ctr_no, prems, ...} : coeqn_data_disc) =
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   676
  if is_none (#pred (nth ctr_specs ctr_no)) then I else
54068
447354985f6a generate optimized DNF formula
blanchet
parents: 54067
diff changeset
   677
    s_conjs prems
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   678
    |> curry subst_bounds (List.rev fun_args)
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   679
    |> HOLogic.tupled_lambda (HOLogic.mk_tuple fun_args)
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   680
    |> K |> nth_map (the (#pred (nth ctr_specs ctr_no)));
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   681
54153
67487a607ce2 avoid 'co_' prefix with underscore meaning 'co', since it is our only possible identifier representation of '(co)'
blanchet
parents: 54145
diff changeset
   682
fun build_corec_arg_no_call (sel_eqns : coeqn_data_sel list) sel =
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   683
  find_first (equal sel o #sel) sel_eqns
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   684
  |> try (fn SOME {fun_args, rhs_term, ...} => abs_tuple fun_args rhs_term)
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   685
  |> the_default undef_const
53411
ab4edf89992f support indirect corecursion
panny
parents: 53401
diff changeset
   686
  |> K;
53360
7ffc4a746a73 handle direct corecursion
panny
parents: 53358
diff changeset
   687
54153
67487a607ce2 avoid 'co_' prefix with underscore meaning 'co', since it is our only possible identifier representation of '(co)'
blanchet
parents: 54145
diff changeset
   688
fun build_corec_args_mutual_call lthy has_call (sel_eqns : coeqn_data_sel list) sel =
53360
7ffc4a746a73 handle direct corecursion
panny
parents: 53358
diff changeset
   689
  let
53411
ab4edf89992f support indirect corecursion
panny
parents: 53401
diff changeset
   690
    val maybe_sel_eqn = find_first (equal sel o #sel) sel_eqns;
53360
7ffc4a746a73 handle direct corecursion
panny
parents: 53358
diff changeset
   691
  in
53876
fabf04d43a75 simplified code
panny
parents: 53875
diff changeset
   692
    if is_none maybe_sel_eqn then (I, I, I) else
fabf04d43a75 simplified code
panny
parents: 53875
diff changeset
   693
    let
fabf04d43a75 simplified code
panny
parents: 53875
diff changeset
   694
      val {fun_args, rhs_term, ... } = the maybe_sel_eqn;
54097
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   695
      val bound_Ts = List.rev (map fastype_of fun_args);
53890
5f647a5bd46e thread through bound types
blanchet
parents: 53889
diff changeset
   696
      fun rewrite_q _ t = if has_call t then @{term False} else @{term True};
5f647a5bd46e thread through bound types
blanchet
parents: 53889
diff changeset
   697
      fun rewrite_g _ t = if has_call t then undef_const else t;
53899
e55b634ff9fb simplified code
panny
parents: 53890
diff changeset
   698
      fun rewrite_h bound_Ts t =
e55b634ff9fb simplified code
panny
parents: 53890
diff changeset
   699
        if has_call t then mk_tuple1 bound_Ts (snd (strip_comb t)) else undef_const;
54102
82ada0a92dde tuned names
blanchet
parents: 54101
diff changeset
   700
      fun massage f _ = massage_mutual_corec_call lthy has_call f bound_Ts rhs_term
54097
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   701
        |> abs_tuple fun_args;
53876
fabf04d43a75 simplified code
panny
parents: 53875
diff changeset
   702
    in
fabf04d43a75 simplified code
panny
parents: 53875
diff changeset
   703
      (massage rewrite_q,
fabf04d43a75 simplified code
panny
parents: 53875
diff changeset
   704
       massage rewrite_g,
fabf04d43a75 simplified code
panny
parents: 53875
diff changeset
   705
       massage rewrite_h)
fabf04d43a75 simplified code
panny
parents: 53875
diff changeset
   706
    end
53360
7ffc4a746a73 handle direct corecursion
panny
parents: 53358
diff changeset
   707
  end;
7ffc4a746a73 handle direct corecursion
panny
parents: 53358
diff changeset
   708
54153
67487a607ce2 avoid 'co_' prefix with underscore meaning 'co', since it is our only possible identifier representation of '(co)'
blanchet
parents: 54145
diff changeset
   709
fun build_corec_arg_nested_call lthy has_call (sel_eqns : coeqn_data_sel list) sel =
53411
ab4edf89992f support indirect corecursion
panny
parents: 53401
diff changeset
   710
  let
ab4edf89992f support indirect corecursion
panny
parents: 53401
diff changeset
   711
    val maybe_sel_eqn = find_first (equal sel o #sel) sel_eqns;
53899
e55b634ff9fb simplified code
panny
parents: 53890
diff changeset
   712
  in
e55b634ff9fb simplified code
panny
parents: 53890
diff changeset
   713
    if is_none maybe_sel_eqn then I else
e55b634ff9fb simplified code
panny
parents: 53890
diff changeset
   714
    let
e55b634ff9fb simplified code
panny
parents: 53890
diff changeset
   715
      val {fun_args, rhs_term, ...} = the maybe_sel_eqn;
54097
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   716
      val bound_Ts = List.rev (map fastype_of fun_args);
53899
e55b634ff9fb simplified code
panny
parents: 53890
diff changeset
   717
      fun rewrite bound_Ts U T (Abs (v, V, b)) = Abs (v, V, rewrite (V :: bound_Ts) U T b)
e55b634ff9fb simplified code
panny
parents: 53890
diff changeset
   718
        | rewrite bound_Ts U T (t as _ $ _) =
e55b634ff9fb simplified code
panny
parents: 53890
diff changeset
   719
          let val (u, vs) = strip_comb t in
e55b634ff9fb simplified code
panny
parents: 53890
diff changeset
   720
            if is_Free u andalso has_call u then
e55b634ff9fb simplified code
panny
parents: 53890
diff changeset
   721
              Inr_const U T $ mk_tuple1 bound_Ts vs
e55b634ff9fb simplified code
panny
parents: 53890
diff changeset
   722
            else if try (fst o dest_Const) u = SOME @{const_name prod_case} then
e55b634ff9fb simplified code
panny
parents: 53890
diff changeset
   723
              map (rewrite bound_Ts U T) vs |> chop 1 |>> HOLogic.mk_split o the_single |> list_comb
e55b634ff9fb simplified code
panny
parents: 53890
diff changeset
   724
            else
e55b634ff9fb simplified code
panny
parents: 53890
diff changeset
   725
              list_comb (rewrite bound_Ts U T u, map (rewrite bound_Ts U T) vs)
e55b634ff9fb simplified code
panny
parents: 53890
diff changeset
   726
          end
e55b634ff9fb simplified code
panny
parents: 53890
diff changeset
   727
        | rewrite _ U T t =
e55b634ff9fb simplified code
panny
parents: 53890
diff changeset
   728
          if is_Free t andalso has_call t then Inr_const U T $ HOLogic.unit else t;
e55b634ff9fb simplified code
panny
parents: 53890
diff changeset
   729
      fun massage t =
54097
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   730
        rhs_term
54102
82ada0a92dde tuned names
blanchet
parents: 54101
diff changeset
   731
        |> massage_nested_corec_call lthy has_call rewrite bound_Ts (range_type (fastype_of t))
53735
99331dac1e1c simplified code; eliminated some dummyTs
panny
parents: 53734
diff changeset
   732
        |> abs_tuple fun_args;
53899
e55b634ff9fb simplified code
panny
parents: 53890
diff changeset
   733
    in
e55b634ff9fb simplified code
panny
parents: 53890
diff changeset
   734
      massage
e55b634ff9fb simplified code
panny
parents: 53890
diff changeset
   735
    end
53411
ab4edf89992f support indirect corecursion
panny
parents: 53401
diff changeset
   736
  end;
53360
7ffc4a746a73 handle direct corecursion
panny
parents: 53358
diff changeset
   737
54153
67487a607ce2 avoid 'co_' prefix with underscore meaning 'co', since it is our only possible identifier representation of '(co)'
blanchet
parents: 54145
diff changeset
   738
fun build_corec_args_sel lthy has_call (all_sel_eqns : coeqn_data_sel list)
54002
01c8f9d3b084 made SML/NJ happy
blanchet
parents: 54001
diff changeset
   739
    (ctr_spec : corec_ctr_spec) =
53341
63015d035301 handle selector formulae with no corecursive calls
panny
parents: 53335
diff changeset
   740
  let val sel_eqns = filter (equal (#ctr ctr_spec) o #ctr) all_sel_eqns in
63015d035301 handle selector formulae with no corecursive calls
panny
parents: 53335
diff changeset
   741
    if null sel_eqns then I else
63015d035301 handle selector formulae with no corecursive calls
panny
parents: 53335
diff changeset
   742
      let
63015d035301 handle selector formulae with no corecursive calls
panny
parents: 53335
diff changeset
   743
        val sel_call_list = #sels ctr_spec ~~ #calls ctr_spec;
63015d035301 handle selector formulae with no corecursive calls
panny
parents: 53335
diff changeset
   744
63015d035301 handle selector formulae with no corecursive calls
panny
parents: 53335
diff changeset
   745
        val no_calls' = map_filter (try (apsnd (fn No_Corec n => n))) sel_call_list;
54102
82ada0a92dde tuned names
blanchet
parents: 54101
diff changeset
   746
        val mutual_calls' = map_filter (try (apsnd (fn Mutual_Corec n => n))) sel_call_list;
82ada0a92dde tuned names
blanchet
parents: 54101
diff changeset
   747
        val nested_calls' = map_filter (try (apsnd (fn Nested_Corec n => n))) sel_call_list;
53341
63015d035301 handle selector formulae with no corecursive calls
panny
parents: 53335
diff changeset
   748
      in
53360
7ffc4a746a73 handle direct corecursion
panny
parents: 53358
diff changeset
   749
        I
53735
99331dac1e1c simplified code; eliminated some dummyTs
panny
parents: 53734
diff changeset
   750
        #> fold (fn (sel, n) => nth_map n (build_corec_arg_no_call sel_eqns sel)) no_calls'
53360
7ffc4a746a73 handle direct corecursion
panny
parents: 53358
diff changeset
   751
        #> fold (fn (sel, (q, g, h)) =>
54102
82ada0a92dde tuned names
blanchet
parents: 54101
diff changeset
   752
          let val (fq, fg, fh) = build_corec_args_mutual_call lthy has_call sel_eqns sel in
82ada0a92dde tuned names
blanchet
parents: 54101
diff changeset
   753
            nth_map q fq o nth_map g fg o nth_map h fh end) mutual_calls'
53360
7ffc4a746a73 handle direct corecursion
panny
parents: 53358
diff changeset
   754
        #> fold (fn (sel, n) => nth_map n
54102
82ada0a92dde tuned names
blanchet
parents: 54101
diff changeset
   755
          (build_corec_arg_nested_call lthy has_call sel_eqns sel)) nested_calls'
53341
63015d035301 handle selector formulae with no corecursive calls
panny
parents: 53335
diff changeset
   756
      end
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   757
  end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   758
54153
67487a607ce2 avoid 'co_' prefix with underscore meaning 'co', since it is our only possible identifier representation of '(co)'
blanchet
parents: 54145
diff changeset
   759
fun build_codefs lthy bs mxs has_call arg_Tss (corec_specs : corec_spec list)
67487a607ce2 avoid 'co_' prefix with underscore meaning 'co', since it is our only possible identifier representation of '(co)'
blanchet
parents: 54145
diff changeset
   760
    (disc_eqnss : coeqn_data_disc list list) (sel_eqnss : coeqn_data_sel list list) =
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   761
  let
54160
a179353111db generate callssss
panny
parents: 54157
diff changeset
   762
    val corecs = map #corec corec_specs;
a179353111db generate callssss
panny
parents: 54157
diff changeset
   763
    val ctr_specss = map #ctr_specs corec_specs;
53360
7ffc4a746a73 handle direct corecursion
panny
parents: 53358
diff changeset
   764
    val corec_args = hd corecs
7ffc4a746a73 handle direct corecursion
panny
parents: 53358
diff changeset
   765
      |> fst o split_last o binder_types o fastype_of
7ffc4a746a73 handle direct corecursion
panny
parents: 53358
diff changeset
   766
      |> map (Const o pair @{const_name undefined})
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   767
      |> fold2 (fold o build_corec_arg_disc) ctr_specss disc_eqnss
53360
7ffc4a746a73 handle direct corecursion
panny
parents: 53358
diff changeset
   768
      |> fold2 (fold o build_corec_args_sel lthy has_call) sel_eqnss ctr_specss;
53735
99331dac1e1c simplified code; eliminated some dummyTs
panny
parents: 53734
diff changeset
   769
    fun currys [] t = t
99331dac1e1c simplified code; eliminated some dummyTs
panny
parents: 53734
diff changeset
   770
      | currys Ts t = t $ mk_tuple1 (List.rev Ts) (map Bound (length Ts - 1 downto 0))
99331dac1e1c simplified code; eliminated some dummyTs
panny
parents: 53734
diff changeset
   771
          |> fold_rev (Term.abs o pair Name.uu) Ts;
53401
2101a97e6220 various refactoring;
panny
parents: 53360
diff changeset
   772
53856
54c8dee1295a commented out debugging output in "primcorec"
blanchet
parents: 53835
diff changeset
   773
(*
53360
7ffc4a746a73 handle direct corecursion
panny
parents: 53358
diff changeset
   774
val _ = tracing ("corecursor arguments:\n    \<cdot> " ^
53411
ab4edf89992f support indirect corecursion
panny
parents: 53401
diff changeset
   775
 space_implode "\n    \<cdot> " (map (Syntax.string_of_term lthy) corec_args));
53856
54c8dee1295a commented out debugging output in "primcorec"
blanchet
parents: 53835
diff changeset
   776
*)
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   777
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   778
    val exclss' =
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   779
      disc_eqnss
53822
6304b12c7627 add "primcorec" command (cf. ae7f50e70c09)
panny
parents: 53811
diff changeset
   780
      |> map (map (fn x => (#fun_args x, #ctr_no x, #prems x, #auto_gen x))
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   781
        #> fst o (fn xs => fold_map (fn x => fn ys => ((x, ys), ys @ [x])) xs [])
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   782
        #> maps (uncurry (map o pair)
53822
6304b12c7627 add "primcorec" command (cf. ae7f50e70c09)
panny
parents: 53811
diff changeset
   783
          #> map (fn ((fun_args, c, x, a), (_, c', y, a')) =>
54068
447354985f6a generate optimized DNF formula
blanchet
parents: 54067
diff changeset
   784
              ((c, c', a orelse a'), (x, s_not (s_conjs y)))
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   785
            ||> apfst (map HOLogic.mk_Trueprop) o apsnd HOLogic.mk_Trueprop
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   786
            ||> Logic.list_implies
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   787
            ||> curry Logic.list_all (map dest_Free fun_args))))
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   788
  in
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   789
    map (list_comb o rpair corec_args) corecs
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   790
    |> map2 (fn Ts => fn t => if length Ts = 0 then t $ HOLogic.unit else t) arg_Tss
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   791
    |> map2 currys arg_Tss
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   792
    |> Syntax.check_terms lthy
54155
b964fad0cbbd conceal more ugly constructions
blanchet
parents: 54154
diff changeset
   793
    |> map3 (fn b => fn mx => fn t => ((b, mx), ((Binding.conceal (Thm.def_binding b), []), t)))
b964fad0cbbd conceal more ugly constructions
blanchet
parents: 54154
diff changeset
   794
      bs mxs
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   795
    |> rpair exclss'
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   796
  end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   797
54002
01c8f9d3b084 made SML/NJ happy
blanchet
parents: 54001
diff changeset
   798
fun mk_real_disc_eqns fun_binding arg_Ts ({ctr_specs, ...} : corec_spec)
54153
67487a607ce2 avoid 'co_' prefix with underscore meaning 'co', since it is our only possible identifier representation of '(co)'
blanchet
parents: 54145
diff changeset
   799
    (sel_eqns : coeqn_data_sel list) (disc_eqns : coeqn_data_disc list) =
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   800
  if length disc_eqns <> length ctr_specs - 1 then disc_eqns else
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   801
    let
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   802
      val n = 0 upto length ctr_specs
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   803
        |> the o find_first (fn idx => not (exists (equal idx o #ctr_no) disc_eqns));
53722
e176d6d3345f generate more theorems (e.g. for types with only one constructor)
panny
parents: 53720
diff changeset
   804
      val fun_args = (try (#fun_args o hd) disc_eqns, try (#fun_args o hd) sel_eqns)
e176d6d3345f generate more theorems (e.g. for types with only one constructor)
panny
parents: 53720
diff changeset
   805
        |> the_default (map (curry Free Name.uu) arg_Ts) o merge_options;
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   806
      val extra_disc_eqn = {
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   807
        fun_name = Binding.name_of fun_binding,
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   808
        fun_T = arg_Ts ---> body_type (fastype_of (#ctr (hd ctr_specs))),
53722
e176d6d3345f generate more theorems (e.g. for types with only one constructor)
panny
parents: 53720
diff changeset
   809
        fun_args = fun_args,
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   810
        ctr = #ctr (nth ctr_specs n),
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   811
        ctr_no = n,
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   812
        disc = #disc (nth ctr_specs n),
54067
7be49e2bfccc rationalized negation code
blanchet
parents: 54065
diff changeset
   813
        prems = maps (s_not_conj o #prems) disc_eqns,
53822
6304b12c7627 add "primcorec" command (cf. ae7f50e70c09)
panny
parents: 53811
diff changeset
   814
        auto_gen = true,
54097
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   815
        maybe_ctr_rhs = NONE,
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   816
        maybe_code_rhs = NONE,
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   817
        user_eqn = undef_const};
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   818
    in
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   819
      chop n disc_eqns ||> cons extra_disc_eqn |> (op @)
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   820
    end;
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   821
54160
a179353111db generate callssss
panny
parents: 54157
diff changeset
   822
fun find_corec_calls has_call basic_ctr_specs {ctr, sel, rhs_term, ...} =
a179353111db generate callssss
panny
parents: 54157
diff changeset
   823
  let
a179353111db generate callssss
panny
parents: 54157
diff changeset
   824
    val sel_no = find_first (equal ctr o #ctr) basic_ctr_specs
a179353111db generate callssss
panny
parents: 54157
diff changeset
   825
      |> find_index (equal sel) o #sels o the;
a179353111db generate callssss
panny
parents: 54157
diff changeset
   826
    fun find (Abs (_, _, b)) = find b
a179353111db generate callssss
panny
parents: 54157
diff changeset
   827
      | find (t as _ $ _) = strip_comb t |>> find ||> maps find |> (op @)
a179353111db generate callssss
panny
parents: 54157
diff changeset
   828
      | find f = if is_Free f andalso has_call f then [f] else [];
a179353111db generate callssss
panny
parents: 54157
diff changeset
   829
  in
a179353111db generate callssss
panny
parents: 54157
diff changeset
   830
    find rhs_term
a179353111db generate callssss
panny
parents: 54157
diff changeset
   831
    |> K |> nth_map sel_no |> AList.map_entry (op =) ctr
a179353111db generate callssss
panny
parents: 54157
diff changeset
   832
  end;
a179353111db generate callssss
panny
parents: 54157
diff changeset
   833
54065
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
   834
fun add_primcorec simple seq fixes specs of_specs lthy =
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   835
  let
53352
43a1cc050943 honor mixfix specifications
traytel
parents: 53341
diff changeset
   836
    val (bs, mxs) = map_split (apfst fst) fixes;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   837
    val (arg_Ts, res_Ts) = map (strip_type o snd o fst #>> HOLogic.mk_tupleT) fixes |> split_list;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   838
54160
a179353111db generate callssss
panny
parents: 54157
diff changeset
   839
    val fun_names = map Binding.name_of bs;
a179353111db generate callssss
panny
parents: 54157
diff changeset
   840
    val basic_ctr_specss = map (basic_corec_specs_of lthy) res_Ts;
a179353111db generate callssss
panny
parents: 54157
diff changeset
   841
    val has_call = exists_subterm (map (fst #>> Binding.name_of #> Free) fixes |> member (op =));
a179353111db generate callssss
panny
parents: 54157
diff changeset
   842
    val eqns_data =
a179353111db generate callssss
panny
parents: 54157
diff changeset
   843
      fold_map2 (dissect_coeqn lthy seq has_call fun_names basic_ctr_specss) (map snd specs)
a179353111db generate callssss
panny
parents: 54157
diff changeset
   844
        of_specs []
a179353111db generate callssss
panny
parents: 54157
diff changeset
   845
      |> flat o fst;
a179353111db generate callssss
panny
parents: 54157
diff changeset
   846
a179353111db generate callssss
panny
parents: 54157
diff changeset
   847
    val callssss =
a179353111db generate callssss
panny
parents: 54157
diff changeset
   848
      map_filter (try (fn Sel x => x)) eqns_data
a179353111db generate callssss
panny
parents: 54157
diff changeset
   849
      |> partition_eq ((op =) o pairself #fun_name)
a179353111db generate callssss
panny
parents: 54157
diff changeset
   850
      |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
54161
panny
parents: 54160 54159
diff changeset
   851
      |> map (flat o snd)
panny
parents: 54160 54159
diff changeset
   852
      |> map2 (fold o find_corec_calls has_call) basic_ctr_specss
54160
a179353111db generate callssss
panny
parents: 54157
diff changeset
   853
      |> map2 (curry (op |>)) (map (map (fn {ctr, sels, ...} =>
a179353111db generate callssss
panny
parents: 54157
diff changeset
   854
        (ctr, map (K []) sels))) basic_ctr_specss);
a179353111db generate callssss
panny
parents: 54157
diff changeset
   855
a179353111db generate callssss
panny
parents: 54157
diff changeset
   856
(*
a179353111db generate callssss
panny
parents: 54157
diff changeset
   857
val _ = tracing ("callssss = " ^ @{make_string} callssss);
a179353111db generate callssss
panny
parents: 54157
diff changeset
   858
*)
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   859
53830
ed2eb7df2aac don't note more induction principles than there are functions + tuning
blanchet
parents: 53822
diff changeset
   860
    val ((n2m, corec_specs', _, coinduct_thm, strong_coinduct_thm, coinduct_thms,
53797
576f9637dc7a note coinduct theorems in "primcorec"
blanchet
parents: 53795
diff changeset
   861
          strong_coinduct_thms), lthy') =
53794
af7d1533a25b undid copy-paste
blanchet
parents: 53793
diff changeset
   862
      corec_specs_of bs arg_Ts res_Ts (get_indices fixes) callssss lthy;
53830
ed2eb7df2aac don't note more induction principles than there are functions + tuning
blanchet
parents: 53822
diff changeset
   863
    val actual_nn = length bs;
ed2eb7df2aac don't note more induction principles than there are functions + tuning
blanchet
parents: 53822
diff changeset
   864
    val corec_specs = take actual_nn corec_specs'; (*###*)
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   865
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   866
    val disc_eqnss' = map_filter (try (fn Disc x => x)) eqns_data
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   867
      |> partition_eq ((op =) o pairself #fun_name)
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   868
      |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   869
      |> map (sort ((op <) o pairself #ctr_no |> make_ord) o flat o snd);
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   870
    val _ = disc_eqnss' |> map (fn x =>
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   871
      let val d = duplicates ((op =) o pairself #ctr_no) x in null d orelse
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   872
        primrec_error_eqns "excess discriminator equations in definition"
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   873
          (maps (fn t => filter (equal (#ctr_no t) o #ctr_no) x) d |> map #user_eqn) end);
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   874
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   875
    val sel_eqnss = map_filter (try (fn Sel x => x)) eqns_data
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   876
      |> partition_eq ((op =) o pairself #fun_name)
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   877
      |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   878
      |> map (flat o snd);
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   879
53360
7ffc4a746a73 handle direct corecursion
panny
parents: 53358
diff changeset
   880
    val arg_Tss = map (binder_types o snd o fst) fixes;
53722
e176d6d3345f generate more theorems (e.g. for types with only one constructor)
panny
parents: 53720
diff changeset
   881
    val disc_eqnss = map5 mk_real_disc_eqns bs arg_Tss corec_specs sel_eqnss disc_eqnss';
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   882
    val (defs, exclss') =
54153
67487a607ce2 avoid 'co_' prefix with underscore meaning 'co', since it is our only possible identifier representation of '(co)'
blanchet
parents: 54145
diff changeset
   883
      build_codefs lthy' bs mxs has_call arg_Tss corec_specs disc_eqnss sel_eqnss;
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   884
53923
blanchet
parents: 53918
diff changeset
   885
    fun excl_tac (c, c', a) =
54065
e30e63d05e58 process code-style inputs
panny
parents: 54044
diff changeset
   886
      if a orelse c = c' orelse seq then
54044
93ab44e992ae strengthened tactic (primcorec_sel_tac) + tuning
blanchet
parents: 54030
diff changeset
   887
        SOME (K (HEADGOAL (mk_primcorec_assumption_tac lthy [])))
93ab44e992ae strengthened tactic (primcorec_sel_tac) + tuning
blanchet
parents: 54030
diff changeset
   888
      else if simple then
93ab44e992ae strengthened tactic (primcorec_sel_tac) + tuning
blanchet
parents: 54030
diff changeset
   889
        SOME (K (auto_tac lthy))
93ab44e992ae strengthened tactic (primcorec_sel_tac) + tuning
blanchet
parents: 54030
diff changeset
   890
      else
93ab44e992ae strengthened tactic (primcorec_sel_tac) + tuning
blanchet
parents: 54030
diff changeset
   891
        NONE;
53822
6304b12c7627 add "primcorec" command (cf. ae7f50e70c09)
panny
parents: 53811
diff changeset
   892
53856
54c8dee1295a commented out debugging output in "primcorec"
blanchet
parents: 53835
diff changeset
   893
(*
53822
6304b12c7627 add "primcorec" command (cf. ae7f50e70c09)
panny
parents: 53811
diff changeset
   894
val _ = tracing ("exclusiveness properties:\n    \<cdot> " ^
6304b12c7627 add "primcorec" command (cf. ae7f50e70c09)
panny
parents: 53811
diff changeset
   895
 space_implode "\n    \<cdot> " (maps (map (Syntax.string_of_term lthy o snd)) exclss'));
53856
54c8dee1295a commented out debugging output in "primcorec"
blanchet
parents: 53835
diff changeset
   896
*)
53822
6304b12c7627 add "primcorec" command (cf. ae7f50e70c09)
panny
parents: 53811
diff changeset
   897
6304b12c7627 add "primcorec" command (cf. ae7f50e70c09)
panny
parents: 53811
diff changeset
   898
    val exclss'' = exclss' |> map (map (fn (idx, t) =>
53923
blanchet
parents: 53918
diff changeset
   899
      (idx, (Option.map (Goal.prove lthy [] [] t) (excl_tac idx), t))));
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   900
    val taut_thmss = map (map (apsnd (the o fst)) o filter (is_some o fst o snd)) exclss'';
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   901
    val (obligation_idxss, obligationss) = exclss''
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   902
      |> map (map (apsnd (rpair [] o snd)) o filter (is_none o fst o snd))
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   903
      |> split_list o map split_list;
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   904
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   905
    fun prove thmss' def_thms' lthy =
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   906
      let
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   907
        val def_thms = map (snd o snd) def_thms';
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   908
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   909
        val exclss' = map (op ~~) (obligation_idxss ~~ thmss');
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   910
        fun mk_exclsss excls n =
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   911
          (excls, map (fn k => replicate k [TrueI] @ replicate (n - k) []) (0 upto n - 1))
53822
6304b12c7627 add "primcorec" command (cf. ae7f50e70c09)
panny
parents: 53811
diff changeset
   912
          |-> fold (fn ((c, c', _), thm) => nth_map c (nth_map c' (K [thm])));
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   913
        val exclssss = (exclss' ~~ taut_thmss |> map (op @), fun_names ~~ corec_specs)
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   914
          |-> map2 (fn excls => fn (_, {ctr_specs, ...}) => mk_exclsss excls (length ctr_specs));
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   915
54002
01c8f9d3b084 made SML/NJ happy
blanchet
parents: 54001
diff changeset
   916
        fun prove_disc ({ctr_specs, ...} : corec_spec) exclsss
54153
67487a607ce2 avoid 'co_' prefix with underscore meaning 'co', since it is our only possible identifier representation of '(co)'
blanchet
parents: 54145
diff changeset
   917
            ({fun_name, fun_T, fun_args, ctr_no, prems, ...} : coeqn_data_disc) =
53722
e176d6d3345f generate more theorems (e.g. for types with only one constructor)
panny
parents: 53720
diff changeset
   918
          if Term.aconv_untyped (#disc (nth ctr_specs ctr_no), @{term "\<lambda>x. x = x"}) then [] else
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   919
            let
53722
e176d6d3345f generate more theorems (e.g. for types with only one constructor)
panny
parents: 53720
diff changeset
   920
              val {disc_corec, ...} = nth ctr_specs ctr_no;
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   921
              val k = 1 + ctr_no;
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   922
              val m = length prems;
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   923
              val t =
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   924
                list_comb (Free (fun_name, fun_T), map Bound (length fun_args - 1 downto 0))
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   925
                |> curry betapply (#disc (nth ctr_specs ctr_no)) (*###*)
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   926
                |> HOLogic.mk_Trueprop
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   927
                |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   928
                |> curry Logic.list_all (map dest_Free fun_args);
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   929
            in
54097
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   930
              if prems = [@{term False}] then [] else
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   931
              mk_primcorec_disc_tac lthy def_thms disc_corec k m exclsss
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   932
              |> K |> Goal.prove lthy [] [] t
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   933
              |> pair (#disc (nth ctr_specs ctr_no))
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   934
              |> single
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   935
            end;
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   936
54002
01c8f9d3b084 made SML/NJ happy
blanchet
parents: 54001
diff changeset
   937
        fun prove_sel ({nested_maps, nested_map_idents, nested_map_comps, ctr_specs, ...}
54153
67487a607ce2 avoid 'co_' prefix with underscore meaning 'co', since it is our only possible identifier representation of '(co)'
blanchet
parents: 54145
diff changeset
   938
            : corec_spec) (disc_eqns : coeqn_data_disc list) exclsss
67487a607ce2 avoid 'co_' prefix with underscore meaning 'co', since it is our only possible identifier representation of '(co)'
blanchet
parents: 54145
diff changeset
   939
            ({fun_name, fun_T, fun_args, ctr, sel, rhs_term, ...} : coeqn_data_sel) =
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   940
          let
53909
7c10e75e62b3 use needed case theorems
blanchet
parents: 53903
diff changeset
   941
            val SOME ctr_spec = find_first (equal ctr o #ctr) ctr_specs;
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   942
            val ctr_no = find_index (equal ctr o #ctr) ctr_specs;
54067
7be49e2bfccc rationalized negation code
blanchet
parents: 54065
diff changeset
   943
            val prems = the_default (maps (s_not_conj o #prems) disc_eqns)
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   944
                (find_first (equal ctr_no o #ctr_no) disc_eqns |> Option.map #prems);
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   945
            val sel_corec = find_index (equal sel) (#sels ctr_spec)
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   946
              |> nth (#sel_corecs ctr_spec);
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   947
            val k = 1 + ctr_no;
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   948
            val m = length prems;
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   949
            val t =
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   950
              list_comb (Free (fun_name, fun_T), map Bound (length fun_args - 1 downto 0))
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   951
              |> curry betapply sel
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   952
              |> rpair (abstract (List.rev fun_args) rhs_term)
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   953
              |> HOLogic.mk_Trueprop o HOLogic.mk_eq
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   954
              |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   955
              |> curry Logic.list_all (map dest_Free fun_args);
53925
blanchet
parents: 53923
diff changeset
   956
            val (distincts, _, sel_splits, sel_split_asms) = case_thms_of_term lthy [] rhs_term;
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   957
          in
53918
0fc622be0185 use new "sel_split(_asm)" to avoid giving rise to quantifiers, which would in turn require relying on injectivity
blanchet
parents: 53910
diff changeset
   958
            mk_primcorec_sel_tac lthy def_thms distincts sel_splits sel_split_asms nested_maps
53910
2c5055a3583d strengthen tactic
blanchet
parents: 53909
diff changeset
   959
              nested_map_idents nested_map_comps sel_corec k m exclsss
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   960
            |> K |> Goal.prove lthy [] [] t
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   961
            |> pair sel
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   962
          end;
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
   963
54153
67487a607ce2 avoid 'co_' prefix with underscore meaning 'co', since it is our only possible identifier representation of '(co)'
blanchet
parents: 54145
diff changeset
   964
        fun prove_ctr disc_alist sel_alist (disc_eqns : coeqn_data_disc list)
67487a607ce2 avoid 'co_' prefix with underscore meaning 'co', since it is our only possible identifier representation of '(co)'
blanchet
parents: 54145
diff changeset
   965
            (sel_eqns : coeqn_data_sel list) ({ctr, disc, sels, collapse, ...} : corec_ctr_spec) =
54097
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   966
          (* don't try to prove theorems when some sel_eqns are missing *)
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   967
          if not (exists (equal ctr o #ctr) disc_eqns)
53722
e176d6d3345f generate more theorems (e.g. for types with only one constructor)
panny
parents: 53720
diff changeset
   968
              andalso not (exists (equal ctr o #ctr) sel_eqns)
54097
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   969
            orelse
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   970
              filter (equal ctr o #ctr) sel_eqns
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   971
              |> fst o finds ((op =) o apsnd #sel) sels
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   972
              |> exists (null o snd)
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   973
          then [] else
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   974
            let
54097
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   975
              val (fun_name, fun_T, fun_args, prems, maybe_rhs) =
53722
e176d6d3345f generate more theorems (e.g. for types with only one constructor)
panny
parents: 53720
diff changeset
   976
                (find_first (equal ctr o #ctr) disc_eqns, find_first (equal ctr o #ctr) sel_eqns)
54097
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   977
                |>> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, #prems x,
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   978
                  #maybe_ctr_rhs x))
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   979
                ||> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, [], NONE))
53722
e176d6d3345f generate more theorems (e.g. for types with only one constructor)
panny
parents: 53720
diff changeset
   980
                |> the o merge_options;
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   981
              val m = length prems;
54097
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   982
              val t = (if is_some maybe_rhs then the maybe_rhs else
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   983
                  filter (equal ctr o #ctr) sel_eqns
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   984
                  |> fst o finds ((op =) o apsnd #sel) sels
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   985
                  |> map (snd #> (fn [x] => (List.rev (#fun_args x), #rhs_term x)) #-> abstract)
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   986
                  |> curry list_comb ctr)
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   987
                |> curry HOLogic.mk_eq (list_comb (Free (fun_name, fun_T),
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   988
                  map Bound (length fun_args - 1 downto 0)))
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   989
                |> HOLogic.mk_Trueprop
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   990
                |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   991
                |> curry Logic.list_all (map dest_Free fun_args);
53791
0ddf045113c9 tuned code
blanchet
parents: 53753
diff changeset
   992
              val maybe_disc_thm = AList.lookup (op =) disc_alist disc;
0ddf045113c9 tuned code
blanchet
parents: 53753
diff changeset
   993
              val sel_thms = map snd (filter (member (op =) sels o fst) sel_alist);
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
   994
            in
54097
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   995
              if prems = [@{term False}] then [] else
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   996
                mk_primcorec_ctr_of_dtr_tac lthy m collapse maybe_disc_thm sel_thms
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   997
                |> K |> Goal.prove lthy [] [] t
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   998
                |> pair ctr
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
   999
                |> single
53876
fabf04d43a75 simplified code
panny
parents: 53875
diff changeset
  1000
            end;
53720
03fac7082137 generate constructor view theorems
panny
parents: 53699
diff changeset
  1001
54098
07a8145aaeba pass the right theorems to tactic
panny
parents: 54097
diff changeset
  1002
        fun prove_code disc_eqns sel_eqns ctr_alist {ctr_specs, ...} =
07a8145aaeba pass the right theorems to tactic
panny
parents: 54097
diff changeset
  1003
          let
54097
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1004
            val (fun_name, fun_T, fun_args, maybe_rhs) =
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1005
              (find_first (member (op =) (map #ctr ctr_specs) o #ctr) disc_eqns,
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1006
               find_first (member (op =) (map #ctr ctr_specs) o #ctr) sel_eqns)
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1007
              |>> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, #maybe_code_rhs x))
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1008
              ||> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, NONE))
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1009
              |> the o merge_options;
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1010
54120
c2f18fd05414 use Code.abort instead of undefined in auto-generated equations
panny
parents: 54106
diff changeset
  1011
            val lhs = list_comb (Free (fun_name, fun_T), map Bound (length fun_args - 1 downto 0));
54097
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1012
            val maybe_rhs' = if is_some maybe_rhs then maybe_rhs else
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1013
              let
54098
07a8145aaeba pass the right theorems to tactic
panny
parents: 54097
diff changeset
  1014
                fun prove_code_ctr {ctr, sels, ...} =
54097
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1015
                  if not (exists (equal ctr o fst) ctr_alist) then NONE else
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1016
                    let
54099
0b58c15ad284 compile -- fix typo introduced in 07a8145aaeba
panny
parents: 54098
diff changeset
  1017
                      val prems = find_first (equal ctr o #ctr) disc_eqns
0b58c15ad284 compile -- fix typo introduced in 07a8145aaeba
panny
parents: 54098
diff changeset
  1018
                        |> Option.map #prems |> the_default [];
54097
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1019
                      val t =
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1020
                        filter (equal ctr o #ctr) sel_eqns
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1021
                        |> fst o finds ((op =) o apsnd #sel) sels
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1022
                        |> map (snd #> (fn [x] => (List.rev (#fun_args x), #rhs_term x)) #-> abstract)
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1023
                        |> curry list_comb ctr;
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1024
                    in
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1025
                      SOME (prems, t)
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1026
                    end;
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1027
                val maybe_ctr_conds_argss = map prove_code_ctr ctr_specs;
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1028
              in
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1029
                if exists is_none maybe_ctr_conds_argss then NONE else
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1030
                  fold_rev (fn SOME (prems, u) => fn t => mk_If (s_conjs prems) u t)
54120
c2f18fd05414 use Code.abort instead of undefined in auto-generated equations
panny
parents: 54106
diff changeset
  1031
                    maybe_ctr_conds_argss
c2f18fd05414 use Code.abort instead of undefined in auto-generated equations
panny
parents: 54106
diff changeset
  1032
                    (Const (@{const_name Code.abort}, @{typ String.literal} -->
c2f18fd05414 use Code.abort instead of undefined in auto-generated equations
panny
parents: 54106
diff changeset
  1033
                        (@{typ unit} --> body_type fun_T) --> body_type fun_T) $
c2f18fd05414 use Code.abort instead of undefined in auto-generated equations
panny
parents: 54106
diff changeset
  1034
                      @{term "STR []"} $ (* FIXME *)
c2f18fd05414 use Code.abort instead of undefined in auto-generated equations
panny
parents: 54106
diff changeset
  1035
                      absdummy @{typ unit} (incr_boundvars 1 lhs))
54097
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1036
                  |> SOME
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1037
              end;
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1038
          in
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1039
            if is_none maybe_rhs' then [] else
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1040
              let
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1041
                val rhs = the maybe_rhs';
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1042
                val bound_Ts = List.rev (map fastype_of fun_args);
54172
9c276e656712 reintroduced parts of 07a8145aaeba that a22ded8a7f7d wrongly took out
blanchet
parents: 54170
diff changeset
  1043
                val (raw_rhs, ctr_thms) =
9c276e656712 reintroduced parts of 07a8145aaeba that a22ded8a7f7d wrongly took out
blanchet
parents: 54170
diff changeset
  1044
                  if is_some maybe_rhs then
9c276e656712 reintroduced parts of 07a8145aaeba that a22ded8a7f7d wrongly took out
blanchet
parents: 54170
diff changeset
  1045
                    let
9c276e656712 reintroduced parts of 07a8145aaeba that a22ded8a7f7d wrongly took out
blanchet
parents: 54170
diff changeset
  1046
                      val raw_rhs = expand_corec_code_rhs lthy has_call bound_Ts rhs;
9c276e656712 reintroduced parts of 07a8145aaeba that a22ded8a7f7d wrongly took out
blanchet
parents: 54170
diff changeset
  1047
                      val cond_ctrs =
9c276e656712 reintroduced parts of 07a8145aaeba that a22ded8a7f7d wrongly took out
blanchet
parents: 54170
diff changeset
  1048
                        fold_rev_corec_code_rhs lthy (K oo (cons oo pair)) bound_Ts raw_rhs [];
9c276e656712 reintroduced parts of 07a8145aaeba that a22ded8a7f7d wrongly took out
blanchet
parents: 54170
diff changeset
  1049
                      val ctr_thms = map (the o AList.lookup (op =) ctr_alist o snd) cond_ctrs;
9c276e656712 reintroduced parts of 07a8145aaeba that a22ded8a7f7d wrongly took out
blanchet
parents: 54170
diff changeset
  1050
                    in (raw_rhs, ctr_thms) end
9c276e656712 reintroduced parts of 07a8145aaeba that a22ded8a7f7d wrongly took out
blanchet
parents: 54170
diff changeset
  1051
                  else
9c276e656712 reintroduced parts of 07a8145aaeba that a22ded8a7f7d wrongly took out
blanchet
parents: 54170
diff changeset
  1052
                    (rhs, map snd ctr_alist);
9c276e656712 reintroduced parts of 07a8145aaeba that a22ded8a7f7d wrongly took out
blanchet
parents: 54170
diff changeset
  1053
54097
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1054
                val ms = map (Logic.count_prems o prop_of) ctr_thms;
54172
9c276e656712 reintroduced parts of 07a8145aaeba that a22ded8a7f7d wrongly took out
blanchet
parents: 54170
diff changeset
  1055
                val (raw_t, t) = (raw_rhs, rhs)
9c276e656712 reintroduced parts of 07a8145aaeba that a22ded8a7f7d wrongly took out
blanchet
parents: 54170
diff changeset
  1056
                  |> pairself
9c276e656712 reintroduced parts of 07a8145aaeba that a22ded8a7f7d wrongly took out
blanchet
parents: 54170
diff changeset
  1057
                    (curry HOLogic.mk_eq (list_comb (Free (fun_name, fun_T),
9c276e656712 reintroduced parts of 07a8145aaeba that a22ded8a7f7d wrongly took out
blanchet
parents: 54170
diff changeset
  1058
                      map Bound (length fun_args - 1 downto 0)))
9c276e656712 reintroduced parts of 07a8145aaeba that a22ded8a7f7d wrongly took out
blanchet
parents: 54170
diff changeset
  1059
                    #> HOLogic.mk_Trueprop
9c276e656712 reintroduced parts of 07a8145aaeba that a22ded8a7f7d wrongly took out
blanchet
parents: 54170
diff changeset
  1060
                    #> curry Logic.list_all (map dest_Free fun_args));
54098
07a8145aaeba pass the right theorems to tactic
panny
parents: 54097
diff changeset
  1061
                val (distincts, discIs, sel_splits, sel_split_asms) =
54172
9c276e656712 reintroduced parts of 07a8145aaeba that a22ded8a7f7d wrongly took out
blanchet
parents: 54170
diff changeset
  1062
                  case_thms_of_term lthy bound_Ts raw_rhs;
54098
07a8145aaeba pass the right theorems to tactic
panny
parents: 54097
diff changeset
  1063
54172
9c276e656712 reintroduced parts of 07a8145aaeba that a22ded8a7f7d wrongly took out
blanchet
parents: 54170
diff changeset
  1064
                val raw_code_thm = mk_primcorec_raw_code_of_ctr_tac lthy distincts discIs sel_splits
54097
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1065
                    sel_split_asms ms ctr_thms
54172
9c276e656712 reintroduced parts of 07a8145aaeba that a22ded8a7f7d wrongly took out
blanchet
parents: 54170
diff changeset
  1066
                  |> K |> Goal.prove lthy [] [] raw_t;
9c276e656712 reintroduced parts of 07a8145aaeba that a22ded8a7f7d wrongly took out
blanchet
parents: 54170
diff changeset
  1067
val _ = tracing ("raw code theorem: " ^ Syntax.string_of_term lthy (prop_of raw_code_thm));
9c276e656712 reintroduced parts of 07a8145aaeba that a22ded8a7f7d wrongly took out
blanchet
parents: 54170
diff changeset
  1068
              in
9c276e656712 reintroduced parts of 07a8145aaeba that a22ded8a7f7d wrongly took out
blanchet
parents: 54170
diff changeset
  1069
                mk_primcorec_code_of_raw_code_tac lthy sel_splits raw_code_thm
9c276e656712 reintroduced parts of 07a8145aaeba that a22ded8a7f7d wrongly took out
blanchet
parents: 54170
diff changeset
  1070
                |> K |> Goal.prove lthy [] [] t
9c276e656712 reintroduced parts of 07a8145aaeba that a22ded8a7f7d wrongly took out
blanchet
parents: 54170
diff changeset
  1071
|> tap (tracing o curry (op ^) "code theorem: " o Syntax.string_of_term lthy o prop_of)
9c276e656712 reintroduced parts of 07a8145aaeba that a22ded8a7f7d wrongly took out
blanchet
parents: 54170
diff changeset
  1072
                |> single
9c276e656712 reintroduced parts of 07a8145aaeba that a22ded8a7f7d wrongly took out
blanchet
parents: 54170
diff changeset
  1073
              end
54106
e5f853482006 keep temporary error handling in there until code equations are properly generated
blanchet
parents: 54102
diff changeset
  1074
handle ERROR s => (warning s; []) (*###*)
e5f853482006 keep temporary error handling in there until code equations are properly generated
blanchet
parents: 54102
diff changeset
  1075
           end;
54097
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1076
53791
0ddf045113c9 tuned code
blanchet
parents: 53753
diff changeset
  1077
        val disc_alists = map3 (maps oo prove_disc) corec_specs exclssss disc_eqnss;
0ddf045113c9 tuned code
blanchet
parents: 53753
diff changeset
  1078
        val sel_alists = map4 (map ooo prove_sel) corec_specs disc_eqnss exclssss sel_eqnss;
0ddf045113c9 tuned code
blanchet
parents: 53753
diff changeset
  1079
        val disc_thmss = map (map snd) disc_alists;
0ddf045113c9 tuned code
blanchet
parents: 53753
diff changeset
  1080
        val sel_thmss = map (map snd) sel_alists;
54097
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1081
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1082
        val ctr_alists = map5 (maps oooo prove_ctr) disc_alists sel_alists disc_eqnss sel_eqnss
53744
9db52ae90009 setting the stage for safe constructor simp rules
blanchet
parents: 53743
diff changeset
  1083
          (map #ctr_specs corec_specs);
54097
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1084
        val ctr_thmss = map (map snd) ctr_alists;
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1085
92c5bd3b342d prove user-supplied equations for ctr and code reductions, preserving "let"s, "case"s etc.;
panny
parents: 54074
diff changeset
  1086
        val code_thmss = map4 prove_code disc_eqnss sel_eqnss ctr_alists corec_specs;
53791
0ddf045113c9 tuned code
blanchet
parents: 53753
diff changeset
  1087
54030
732b53d9b720 don't register equations of the form 'f x = ...' as simp rules, even if they are safe (noncorecursive), because they unfold too aggresively concepts users are likely to want to stay folded
blanchet
parents: 54028
diff changeset
  1088
        val simp_thmss = map2 append disc_thmss sel_thmss
53795
dfa1108368ad generate "simps" from "primcorec"
blanchet
parents: 53794
diff changeset
  1089
53797
576f9637dc7a note coinduct theorems in "primcorec"
blanchet
parents: 53795
diff changeset
  1090
        val common_name = mk_common_name fun_names;
576f9637dc7a note coinduct theorems in "primcorec"
blanchet
parents: 53795
diff changeset
  1091
53791
0ddf045113c9 tuned code
blanchet
parents: 53753
diff changeset
  1092
        val notes =
53830
ed2eb7df2aac don't note more induction principles than there are functions + tuning
blanchet
parents: 53822
diff changeset
  1093
          [(coinductN, map (if n2m then single else K []) coinduct_thms, []),
54145
297d1c603999 make sure that registered code equations are actually equations
blanchet
parents: 54133
diff changeset
  1094
           (codeN, code_thmss, code_nitpicksimp_attrs),
53797
576f9637dc7a note coinduct theorems in "primcorec"
blanchet
parents: 53795
diff changeset
  1095
           (ctrN, ctr_thmss, []),
53791
0ddf045113c9 tuned code
blanchet
parents: 53753
diff changeset
  1096
           (discN, disc_thmss, simp_attrs),
53795
dfa1108368ad generate "simps" from "primcorec"
blanchet
parents: 53794
diff changeset
  1097
           (selN, sel_thmss, simp_attrs),
53797
576f9637dc7a note coinduct theorems in "primcorec"
blanchet
parents: 53795
diff changeset
  1098
           (simpsN, simp_thmss, []),
53830
ed2eb7df2aac don't note more induction principles than there are functions + tuning
blanchet
parents: 53822
diff changeset
  1099
           (strong_coinductN, map (if n2m then single else K []) strong_coinduct_thms, [])]
53791
0ddf045113c9 tuned code
blanchet
parents: 53753
diff changeset
  1100
          |> maps (fn (thmN, thmss, attrs) =>
0ddf045113c9 tuned code
blanchet
parents: 53753
diff changeset
  1101
            map2 (fn fun_name => fn thms =>
0ddf045113c9 tuned code
blanchet
parents: 53753
diff changeset
  1102
                ((Binding.qualify true fun_name (Binding.name thmN), attrs), [(thms, [])]))
53830
ed2eb7df2aac don't note more induction principles than there are functions + tuning
blanchet
parents: 53822
diff changeset
  1103
              fun_names (take actual_nn thmss))
53791
0ddf045113c9 tuned code
blanchet
parents: 53753
diff changeset
  1104
          |> filter_out (null o fst o hd o snd);
53797
576f9637dc7a note coinduct theorems in "primcorec"
blanchet
parents: 53795
diff changeset
  1105
576f9637dc7a note coinduct theorems in "primcorec"
blanchet
parents: 53795
diff changeset
  1106
        val common_notes =
53830
ed2eb7df2aac don't note more induction principles than there are functions + tuning
blanchet
parents: 53822
diff changeset
  1107
          [(coinductN, if n2m then [coinduct_thm] else [], []),
ed2eb7df2aac don't note more induction principles than there are functions + tuning
blanchet
parents: 53822
diff changeset
  1108
           (strong_coinductN, if n2m then [strong_coinduct_thm] else [], [])]
53797
576f9637dc7a note coinduct theorems in "primcorec"
blanchet
parents: 53795
diff changeset
  1109
          |> filter_out (null o #2)
576f9637dc7a note coinduct theorems in "primcorec"
blanchet
parents: 53795
diff changeset
  1110
          |> map (fn (thmN, thms, attrs) =>
576f9637dc7a note coinduct theorems in "primcorec"
blanchet
parents: 53795
diff changeset
  1111
            ((Binding.qualify true common_name (Binding.name thmN), attrs), [(thms, [])]));
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
  1112
      in
54030
732b53d9b720 don't register equations of the form 'f x = ...' as simp rules, even if they are safe (noncorecursive), because they unfold too aggresively concepts users are likely to want to stay folded
blanchet
parents: 54028
diff changeset
  1113
        lthy |> Local_Theory.notes (notes @ common_notes) |> snd
53654
8b9ea4420f81 prove simp theorems for newly generated definitions
panny
parents: 53411
diff changeset
  1114
      end;
53822
6304b12c7627 add "primcorec" command (cf. ae7f50e70c09)
panny
parents: 53811
diff changeset
  1115
6304b12c7627 add "primcorec" command (cf. ae7f50e70c09)
panny
parents: 53811
diff changeset
  1116
    fun after_qed thmss' = fold_map Local_Theory.define defs #-> prove thmss';
6304b12c7627 add "primcorec" command (cf. ae7f50e70c09)
panny
parents: 53811
diff changeset
  1117
6304b12c7627 add "primcorec" command (cf. ae7f50e70c09)
panny
parents: 53811
diff changeset
  1118
    val _ = if not simple orelse forall null obligationss then () else
6304b12c7627 add "primcorec" command (cf. ae7f50e70c09)
panny
parents: 53811
diff changeset
  1119
      primrec_error "need exclusiveness proofs - use primcorecursive instead of primcorec";
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
  1120
  in
53822
6304b12c7627 add "primcorec" command (cf. ae7f50e70c09)
panny
parents: 53811
diff changeset
  1121
    if simple then
6304b12c7627 add "primcorec" command (cf. ae7f50e70c09)
panny
parents: 53811
diff changeset
  1122
      lthy'
6304b12c7627 add "primcorec" command (cf. ae7f50e70c09)
panny
parents: 53811
diff changeset
  1123
      |> after_qed (map (fn [] => []) obligationss)
6304b12c7627 add "primcorec" command (cf. ae7f50e70c09)
panny
parents: 53811
diff changeset
  1124
      |> pair NONE o SOME
6304b12c7627 add "primcorec" command (cf. ae7f50e70c09)
panny
parents: 53811
diff changeset
  1125
    else
6304b12c7627 add "primcorec" command (cf. ae7f50e70c09)
panny
parents: 53811
diff changeset
  1126
      lthy'
6304b12c7627 add "primcorec" command (cf. ae7f50e70c09)
panny
parents: 53811
diff changeset
  1127
      |> Proof.theorem NONE after_qed obligationss
6304b12c7627 add "primcorec" command (cf. ae7f50e70c09)
panny
parents: 53811
diff changeset
  1128
      |> Proof.refine (Method.primitive_text I)
6304b12c7627 add "primcorec" command (cf. ae7f50e70c09)
panny
parents: 53811
diff changeset
  1129
      |> Seq.hd
6304b12c7627 add "primcorec" command (cf. ae7f50e70c09)
panny
parents: 53811
diff changeset
  1130
      |> rpair NONE o SOME
6304b12c7627 add "primcorec" command (cf. ae7f50e70c09)
panny
parents: 53811
diff changeset
  1131
  end;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
  1132
53831
80423b9080cf support "of" syntax to disambiguate selector equations
panny
parents: 53830
diff changeset
  1133
fun add_primcorec_ursive_cmd simple seq (raw_fixes, raw_specs') lthy =
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
  1134
  let
53831
80423b9080cf support "of" syntax to disambiguate selector equations
panny
parents: 53830
diff changeset
  1135
    val (raw_specs, of_specs) = split_list raw_specs' ||> map (Option.map (Syntax.read_term lthy));
80423b9080cf support "of" syntax to disambiguate selector equations
panny
parents: 53830
diff changeset
  1136
    val ((fixes, specs), _) = Specification.read_spec raw_fixes raw_specs lthy;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
  1137
  in
53831
80423b9080cf support "of" syntax to disambiguate selector equations
panny
parents: 53830
diff changeset
  1138
    add_primcorec simple seq fixes specs of_specs lthy
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
  1139
    handle ERROR str => primrec_error str
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
  1140
  end
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
  1141
  handle Primrec_Error (str, eqns) =>
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
  1142
    if null eqns
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
  1143
    then error ("primcorec error:\n  " ^ str)
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
  1144
    else error ("primcorec error:\n  " ^ str ^ "\nin\n  " ^
53822
6304b12c7627 add "primcorec" command (cf. ae7f50e70c09)
panny
parents: 53811
diff changeset
  1145
      space_implode "\n  " (map (quote o Syntax.string_of_term lthy) eqns));
6304b12c7627 add "primcorec" command (cf. ae7f50e70c09)
panny
parents: 53811
diff changeset
  1146
6304b12c7627 add "primcorec" command (cf. ae7f50e70c09)
panny
parents: 53811
diff changeset
  1147
val add_primcorecursive_cmd = (the o fst) ooo add_primcorec_ursive_cmd false;
6304b12c7627 add "primcorec" command (cf. ae7f50e70c09)
panny
parents: 53811
diff changeset
  1148
val add_primcorec_cmd = (the o snd) ooo add_primcorec_ursive_cmd true;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
  1149
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
  1150
end;