src/HOL/BNF/Tools/bnf_fp_rec_sugar.ML
changeset 53720 03fac7082137
parent 53699 7d32f33d340d
child 53722 e176d6d3345f
equal deleted inserted replaced
53719:edbd6bc472b4 53720:03fac7082137
    33 val free_name = try (fn Free (v, _) => v);
    33 val free_name = try (fn Free (v, _) => v);
    34 val const_name = try (fn Const (v, _) => v);
    34 val const_name = try (fn Const (v, _) => v);
    35 val undef_const = Const (@{const_name undefined}, dummyT);
    35 val undef_const = Const (@{const_name undefined}, dummyT);
    36 
    36 
    37 fun permute_args n t = list_comb (t, map Bound (0 :: (n downto 1)))
    37 fun permute_args n t = list_comb (t, map Bound (0 :: (n downto 1)))
    38   |> fold (K (fn u => Abs (Name.uu, dummyT, u))) (0 upto n);
    38   |> fold (K (Term.abs (Name.uu, dummyT))) (0 upto n);
    39 val abs_tuple = HOLogic.tupled_lambda o HOLogic.mk_tuple;
    39 val abs_tuple = HOLogic.tupled_lambda o HOLogic.mk_tuple;
    40 fun drop_All t = subst_bounds (strip_qnt_vars @{const_name all} t |> map Free |> rev,
    40 fun drop_All t = subst_bounds (strip_qnt_vars @{const_name all} t |> map Free |> rev,
    41   strip_qnt_body @{const_name all} t)
    41   strip_qnt_body @{const_name all} t)
    42 fun mk_not @{const True} = @{const False}
    42 fun mk_not @{const True} = @{const False}
    43   | mk_not @{const False} = @{const True}
    43   | mk_not @{const False} = @{const True}
    44   | mk_not (@{const Not} $ t) = t
    44   | mk_not (@{const Not} $ t) = t
    45   | mk_not (@{const Trueprop} $ t) = @{const Trueprop} $ mk_not t
    45   | mk_not (@{const Trueprop} $ t) = @{const Trueprop} $ mk_not t
    46   | mk_not t = HOLogic.mk_not t
    46   | mk_not t = HOLogic.mk_not t
    47 val mk_conjs = try (foldr1 HOLogic.mk_conj) #> the_default @{const True};
    47 val mk_conjs = try (foldr1 HOLogic.mk_conj) #> the_default @{const True};
    48 val mk_disjs = try (foldr1 HOLogic.mk_disj) #> the_default @{const False};
    48 val mk_disjs = try (foldr1 HOLogic.mk_disj) #> the_default @{const False};
    49 
       
    50 fun invert_prems [t] = map mk_not (HOLogic.disjuncts t)
    49 fun invert_prems [t] = map mk_not (HOLogic.disjuncts t)
    51   | invert_prems ts = [mk_disjs (map mk_not ts)];
    50   | invert_prems ts = [mk_disjs (map mk_not ts)];
       
    51 fun invert_prems_disj [t] = map mk_not (HOLogic.disjuncts t)
       
    52   | invert_prems_disj ts = [mk_disjs (map (mk_conjs o map mk_not o HOLogic.disjuncts) ts)];
       
    53 fun abstract vs =
       
    54   let fun a n (t $ u) = a n t $ a n u
       
    55         | a n (Abs (v, T, b)) = Abs (v, T, a (n + 1) b)
       
    56         | a n t = let val idx = find_index (equal t) vs in
       
    57             if idx < 0 then t else Bound (n + idx) end
       
    58   in a 0 end;
    52 
    59 
    53 val simp_attrs = @{attributes [simp]};
    60 val simp_attrs = @{attributes [simp]};
    54 
       
    55 fun abstract n vs (t $ u) = abstract n vs t $ abstract n vs u
       
    56   | abstract n vs (Abs (v, T, b)) = Abs (v, T, abstract (n + 1) vs b)
       
    57   | abstract n vs t = let val idx = find_index (equal t) vs in
       
    58       if idx < 0 then t else Bound (n + idx) end;
       
    59 
    61 
    60 
    62 
    61 (* Primrec *)
    63 (* Primrec *)
    62 
    64 
    63 type eqn_data = {
    65 type eqn_data = {
   381 
   383 
   382 (* Primcorec *)
   384 (* Primcorec *)
   383 
   385 
   384 type co_eqn_data_disc = {
   386 type co_eqn_data_disc = {
   385   fun_name: string,
   387   fun_name: string,
       
   388   fun_T: typ,
   386   fun_args: term list,
   389   fun_args: term list,
       
   390   ctr: term,
   387   ctr_no: int, (*###*)
   391   ctr_no: int, (*###*)
       
   392   disc: term,
   388   prems: term list,
   393   prems: term list,
   389   user_eqn: term
   394   user_eqn: term
   390 };
   395 };
   391 type co_eqn_data_sel = {
   396 type co_eqn_data_sel = {
   392   fun_name: string,
   397   fun_name: string,
       
   398   fun_T: typ,
   393   fun_args: term list,
   399   fun_args: term list,
   394   ctr: term,
   400   ctr: term,
   395   sel: term,
   401   sel: term,
   396   rhs_term: term,
   402   rhs_term: term,
   397   user_eqn: term
   403   user_eqn: term
   398 };
   404 };
   399 datatype co_eqn_data =
   405 datatype co_eqn_data =
   400   Disc of co_eqn_data_disc |
   406   Disc of co_eqn_data_disc |
   401   Sel of co_eqn_data_sel;
   407   Sel of co_eqn_data_sel;
   402 
   408 
   403 fun co_dissect_eqn_disc sequential fun_names corec_specs prems' concl matchedss =
   409 fun co_dissect_eqn_disc sequential fun_names corec_specs prems' concl matchedsss =
   404   let
   410   let
   405     fun find_subterm p = let (* FIXME \<exists>? *)
   411     fun find_subterm p = let (* FIXME \<exists>? *)
   406       fun f (t as u $ v) = if p t then SOME t else merge_options (f u, f v)
   412       fun f (t as u $ v) = if p t then SOME t else merge_options (f u, f v)
   407         | f t = if p t then SOME t else NONE
   413         | f t = if p t then SOME t else NONE
   408       in f end;
   414       in f end;
   409 
   415 
   410     val applied_fun = concl
   416     val applied_fun = concl
   411       |> find_subterm (member ((op =) o apsnd SOME) fun_names o try (fst o dest_Free o head_of))
   417       |> find_subterm (member ((op =) o apsnd SOME) fun_names o try (fst o dest_Free o head_of))
   412       |> the
   418       |> the
   413       handle Option.Option => primrec_error_eqn "malformed discriminator equation" concl;
   419       handle Option.Option => primrec_error_eqn "malformed discriminator equation" concl;
   414     val (fun_name, fun_args) = strip_comb applied_fun |>> fst o dest_Free;
   420     val ((fun_name, fun_T), fun_args) = strip_comb applied_fun |>> dest_Free;
   415     val corec_spec = the (AList.lookup (op =) (fun_names ~~ corec_specs) fun_name);
   421     val {ctr_specs, ...} = the (AList.lookup (op =) (fun_names ~~ corec_specs) fun_name);
   416 
   422 
   417     val discs = #ctr_specs corec_spec |> map #disc;
   423     val discs = map #disc ctr_specs;
   418     val ctrs = #ctr_specs corec_spec |> map #ctr;
   424     val ctrs = map #ctr ctr_specs;
   419     val not_disc = head_of concl = @{term Not};
   425     val not_disc = head_of concl = @{term Not};
   420     val _ = not_disc andalso length ctrs <> 2 andalso
   426     val _ = not_disc andalso length ctrs <> 2 andalso
   421       primrec_error_eqn "\<not>ed discriminator for a type with \<noteq> 2 constructors" concl;
   427       primrec_error_eqn "\<not>ed discriminator for a type with \<noteq> 2 constructors" concl;
   422     val disc = find_subterm (member (op =) discs o head_of) concl;
   428     val disc = find_subterm (member (op =) discs o head_of) concl;
   423     val eq_ctr0 = concl |> perhaps (try (HOLogic.dest_not)) |> try (HOLogic.dest_eq #> snd)
   429     val eq_ctr0 = concl |> perhaps (try (HOLogic.dest_not)) |> try (HOLogic.dest_eq #> snd)
   426     val _ = is_some disc orelse is_some eq_ctr0 orelse
   432     val _ = is_some disc orelse is_some eq_ctr0 orelse
   427       primrec_error_eqn "no discriminator in equation" concl;
   433       primrec_error_eqn "no discriminator in equation" concl;
   428     val ctr_no' =
   434     val ctr_no' =
   429       if is_none disc then the eq_ctr0 else find_index (equal (head_of (the disc))) discs;
   435       if is_none disc then the eq_ctr0 else find_index (equal (head_of (the disc))) discs;
   430     val ctr_no = if not_disc then 1 - ctr_no' else ctr_no';
   436     val ctr_no = if not_disc then 1 - ctr_no' else ctr_no';
       
   437     val ctr = #ctr (nth ctr_specs ctr_no);
   431 
   438 
   432     val catch_all = try (fst o dest_Free o the_single) prems' = SOME Name.uu_;
   439     val catch_all = try (fst o dest_Free o the_single) prems' = SOME Name.uu_;
   433     val matcheds = AList.lookup (op =) matchedss fun_name |> the_default [];
   440     val matchedss = AList.lookup (op =) matchedsss fun_name |> the_default [];
   434     val prems = map (abstract 0 (List.rev fun_args)) prems';
   441     val prems = map (abstract (List.rev fun_args)) prems';
   435     val real_prems = (if catch_all orelse sequential then invert_prems matcheds else []) @
   442     val real_prems =
       
   443       (if catch_all orelse sequential then maps invert_prems_disj matchedss else []) @
   436       (if catch_all then [] else prems);
   444       (if catch_all then [] else prems);
   437 
   445 
   438     val matchedss' = AList.delete (op =) fun_name matchedss
   446     val matchedsss' = AList.delete (op =) fun_name matchedsss
   439       |> cons (fun_name, if sequential then prems @ matcheds else real_prems @ matcheds);
   447       |> cons (fun_name, if sequential then matchedss @ [prems] else matchedss @ [real_prems]);
   440 
   448 
   441     val user_eqn =
   449     val user_eqn =
   442       (real_prems, betapply (#disc (nth (#ctr_specs corec_spec) ctr_no), applied_fun))
   450       (real_prems, betapply (#disc (nth ctr_specs ctr_no), applied_fun))
   443       |>> map HOLogic.mk_Trueprop ||> HOLogic.mk_Trueprop
   451       |>> map HOLogic.mk_Trueprop ||> HOLogic.mk_Trueprop
   444       |> Logic.list_implies;
   452       |> Logic.list_implies;
   445   in
   453   in
   446     (Disc {
   454     (Disc {
   447       fun_name = fun_name,
   455       fun_name = fun_name,
       
   456       fun_T = fun_T,
   448       fun_args = fun_args,
   457       fun_args = fun_args,
       
   458       ctr = ctr,
   449       ctr_no = ctr_no,
   459       ctr_no = ctr_no,
       
   460       disc = #disc (nth ctr_specs ctr_no),
   450       prems = real_prems,
   461       prems = real_prems,
   451       user_eqn = user_eqn
   462       user_eqn = user_eqn
   452     }, matchedss')
   463     }, matchedsss')
   453   end;
   464   end;
   454 
   465 
   455 fun co_dissect_eqn_sel fun_names corec_specs eqn' eqn =
   466 fun co_dissect_eqn_sel fun_names corec_specs eqn' eqn =
   456   let
   467   let
   457     val (lhs, rhs) = HOLogic.dest_eq eqn
   468     val (lhs, rhs) = HOLogic.dest_eq eqn
   458       handle TERM _ =>
   469       handle TERM _ =>
   459         primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn;
   470         primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn;
   460     val sel = head_of lhs;
   471     val sel = head_of lhs;
   461     val (fun_name, fun_args) = dest_comb lhs |> snd |> strip_comb |> apfst (fst o dest_Free)
   472     val ((fun_name, fun_T), fun_args) = dest_comb lhs |> snd |> strip_comb |> apfst dest_Free
   462       handle TERM _ =>
   473       handle TERM _ =>
   463         primrec_error_eqn "malformed selector argument in left-hand side" eqn;
   474         primrec_error_eqn "malformed selector argument in left-hand side" eqn;
   464     val corec_spec = the (AList.lookup (op =) (fun_names ~~ corec_specs) fun_name)
   475     val corec_spec = the (AList.lookup (op =) (fun_names ~~ corec_specs) fun_name)
   465       handle Option.Option => primrec_error_eqn "malformed selector argument in left-hand side" eqn;
   476       handle Option.Option => primrec_error_eqn "malformed selector argument in left-hand side" eqn;
   466     val (ctr_spec, sel) = #ctr_specs corec_spec
   477     val (ctr_spec, sel) = #ctr_specs corec_spec
   468       |>> nth (#ctr_specs corec_spec);
   479       |>> nth (#ctr_specs corec_spec);
   469     val user_eqn = drop_All eqn';
   480     val user_eqn = drop_All eqn';
   470   in
   481   in
   471     Sel {
   482     Sel {
   472       fun_name = fun_name,
   483       fun_name = fun_name,
       
   484       fun_T = fun_T,
   473       fun_args = fun_args,
   485       fun_args = fun_args,
   474       ctr = #ctr ctr_spec,
   486       ctr = #ctr ctr_spec,
   475       sel = sel,
   487       sel = sel,
   476       rhs_term = rhs,
   488       rhs_term = rhs,
   477       user_eqn = user_eqn
   489       user_eqn = user_eqn
   478     }
   490     }
   479   end;
   491   end;
   480 
   492 
   481 fun co_dissect_eqn_ctr sequential fun_names corec_specs eqn' imp_prems imp_rhs matchedss =
   493 fun co_dissect_eqn_ctr sequential fun_names corec_specs eqn' imp_prems imp_rhs matchedsss =
   482   let 
   494   let 
   483     val (lhs, rhs) = HOLogic.dest_eq imp_rhs;
   495     val (lhs, rhs) = HOLogic.dest_eq imp_rhs;
   484     val fun_name = head_of lhs |> fst o dest_Free;
   496     val fun_name = head_of lhs |> fst o dest_Free;
   485     val corec_spec = the (AList.lookup (op =) (fun_names ~~ corec_specs) fun_name);
   497     val {ctr_specs, ...} = the (AList.lookup (op =) (fun_names ~~ corec_specs) fun_name);
   486     val (ctr, ctr_args) = strip_comb rhs;
   498     val (ctr, ctr_args) = strip_comb rhs;
   487     val ctr_spec = the (find_first (equal ctr o #ctr) (#ctr_specs corec_spec))
   499     val {disc, sels, ...} = the (find_first (equal ctr o #ctr) ctr_specs)
   488       handle Option.Option => primrec_error_eqn "not a constructor" ctr;
   500       handle Option.Option => primrec_error_eqn "not a constructor" ctr;
   489 
   501 
   490     val disc_imp_rhs = betapply (#disc ctr_spec, lhs);
   502     val disc_imp_rhs = betapply (disc, lhs);
   491     val (maybe_eqn_data_disc, matchedss') = if length (#ctr_specs corec_spec) = 1
   503     val (maybe_eqn_data_disc, matchedsss') = if length ctr_specs = 1
   492       then (NONE, matchedss)
   504       then (NONE, matchedsss)
   493       else apfst SOME (co_dissect_eqn_disc
   505       else apfst SOME (co_dissect_eqn_disc
   494           sequential fun_names corec_specs imp_prems disc_imp_rhs matchedss);
   506           sequential fun_names corec_specs imp_prems disc_imp_rhs matchedsss);
   495 
   507 
   496     val sel_imp_rhss = (#sels ctr_spec ~~ ctr_args)
   508     val sel_imp_rhss = (sels ~~ ctr_args)
   497       |> map (fn (sel, ctr_arg) => HOLogic.mk_eq (betapply (sel, lhs), ctr_arg));
   509       |> map (fn (sel, ctr_arg) => HOLogic.mk_eq (betapply (sel, lhs), ctr_arg));
   498 
   510 
   499 val _ = tracing ("reduced\n    " ^ Syntax.string_of_term @{context} imp_rhs ^ "\nto\n    \<cdot> " ^
   511 val _ = tracing ("reduced\n    " ^ Syntax.string_of_term @{context} imp_rhs ^ "\nto\n    \<cdot> " ^
   500  (is_some maybe_eqn_data_disc ? K (Syntax.string_of_term @{context} disc_imp_rhs ^ "\n    \<cdot> ")) "" ^
   512  (is_some maybe_eqn_data_disc ? K (Syntax.string_of_term @{context} disc_imp_rhs ^ "\n    \<cdot> ")) "" ^
   501  space_implode "\n    \<cdot> " (map (Syntax.string_of_term @{context}) sel_imp_rhss));
   513  space_implode "\n    \<cdot> " (map (Syntax.string_of_term @{context}) sel_imp_rhss));
   502 
   514 
   503     val eqns_data_sel =
   515     val eqns_data_sel =
   504       map (co_dissect_eqn_sel fun_names corec_specs eqn') sel_imp_rhss;
   516       map (co_dissect_eqn_sel fun_names corec_specs eqn') sel_imp_rhss;
   505   in
   517   in
   506     (the_list maybe_eqn_data_disc @ eqns_data_sel, matchedss')
   518     (the_list maybe_eqn_data_disc @ eqns_data_sel, matchedsss')
   507   end;
   519   end;
   508 
   520 
   509 fun co_dissect_eqn sequential fun_names corec_specs eqn' matchedss =
   521 fun co_dissect_eqn sequential fun_names corec_specs eqn' matchedsss =
   510   let
   522   let
   511     val eqn = drop_All eqn'
   523     val eqn = drop_All eqn'
   512       handle TERM _ => primrec_error_eqn "malformed function equation" eqn';
   524       handle TERM _ => primrec_error_eqn "malformed function equation" eqn';
   513     val (imp_prems, imp_rhs) = Logic.strip_horn eqn
   525     val (imp_prems, imp_rhs) = Logic.strip_horn eqn
   514       |> apfst (map HOLogic.dest_Trueprop) o apsnd HOLogic.dest_Trueprop;
   526       |> apfst (map HOLogic.dest_Trueprop) o apsnd HOLogic.dest_Trueprop;
   524     val ctrs = maps #ctr_specs corec_specs |> map #ctr;
   536     val ctrs = maps #ctr_specs corec_specs |> map #ctr;
   525   in
   537   in
   526     if member (op =) discs head orelse
   538     if member (op =) discs head orelse
   527       is_some maybe_rhs andalso
   539       is_some maybe_rhs andalso
   528         member (op =) (filter (null o binder_types o fastype_of) ctrs) (the maybe_rhs) then
   540         member (op =) (filter (null o binder_types o fastype_of) ctrs) (the maybe_rhs) then
   529       co_dissect_eqn_disc sequential fun_names corec_specs imp_prems imp_rhs matchedss
   541       co_dissect_eqn_disc sequential fun_names corec_specs imp_prems imp_rhs matchedsss
   530       |>> single
   542       |>> single
   531     else if member (op =) sels head then
   543     else if member (op =) sels head then
   532       ([co_dissect_eqn_sel fun_names corec_specs eqn' imp_rhs], matchedss)
   544       ([co_dissect_eqn_sel fun_names corec_specs eqn' imp_rhs], matchedsss)
   533     else if is_Free head andalso member (op =) fun_names (fst (dest_Free head)) then
   545     else if is_Free head andalso member (op =) fun_names (fst (dest_Free head)) then
   534       co_dissect_eqn_ctr sequential fun_names corec_specs eqn' imp_prems imp_rhs matchedss
   546       co_dissect_eqn_ctr sequential fun_names corec_specs eqn' imp_prems imp_rhs matchedsss
   535     else
   547     else
   536       primrec_error_eqn "malformed function equation" eqn
   548       primrec_error_eqn "malformed function equation" eqn
   537   end;
   549   end;
   538 
   550 
   539 fun build_corec_arg_disc ctr_specs {fun_args, ctr_no, prems, ...} =
   551 fun build_corec_arg_disc ctr_specs {fun_args, ctr_no, prems, ...} =
   541     mk_conjs prems
   553     mk_conjs prems
   542     |> curry subst_bounds (List.rev fun_args)
   554     |> curry subst_bounds (List.rev fun_args)
   543     |> HOLogic.tupled_lambda (HOLogic.mk_tuple fun_args)
   555     |> HOLogic.tupled_lambda (HOLogic.mk_tuple fun_args)
   544     |> K |> nth_map (the (#pred (nth ctr_specs ctr_no)));
   556     |> K |> nth_map (the (#pred (nth ctr_specs ctr_no)));
   545 
   557 
   546 fun build_corec_arg_no_call sel_eqns sel = find_first (equal sel o #sel) sel_eqns
   558 fun build_corec_arg_no_call sel_eqns sel =
   547   |> try (fn SOME sel_eqn => (#fun_args sel_eqn, #rhs_term sel_eqn))
   559   find_first (equal sel o #sel) sel_eqns
   548   |> the_default ([], undef_const)
   560   |> try (fn SOME {fun_args, rhs_term, ...} => abs_tuple fun_args rhs_term)
   549   |-> abs_tuple
   561   |> the_default undef_const
   550   |> K;
   562   |> K;
   551 
   563 
   552 fun build_corec_arg_direct_call lthy has_call sel_eqns sel =
   564 fun build_corec_arg_direct_call lthy has_call sel_eqns sel =
   553   let
   565   let
   554     val maybe_sel_eqn = find_first (equal sel o #sel) sel_eqns;
   566     val maybe_sel_eqn = find_first (equal sel o #sel) sel_eqns;
   571       let
   583       let
   572         fun subst (Abs (v, T, b)) = Abs (v, T, subst b)
   584         fun subst (Abs (v, T, b)) = Abs (v, T, subst b)
   573           | subst (t as _ $ _) =
   585           | subst (t as _ $ _) =
   574             let val (u, vs) = strip_comb t in
   586             let val (u, vs) = strip_comb t in
   575               if is_Free u andalso has_call u then
   587               if is_Free u andalso has_call u then
   576                 Const (@{const_name Inr}, dummyT) $ (*HOLogic.mk_tuple vs*)
   588                 Const (@{const_name Inr}, dummyT) $
   577                   (try (foldr1 (fn (x, y) => Const (@{const_name Pair}, dummyT) $ x $ y)) vs
   589                   (try (foldr1 (fn (x, y) => Const (@{const_name Pair}, dummyT) $ x $ y)) vs
   578                     |> the_default (hd vs))
   590                     |> the_default (hd vs))
   579               else if try (fst o dest_Const) u = SOME @{const_name prod_case} then
   591               else if try (fst o dest_Const) u = SOME @{const_name prod_case} then
   580                 list_comb (u |> map_types (K dummyT), map subst vs)
   592                 list_comb (u |> map_types (K dummyT), map subst vs)
   581               else
   593               else
   584           | subst t = t;
   596           | subst t = t;
   585       in
   597       in
   586         subst
   598         subst
   587       end;
   599       end;
   588     fun massage rhs_term t = massage_indirect_corec_call
   600     fun massage rhs_term t = massage_indirect_corec_call
   589       lthy has_call rewrite [] (fastype_of t |> range_type) rhs_term;
   601       lthy has_call rewrite [] (range_type (fastype_of t)) rhs_term;
   590   in
   602   in
   591     if is_none maybe_sel_eqn then I else
   603     if is_none maybe_sel_eqn then I else
   592       abs_tuple (#fun_args (the maybe_sel_eqn)) o massage (#rhs_term (the maybe_sel_eqn))
   604       abs_tuple (#fun_args (the maybe_sel_eqn)) o massage (#rhs_term (the maybe_sel_eqn))
   593   end;
   605   end;
   594 
   606 
   611         #> fold (fn (sel, n) => nth_map n
   623         #> fold (fn (sel, n) => nth_map n
   612           (build_corec_arg_indirect_call lthy has_call sel_eqns sel)) indirect_calls'
   624           (build_corec_arg_indirect_call lthy has_call sel_eqns sel)) indirect_calls'
   613       end
   625       end
   614   end;
   626   end;
   615 
   627 
   616 fun mk_real_disc_eqns ctr_specs disc_eqns =
       
   617   let
       
   618     val real_disc_eqns =
       
   619       if length disc_eqns = 0 then disc_eqns
       
   620       else if length disc_eqns = length ctr_specs - 1 then
       
   621         let
       
   622           val n = 0 upto length ctr_specs
       
   623             |> the(*###*) o find_first (fn idx => not (exists (equal idx o #ctr_no) disc_eqns));
       
   624           val extra_disc_eqn = {
       
   625             fun_name = #fun_name (hd disc_eqns),
       
   626             fun_args = #fun_args (hd disc_eqns),
       
   627             ctr_no = n,
       
   628             prems = maps (invert_prems o #prems) disc_eqns,
       
   629             user_eqn = Const (@{const_name undefined}, dummyT)};
       
   630         in
       
   631           chop n disc_eqns ||> cons extra_disc_eqn |> (op @)
       
   632         end
       
   633       else disc_eqns;
       
   634   in
       
   635     real_disc_eqns
       
   636   end;
       
   637 
       
   638 fun co_build_defs lthy bs mxs has_call arg_Tss corec_specs disc_eqnss sel_eqnss =
   628 fun co_build_defs lthy bs mxs has_call arg_Tss corec_specs disc_eqnss sel_eqnss =
   639   let
   629   let
   640     val _ = disc_eqnss |> map (fn x =>
       
   641       let val d = duplicates ((op =) o pairself #ctr_no) x in null d orelse
       
   642         primrec_error_eqns "excess discriminator equations in definition"
       
   643           (maps (fn t => filter (equal (#ctr_no t) o #ctr_no) x) d |> map #user_eqn) end);
       
   644     val corec_specs' = take (length bs) corec_specs;
   630     val corec_specs' = take (length bs) corec_specs;
   645     val corecs = map #corec corec_specs';
   631     val corecs = map #corec corec_specs';
   646     val ctr_specss = map #ctr_specs corec_specs';
   632     val ctr_specss = map #ctr_specs corec_specs';
   647     val real_disc_eqnss = map2 mk_real_disc_eqns ctr_specss disc_eqnss;
       
   648     val corec_args = hd corecs
   633     val corec_args = hd corecs
   649       |> fst o split_last o binder_types o fastype_of
   634       |> fst o split_last o binder_types o fastype_of
   650       |> map (Const o pair @{const_name undefined})
   635       |> map (Const o pair @{const_name undefined})
   651       |> fold2 (fold o build_corec_arg_disc) ctr_specss real_disc_eqnss
   636       |> fold2 (fold o build_corec_arg_disc) ctr_specss disc_eqnss
   652       |> fold2 (fold o build_corec_args_sel lthy has_call) sel_eqnss ctr_specss;
   637       |> fold2 (fold o build_corec_args_sel lthy has_call) sel_eqnss ctr_specss;
   653     fun currys Ts t = if length Ts <= 1 then t else
   638     fun currys Ts t = if length Ts <= 1 then t else
   654       t $ foldr1 (fn (u, v) => HOLogic.pair_const dummyT dummyT $ u $ v)
   639       t $ foldr1 (fn (u, v) => HOLogic.pair_const dummyT dummyT $ u $ v)
   655         (length Ts - 1 downto 0 |> map Bound)
   640         (length Ts - 1 downto 0 |> map Bound)
   656       |> fold_rev (fn T => fn u => Abs (Name.uu, T, u)) Ts;
   641       |> fold_rev (Term.abs o pair Name.uu) Ts;
   657 
   642 
   658 val _ = tracing ("corecursor arguments:\n    \<cdot> " ^
   643 val _ = tracing ("corecursor arguments:\n    \<cdot> " ^
   659  space_implode "\n    \<cdot> " (map (Syntax.string_of_term lthy) corec_args));
   644  space_implode "\n    \<cdot> " (map (Syntax.string_of_term lthy) corec_args));
   660 
   645 
   661     val exclss' =
   646     val exclss' =
   662       real_disc_eqnss
   647       disc_eqnss
   663       |> map (map (fn {fun_args, ctr_no, prems, ...} => (fun_args, ctr_no, prems))
   648       |> map (map (fn {fun_args, ctr_no, prems, ...} => (fun_args, ctr_no, prems))
   664         #> fst o (fn xs => fold_map (fn x => fn ys => ((x, ys), ys @ [x])) xs [])
   649         #> fst o (fn xs => fold_map (fn x => fn ys => ((x, ys), ys @ [x])) xs [])
   665         #> maps (uncurry (map o pair)
   650         #> maps (uncurry (map o pair)
   666           #> map (fn ((fun_args, c, x), (_, c', y)) => ((c, c'), (x, mk_not (mk_conjs y)))
   651           #> map (fn ((fun_args, c, x), (_, c', y)) => ((c, c'), (x, mk_not (mk_conjs y)))
   667             ||> apfst (map HOLogic.mk_Trueprop) o apsnd HOLogic.mk_Trueprop
   652             ||> apfst (map HOLogic.mk_Trueprop) o apsnd HOLogic.mk_Trueprop
   674     |> Syntax.check_terms lthy
   659     |> Syntax.check_terms lthy
   675     |> map3 (fn b => fn mx => fn t => ((b, mx), ((Binding.map_name Thm.def_name b, []), t))) bs mxs
   660     |> map3 (fn b => fn mx => fn t => ((b, mx), ((Binding.map_name Thm.def_name b, []), t))) bs mxs
   676     |> rpair exclss'
   661     |> rpair exclss'
   677   end;
   662   end;
   678 
   663 
       
   664 fun mk_real_disc_eqns fun_binding arg_Ts {ctr_specs, ...} disc_eqns =
       
   665   if length disc_eqns <> length ctr_specs - 1 then disc_eqns else
       
   666     let
       
   667       val n = 0 upto length ctr_specs
       
   668         |> the o find_first (fn idx => not (exists (equal idx o #ctr_no) disc_eqns));
       
   669       val extra_disc_eqn = {
       
   670         fun_name = Binding.name_of fun_binding,
       
   671         fun_T = arg_Ts ---> body_type (fastype_of (#ctr (hd ctr_specs))),
       
   672         fun_args = the_default (map (curry Free Name.uu) arg_Ts) (try (#fun_args o hd) disc_eqns),
       
   673         ctr = #ctr (nth ctr_specs n),
       
   674         ctr_no = n,
       
   675         disc = #disc (nth ctr_specs n),
       
   676         prems = maps (invert_prems o #prems) disc_eqns,
       
   677         user_eqn = undef_const};
       
   678     in
       
   679       chop n disc_eqns ||> cons extra_disc_eqn |> (op @)
       
   680     end;
       
   681 
   679 fun add_primcorec sequential fixes specs lthy =
   682 fun add_primcorec sequential fixes specs lthy =
   680   let
   683   let
   681     val (bs, mxs) = map_split (apfst fst) fixes;
   684     val (bs, mxs) = map_split (apfst fst) fixes;
   682     val (arg_Ts, res_Ts) = map (strip_type o snd o fst #>> HOLogic.mk_tupleT) fixes |> split_list;
   685     val (arg_Ts, res_Ts) = map (strip_type o snd o fst #>> HOLogic.mk_tupleT) fixes |> split_list;
   683 
   686 
   697 
   700 
   698     val (eqns_data, _) =
   701     val (eqns_data, _) =
   699       fold_map (co_dissect_eqn sequential fun_names corec_specs) (map snd specs) []
   702       fold_map (co_dissect_eqn sequential fun_names corec_specs) (map snd specs) []
   700       |>> flat;
   703       |>> flat;
   701 
   704 
   702     val disc_eqnss = map_filter (try (fn Disc x => x)) eqns_data
   705     val disc_eqnss' = map_filter (try (fn Disc x => x)) eqns_data
   703       |> partition_eq ((op =) o pairself #fun_name)
   706       |> partition_eq ((op =) o pairself #fun_name)
   704       |> finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names |> fst
   707       |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
   705       |> map (sort ((op <) o pairself #ctr_no |> make_ord) o flat o snd);
   708       |> map (sort ((op <) o pairself #ctr_no |> make_ord) o flat o snd);
       
   709     val _ = disc_eqnss' |> map (fn x =>
       
   710       let val d = duplicates ((op =) o pairself #ctr_no) x in null d orelse
       
   711         primrec_error_eqns "excess discriminator equations in definition"
       
   712           (maps (fn t => filter (equal (#ctr_no t) o #ctr_no) x) d |> map #user_eqn) end);
   706 
   713 
   707     val sel_eqnss = map_filter (try (fn Sel x => x)) eqns_data
   714     val sel_eqnss = map_filter (try (fn Sel x => x)) eqns_data
   708       |> partition_eq ((op =) o pairself #fun_name)
   715       |> partition_eq ((op =) o pairself #fun_name)
   709       |> finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names |> fst
   716       |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
   710       |> map (flat o snd);
   717       |> map (flat o snd);
   711 
   718 
   712     val has_call = exists_subterm (map (fst #>> Binding.name_of #> Free) fixes |> member (op =));
   719     val has_call = exists_subterm (map (fst #>> Binding.name_of #> Free) fixes |> member (op =));
   713     val arg_Tss = map (binder_types o snd o fst) fixes;
   720     val arg_Tss = map (binder_types o snd o fst) fixes;
       
   721     val disc_eqnss = map4 mk_real_disc_eqns bs arg_Tss corec_specs disc_eqnss';
   714     val (defs, exclss') =
   722     val (defs, exclss') =
   715       co_build_defs lthy' bs mxs has_call arg_Tss corec_specs disc_eqnss sel_eqnss;
   723       co_build_defs lthy' bs mxs has_call arg_Tss corec_specs disc_eqnss sel_eqnss;
   716 
   724 
   717     (* try to prove (automatically generated) tautologies by ourselves *)
   725     (* try to prove (automatically generated) tautologies by ourselves *)
   718     val exclss'' = exclss'
   726     val exclss'' = exclss'
   733           |-> fold (fn ((c, c'), thm) => nth_map c (nth_map c' (K [thm])));
   741           |-> fold (fn ((c, c'), thm) => nth_map c (nth_map c' (K [thm])));
   734         val exclssss = (exclss' ~~ taut_thmss |> map (op @), fun_names ~~ corec_specs)
   742         val exclssss = (exclss' ~~ taut_thmss |> map (op @), fun_names ~~ corec_specs)
   735           |-> map2 (fn excls => fn (_, {ctr_specs, ...}) => mk_exclsss excls (length ctr_specs));
   743           |-> map2 (fn excls => fn (_, {ctr_specs, ...}) => mk_exclsss excls (length ctr_specs));
   736 
   744 
   737         fun prove_disc {ctr_specs, ...} exclsss
   745         fun prove_disc {ctr_specs, ...} exclsss
   738             {fun_name, fun_args, ctr_no, prems, ...} =
   746             {fun_name, fun_T, fun_args, ctr_no, prems, user_eqn, ...} =
       
   747           if user_eqn = undef_const then [] else
       
   748             let
       
   749               val disc_corec = nth ctr_specs ctr_no |> #disc_corec;
       
   750               val k = 1 + ctr_no;
       
   751               val m = length prems;
       
   752               val t =
       
   753                 (* FIXME use applied_fun from dissect_\<dots> instead? *)
       
   754                 list_comb (Free (fun_name, fun_T), map Bound (length fun_args - 1 downto 0))
       
   755                 |> curry betapply (#disc (nth ctr_specs ctr_no)) (*###*)
       
   756                 |> HOLogic.mk_Trueprop
       
   757                 |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
       
   758                 |> curry Logic.list_all (map dest_Free fun_args);
       
   759             in
       
   760               mk_primcorec_disc_tac lthy def_thms disc_corec k m exclsss
       
   761               |> K |> Goal.prove lthy [] [] t
       
   762               |> pair (#disc (nth ctr_specs ctr_no))
       
   763               |> single
       
   764             end;
       
   765 
       
   766         fun prove_sel {ctr_specs, nested_maps, nested_map_idents, nested_map_comps, ...}
       
   767             disc_eqns exclsss {fun_name, fun_T, fun_args, ctr, sel, rhs_term, ...} =
   739           let
   768           let
   740             val disc_corec = nth ctr_specs ctr_no |> #disc_corec;
   769             val (SOME ctr_spec) = find_first (equal ctr o #ctr) ctr_specs;
       
   770             val ctr_no = find_index (equal ctr o #ctr) ctr_specs;
       
   771             val prems = the_default (maps (invert_prems o #prems) disc_eqns)
       
   772                 (find_first (equal ctr_no o #ctr_no) disc_eqns |> Option.map #prems);
       
   773             val sel_corec = find_index (equal sel) (#sels ctr_spec)
       
   774               |> nth (#sel_corecs ctr_spec);
   741             val k = 1 + ctr_no;
   775             val k = 1 + ctr_no;
   742             val m = length prems;
   776             val m = length prems;
   743             val t =
   777             val t =
   744               (* FIXME use applied_fun from dissect_\<dots> instead? *)
   778               list_comb (Free (fun_name, fun_T), map Bound (length fun_args - 1 downto 0))
   745               list_comb (Free (fun_name, dummyT), map Bound (length fun_args - 1 downto 0))
   779               |> curry betapply sel
   746               |> curry betapply (#disc (nth ctr_specs ctr_no)) (*###*)
   780               |> rpair (abstract (List.rev fun_args) rhs_term)
   747               |> HOLogic.mk_Trueprop
   781               |> HOLogic.mk_Trueprop o HOLogic.mk_eq
   748               |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
   782               |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
   749               |> curry Logic.list_all (map dest_Free fun_args)
   783               |> curry Logic.list_all (map dest_Free fun_args);
   750               |> Syntax.check_term lthy(*###*);
       
   751           in
       
   752             mk_primcorec_disc_tac lthy def_thms disc_corec k m exclsss
       
   753             |> K |> Goal.prove lthy [] [] t
       
   754           end;
       
   755 
       
   756         (* FIXME don't use user_eqn (cf. constructor view reduction),
       
   757                  instead generate "sel" and "code" theorems ourselves *)
       
   758         fun prove_sel
       
   759           ((fun_name, {ctr_specs, nested_maps, nested_map_idents, nested_map_comps, ...}),
       
   760             disc_eqns) exclsss sel_eqn =
       
   761           let
       
   762             val (SOME ctr_spec) = find_first (equal (#ctr sel_eqn) o #ctr) ctr_specs;
       
   763             val ctr_no = find_index (equal (#ctr sel_eqn) o #ctr) ctr_specs;
       
   764             val prems = the_default (maps (invert_prems o #prems) disc_eqns)
       
   765                 (find_first (equal ctr_no o #ctr_no) disc_eqns |> Option.map #prems);
       
   766             val sel_corec = find_index (equal (#sel sel_eqn)) (#sels ctr_spec)
       
   767               |> nth (#sel_corecs ctr_spec);
       
   768             val k = 1 + ctr_no;
       
   769             val m = length prems;
       
   770             val t = #user_eqn sel_eqn
       
   771               |> abstract 0 (List.rev (#fun_args sel_eqn)) (* FIXME do this in dissect_\<dots> *)
       
   772               |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
       
   773               |> curry Logic.list_all (map dest_Free (#fun_args sel_eqn));
       
   774           in
   784           in
   775             mk_primcorec_ctr_or_sel_tac lthy def_thms sel_corec k m exclsss
   785             mk_primcorec_ctr_or_sel_tac lthy def_thms sel_corec k m exclsss
   776               nested_maps nested_map_idents nested_map_comps
   786               nested_maps nested_map_idents nested_map_comps
   777             |> K |> Goal.prove lthy [] [] t
   787             |> K |> Goal.prove lthy [] [] t
       
   788             |> pair sel
   778           end;
   789           end;
   779 
   790 
   780         val disc_notes =
   791         fun prove_ctr (_, disc_thms) (_, sel_thms') disc_eqns sel_eqns
   781           fun_names ~~
   792             {ctr, disc, sels, collapse, ...} =
   782             map3 (map oo prove_disc) (take (length disc_eqnss) corec_specs) exclssss disc_eqnss
   793           if not (exists (equal ctr o #ctr) disc_eqns)
       
   794 andalso (warning ("no disc_eqn for ctr " ^ Syntax.string_of_term lthy ctr); true)
       
   795             orelse (* don't try to prove theorems where some sel_eqns are missing *)
       
   796               filter (equal ctr o #ctr) sel_eqns
       
   797               |> fst o finds ((op =) o apsnd #sel) sels
       
   798               |> exists (null o snd)
       
   799 andalso (warning ("sel_eqn(s) missing for ctr " ^ Syntax.string_of_term lthy ctr); true)
       
   800             orelse
       
   801               #user_eqn (the (find_first (equal ctr o #ctr) disc_eqns)) = undef_const
       
   802 andalso (warning ("auto-generated disc_eqn for ctr " ^ Syntax.string_of_term lthy ctr); true)
       
   803           then [] else
       
   804             let
       
   805 val _ = tracing ("ctr = " ^ Syntax.string_of_term lthy ctr);
       
   806 val _ = tracing ("disc = " ^ Syntax.string_of_term lthy (#disc (the (find_first (equal ctr o #ctr) disc_eqns))));
       
   807               val {fun_name, fun_T, fun_args, prems, ...} =
       
   808                 the (find_first (equal ctr o #ctr) disc_eqns);
       
   809               val m = length prems;
       
   810               val t = sel_eqns
       
   811                 |> fst o finds ((op =) o apsnd #sel) sels
       
   812                 |> map (snd #> (fn [x] => (List.rev (#fun_args x), #rhs_term x)) #-> abstract)
       
   813                 |> curry list_comb ctr
       
   814                 |> curry HOLogic.mk_eq (list_comb (Free (fun_name, fun_T),
       
   815                   map Bound (length fun_args - 1 downto 0)))
       
   816                 |> HOLogic.mk_Trueprop
       
   817                 |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
       
   818                 |> curry Logic.list_all (map dest_Free fun_args);
       
   819               val disc_thm = the_default TrueI (AList.lookup (op =) disc_thms disc);
       
   820               val sel_thms = map snd (filter (member (op =) sels o fst) sel_thms');
       
   821 val _ = tracing ("t = " ^ Syntax.string_of_term lthy t);
       
   822 val _ = tracing ("m = " ^ @{make_string} m);
       
   823 val _ = tracing ("collapse = " ^ @{make_string} collapse);
       
   824 val _ = tracing ("disc_thm = " ^ @{make_string} disc_thm);
       
   825 val _ = tracing ("sel_thms = " ^ @{make_string} sel_thms);
       
   826             in
       
   827               mk_primcorec_ctr_of_dtr_tac lthy m collapse disc_thm sel_thms
       
   828               |> K |> Goal.prove lthy [] [] t
       
   829               |> single
       
   830           end;
       
   831 
       
   832         val (disc_notes, disc_thmss) =
       
   833           fun_names ~~ map3 (maps oo prove_disc) corec_specs exclssss disc_eqnss
       
   834           |> `(map (fn (fun_name, thms) =>
       
   835             ((Binding.qualify true fun_name (@{binding disc}), simp_attrs), [(map snd thms, [])])));
       
   836         val (sel_notes, sel_thmss) =
       
   837           fun_names ~~ map4 (map ooo prove_sel) corec_specs disc_eqnss exclssss sel_eqnss
       
   838           |> `(map (fn (fun_name, thms) =>
       
   839             ((Binding.qualify true fun_name (@{binding sel}), simp_attrs), [(map snd thms, [])])));
       
   840         val ctr_notes =
       
   841           fun_names ~~ map5 (maps oooo prove_ctr) disc_thmss sel_thmss
       
   842             disc_eqnss sel_eqnss (map #ctr_specs corec_specs)
   783           |> map (fn (fun_name, thms) =>
   843           |> map (fn (fun_name, thms) =>
   784             ((Binding.qualify true fun_name (@{binding disc}), simp_attrs), [(thms, [])]));
   844             ((Binding.qualify true fun_name (@{binding ctr}), simp_attrs), [(thms, [])]));
   785         val sel_notes =
       
   786           fun_names ~~
       
   787             map3 (map oo prove_sel) (fun_names ~~ corec_specs ~~ disc_eqnss) exclssss sel_eqnss
       
   788           |> map (fn (fun_name, thms) =>
       
   789             ((Binding.qualify true fun_name (@{binding sel}), simp_attrs), [(thms, [])]));
       
   790       in
   845       in
   791         lthy |> snd o Local_Theory.notes (disc_notes @ sel_notes)
   846         lthy |> snd o Local_Theory.notes
       
   847           (filter (not o null o fst o the_single o snd) (disc_notes @ sel_notes @ ctr_notes))
   792       end;
   848       end;
   793   in
   849   in
   794     lthy'
   850     lthy'
   795     |> Proof.theorem NONE (curry (op #->) (fold_map Local_Theory.define defs) o prove) obligationss
   851     |> Proof.theorem NONE (curry (op #->) (fold_map Local_Theory.define defs) o prove) obligationss
   796     |> Proof.refine (Method.primitive_text I)
   852     |> Proof.refine (Method.primitive_text I)