more work on FP sugar
authorblanchet
Tue, 04 Sep 2012 13:02:32 +0200
changeset 49119 1f605c36869c
parent 49118 b815fa776b91
child 49120 7f8e69fc6ac9
more work on FP sugar
src/HOL/Codatatype/Tools/bnf_fp_sugar.ML
src/HOL/Codatatype/Tools/bnf_fp_util.ML
src/HOL/Codatatype/Tools/bnf_gfp_util.ML
src/HOL/Codatatype/Tools/bnf_util.ML
src/HOL/Codatatype/Tools/bnf_wrap.ML
--- a/src/HOL/Codatatype/Tools/bnf_fp_sugar.ML	Tue Sep 04 13:02:31 2012 +0200
+++ b/src/HOL/Codatatype/Tools/bnf_fp_sugar.ML	Tue Sep 04 13:02:32 2012 +0200
@@ -12,28 +12,115 @@
 structure BNF_FP_Sugar : BNF_FP_SUGAR =
 struct
 
-fun data_cmd gfp specs lthy =
+open BNF_Util
+open BNF_Wrap
+open BNF_FP_Util
+open BNF_LFP
+open BNF_GFP
+
+fun cannot_merge_types () = error "Mutually recursive (co)datatypes must have same type parameters";
+
+fun merge_type_arg_constrained ctxt (T, c) (T', c') =
+  if T = T' then
+    (case (c, c') of
+      (_, NONE) => (T, c)
+    | (NONE, _) => (T, c')
+    | _ =>
+      if c = c' then
+        (T, c)
+      else
+        error ("Inconsistent sort constraints for type variable " ^
+          quote (Syntax.string_of_typ ctxt T)))
+  else
+    cannot_merge_types ();
+
+fun merge_type_args_constrained ctxt (cAs, cAs') =
+  if length cAs = length cAs' then map2 (merge_type_arg_constrained ctxt) cAs cAs'
+  else cannot_merge_types ();
+
+fun type_args_constrained_of_spec (((cAs, _), _), _) = cAs;
+fun type_name_of_spec (((_, b), _), _) = b;
+fun mixfix_of_spec ((_, mx), _) = mx;
+fun ctr_specs_of_spec (_, ctr_specs) = ctr_specs;
+
+fun disc_of_ctr_spec (((disc, _), _), _) = disc;
+fun ctr_of_ctr_spec (((_, ctr), _), _) = ctr;
+fun args_of_ctr_spec ((_, args), _) = args;
+fun mixfix_of_ctr_spec (_, mx) = mx;
+
+val mk_prod_sum = mk_sumTN o map HOLogic.mk_tupleT;
+
+val lfp_info = bnf_lfp;
+val gfp_info = bnf_gfp;
+
+fun prepare_data prepare_typ construct specs lthy =
   let
+    val constrained_passiveAs =
+      map (map (apfst (prepare_typ lthy)) o type_args_constrained_of_spec) specs
+      |> Library.foldr1 (merge_type_args_constrained lthy);
+    val passiveAs = map fst constrained_passiveAs;
+
+    val _ = (case duplicates (op =) passiveAs of [] => ()
+      | T :: _ => error ("Duplicate type parameter " ^ quote (Syntax.string_of_typ lthy T)));
+
+    (* TODO: check that no type variables occur in the rhss that's not in the lhss *)
+    (* TODO: use sort constraints on type args *)
+
+    val N = length specs;
+
+    val bs = map type_name_of_spec specs;
+    val mixfixes = map mixfix_of_spec specs;
+
+    val _ = (case duplicates Binding.eq_name bs of [] => ()
+      | b :: _ => error ("Duplicate type name declaration " ^ quote (Binding.name_of b)));
+
+    val ctr_specss = map ctr_specs_of_spec specs;
+
+    val disc_namess = map (map disc_of_ctr_spec) ctr_specss;
+    val raw_ctr_namess = map (map ctr_of_ctr_spec) ctr_specss;
+    val ctr_argsss = map (map args_of_ctr_spec) ctr_specss;
+    val ctr_mixfixess = map (map mixfix_of_ctr_spec) ctr_specss;
+
+    val sel_namesss = map (map (map fst)) ctr_argsss;
+    val ctr_Tsss = map (map (map (prepare_typ lthy o snd))) ctr_argsss;
+
+    val (activeAs, _) = lthy |> mk_TFrees N;
+
+    val eqs = map2 (fn TFree A => fn Tss => (A, mk_prod_sum Tss)) activeAs ctr_Tsss;
+
+    val lthy' = fp_bnf construct bs eqs lthy;
+
+    fun wrap_type ((b, disc_names), sel_namess) lthy =
+      let
+        val ctrs = [];
+        val caseof = @{term True};
+        val tacss = [];
+      in
+        wrap tacss ((ctrs, caseof), (disc_names, sel_namess)) lthy
+      end;
   in
-    lthy
+    lthy' |> fold wrap_type (bs ~~ disc_namess ~~ sel_namesss)
   end;
 
+val data_cmd = prepare_data Syntax.read_typ;
+
+val parse_opt_binding_colon = Scan.optional (Parse.binding --| Parse.$$$ ":") no_name
+
 val parse_ctr_arg =
-  Parse.$$$ "(" |-- Scan.option Parse.binding --| Parse.$$$ ":" -- Parse.typ --| Parse.$$$ ")" ||
-  (Parse.typ >> pair NONE);
+  Parse.$$$ "(" |-- parse_opt_binding_colon -- Parse.typ --| Parse.$$$ ")" ||
+  (Parse.typ >> pair no_name);
 
 val parse_single_spec =
   Parse.type_args_constrained -- Parse.binding -- Parse.opt_mixfix --
-  (@{keyword "="} |-- Parse.enum1 "|" (Parse.binding -- Scan.repeat parse_ctr_arg --
-     Parse.opt_mixfix))
-  >> (fn (((vs, t), mx), cons) => ((t, vs, mx), map Parse.triple1 cons));
+  (@{keyword "="} |-- Parse.enum1 "|" (parse_opt_binding_colon -- Parse.binding --
+    Scan.repeat parse_ctr_arg -- Parse.opt_mixfix));
 
 val _ =
   Outer_Syntax.local_theory @{command_spec "data"} "define BNF-based inductive datatypes"
-    (Parse.and_list1 parse_single_spec >> data_cmd false);
+    (Parse.and_list1 parse_single_spec >> data_cmd lfp_info);
 
 val _ =
   Outer_Syntax.local_theory @{command_spec "codata"} "define BNF-based coinductive datatypes"
-    (Parse.and_list1 parse_single_spec >> data_cmd true);
+    (Parse.and_list1 parse_single_spec >> data_cmd gfp_info);
 
 end;
--- a/src/HOL/Codatatype/Tools/bnf_fp_util.ML	Tue Sep 04 13:02:31 2012 +0200
+++ b/src/HOL/Codatatype/Tools/bnf_fp_util.ML	Tue Sep 04 13:02:32 2012 +0200
@@ -186,7 +186,7 @@
 
 (* stolen from "~~/src/HOL/Tools/Datatype/datatype_aux.ML" *)
 fun split_conj_thm th =
-  ((th RS conjunct1)::(split_conj_thm (th RS conjunct2))) handle THM _ => [th];
+  ((th RS conjunct1) :: split_conj_thm (th RS conjunct2)) handle THM _ => [th];
 
 fun split_conj_prems limit th =
   let
--- a/src/HOL/Codatatype/Tools/bnf_gfp_util.ML	Tue Sep 04 13:02:31 2012 +0200
+++ b/src/HOL/Codatatype/Tools/bnf_gfp_util.ML	Tue Sep 04 13:02:32 2012 +0200
@@ -43,7 +43,6 @@
 
   val mk_InN: typ list -> term -> int -> term
 
-  val mk_sumTN: typ list -> typ
   val mk_sum_case: term -> term -> term
   val mk_sum_caseN: term list -> term
 
@@ -192,8 +191,6 @@
       A $ f1 $ f2 $ b1 $ b2
   end;
 
-fun mk_sumTN Ts = Library.foldr1 (mk_sumT) Ts;
-
 fun Inl_const LT RT = Const (@{const_name Inl}, LT --> mk_sumT (LT, RT));
 fun mk_Inl t RT = Inl_const (fastype_of t) RT $ t;
 
--- a/src/HOL/Codatatype/Tools/bnf_util.ML	Tue Sep 04 13:02:31 2012 +0200
+++ b/src/HOL/Codatatype/Tools/bnf_util.ML	Tue Sep 04 13:02:32 2012 +0200
@@ -57,6 +57,7 @@
   val mk_relT: typ * typ -> typ
   val dest_relT: typ -> typ * typ
   val mk_sumT: typ * typ -> typ
+  val mk_sumTN: typ list -> typ
 
   val ctwo: term
   val fst_const: typ -> term
@@ -288,6 +289,7 @@
 val mk_relT = HOLogic.mk_setT o HOLogic.mk_prodT;
 val dest_relT = HOLogic.dest_prodT o HOLogic.dest_setT;
 fun mk_sumT (LT, RT) = Type (@{type_name Sum_Type.sum}, [LT, RT]);
+fun mk_sumTN Ts = Library.foldr1 mk_sumT Ts;
 fun mk_partial_funT (ranT, domT) = domT --> mk_optionT ranT;
 
 
--- a/src/HOL/Codatatype/Tools/bnf_wrap.ML	Tue Sep 04 13:02:31 2012 +0200
+++ b/src/HOL/Codatatype/Tools/bnf_wrap.ML	Tue Sep 04 13:02:32 2012 +0200
@@ -7,8 +7,9 @@
 
 signature BNF_WRAP =
 sig
+  val no_name: binding
   val wrap: ({prems: thm list, context: Proof.context} -> tactic) list list ->
-    (term list * term) * (binding list * binding list list) -> Proof.context -> local_theory
+    (term list * term) * (binding list * binding list list) -> local_theory -> local_theory
 end;
 
 structure BNF_Wrap : BNF_WRAP =
@@ -39,7 +40,7 @@
 val weak_case_cong_thmsN = "weak_case_cong";
 
 val no_name = @{binding "*"};
-val default_name = @{binding _};
+val fallback_name = @{binding _};
 
 fun pad_list x n xs = xs @ replicate (n - length xs) x;
 
@@ -95,11 +96,11 @@
     val ms = map length ctr_Tss;
 
     val disc_names =
-      pad_list default_name n raw_disc_names
+      pad_list fallback_name n raw_disc_names
       |> map2 (fn ctr => fn disc =>
         if Binding.eq_name (disc, no_name) then
           NONE
-        else if Binding.eq_name (disc, default_name) then
+        else if Binding.eq_name (disc, fallback_name) then
           SOME (Binding.name (prefix is_N (Long_Name.base_name (name_of_ctr ctr))))
         else
           SOME disc) ctrs0;
@@ -109,10 +110,10 @@
     val sel_namess =
       pad_list [] n raw_sel_namess
       |> map3 (fn ctr => fn m => map2 (fn l => fn sel =>
-        if Binding.eq_name (sel, default_name) then
+        if Binding.eq_name (sel, fallback_name) then
           Binding.name (mk_un_N m l (Long_Name.base_name (name_of_ctr ctr)))
         else
-          sel) (1 upto m) o pad_list default_name m) ctrs0 ms;
+          sel) (1 upto m) o pad_list fallback_name m) ctrs0 ms;
 
     fun mk_caseof Ts T =
       let val (binders, body) = strip_type (fastype_of caseof0) in