src/HOL/BNF/Tools/bnf_fp_rec_sugar.ML
changeset 53654 8b9ea4420f81
parent 53411 ab4edf89992f
child 53692 2c04e30c2f3e
equal deleted inserted replaced
53648:924579729403 53654:8b9ea4420f81
    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 (fn u => Abs (Name.uu, dummyT, u))) (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,
       
    41   strip_qnt_body @{const_name all} t)
       
    42 fun mk_not @{const True} = @{const False}
       
    43   | mk_not @{const False} = @{const True}
       
    44   | mk_not (@{const Not} $ t) = t
       
    45   | mk_not (@{const Trueprop} $ t) = @{const Trueprop} $ mk_not t
       
    46   | mk_not t = HOLogic.mk_not t
       
    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};
       
    49 
       
    50 fun invert_prems [t] = map mk_not (HOLogic.disjuncts t)
       
    51   | invert_prems ts = [mk_disjs (map mk_not ts)];
    40 
    52 
    41 val simp_attrs = @{attributes [simp]};
    53 val simp_attrs = @{attributes [simp]};
    42 
    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;
    43 
    59 
    44 
    60 
    45 (* Primrec *)
    61 (* Primrec *)
    46 
    62 
    47 type eqn_data = {
    63 type eqn_data = {
    56   user_eqn: term
    72   user_eqn: term
    57 };
    73 };
    58 
    74 
    59 fun dissect_eqn lthy fun_names eqn' =
    75 fun dissect_eqn lthy fun_names eqn' =
    60   let
    76   let
    61     val eqn = subst_bounds (strip_qnt_vars @{const_name all} eqn' |> map Free |> rev,
    77     val eqn = drop_All eqn' |> HOLogic.dest_Trueprop
    62         strip_qnt_body @{const_name all} eqn') |> HOLogic.dest_Trueprop
    78       handle TERM _ =>
    63         handle TERM _ =>
    79         primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn';
    64           primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn';
       
    65     val (lhs, rhs) = HOLogic.dest_eq eqn
    80     val (lhs, rhs) = HOLogic.dest_eq eqn
    66         handle TERM _ =>
    81         handle TERM _ =>
    67           primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn';
    82           primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn';
    68     val (fun_name, args) = strip_comb lhs
    83     val (fun_name, args) = strip_comb lhs
    69       |>> (fn x => if is_Free x then fst (dest_Free x)
    84       |>> (fn x => if is_Free x then fst (dest_Free x)
   340       |> map (fn (thmN, thms, attrs) =>
   355       |> map (fn (thmN, thms, attrs) =>
   341         ((Binding.qualify true common_name (Binding.name thmN), attrs), [(thms, [])]));
   356         ((Binding.qualify true common_name (Binding.name thmN), attrs), [(thms, [])]));
   342   in
   357   in
   343     lthy'
   358     lthy'
   344     |> fold_map Local_Theory.define defs
   359     |> fold_map Local_Theory.define defs
   345     |-> (fn def_thms' => fold_map3 (prove def_thms') (take actual_nn rec_specs)
   360     |-> snd oo (fn def_thms' => fold_map3 (prove def_thms') (take actual_nn rec_specs)
   346       (take actual_nn induct_thms) funs_data)
   361       (take actual_nn induct_thms) funs_data)
   347     |> snd
   362     |> snd o Local_Theory.notes common_notes
   348     |> Local_Theory.notes common_notes |> snd
       
   349   end;
   363   end;
   350 
   364 
   351 fun add_primrec_cmd raw_fixes raw_specs lthy =
   365 fun add_primrec_cmd raw_fixes raw_specs lthy =
   352   let
   366   let
   353     val _ = let val d = duplicates (op =) (map (Binding.name_of o #1) raw_fixes) in null d orelse
   367     val _ = let val d = duplicates (op =) (map (Binding.name_of o #1) raw_fixes) in null d orelse
   369 
   383 
   370 type co_eqn_data_disc = {
   384 type co_eqn_data_disc = {
   371   fun_name: string,
   385   fun_name: string,
   372   fun_args: term list,
   386   fun_args: term list,
   373   ctr_no: int, (*###*)
   387   ctr_no: int, (*###*)
   374   cond: term,
   388   prems: term list,
   375   user_eqn: term
   389   user_eqn: term
   376 };
   390 };
   377 type co_eqn_data_sel = {
   391 type co_eqn_data_sel = {
   378   fun_name: string,
   392   fun_name: string,
   379   fun_args: term list,
   393   fun_args: term list,
   384 };
   398 };
   385 datatype co_eqn_data =
   399 datatype co_eqn_data =
   386   Disc of co_eqn_data_disc |
   400   Disc of co_eqn_data_disc |
   387   Sel of co_eqn_data_sel;
   401   Sel of co_eqn_data_sel;
   388 
   402 
   389 fun co_dissect_eqn_disc sequential fun_name_corec_spec_list eqn' imp_lhs' imp_rhs matched_conds =
   403 fun co_dissect_eqn_disc sequential fun_names corec_specs prems' concl matchedss =
   390   let
   404   let
   391     fun find_subterm p = let (* FIXME \<exists>? *)
   405     fun find_subterm p = let (* FIXME \<exists>? *)
   392       fun f (t as u $ v) = if p t then SOME t else merge_options (f u, f v)
   406       fun f (t as u $ v) = if p t then SOME t else merge_options (f u, f v)
   393         | f t = if p t then SOME t else NONE
   407         | f t = if p t then SOME t else NONE
   394       in f end;
   408       in f end;
   395 
   409 
   396     val fun_name = imp_rhs
   410     val applied_fun = concl
   397       |> perhaps (try HOLogic.dest_not)
   411       |> find_subterm (member ((op =) o apsnd SOME) fun_names o try (fst o dest_Free o head_of))
   398       |> `(try (fst o dest_Free o head_of o snd o dest_comb))
   412       |> the
   399       ||> (try (fst o dest_Free o head_of o fst o HOLogic.dest_eq))
   413       handle Option.Option => primrec_error_eqn "malformed discriminator equation" concl;
   400       |> the o merge_options;
   414     val (fun_name, fun_args) = strip_comb applied_fun |>> fst o dest_Free;
   401     val corec_spec = the (AList.lookup (op =) fun_name_corec_spec_list fun_name)
   415     val corec_spec = the (AList.lookup (op =) (fun_names ~~ corec_specs) fun_name);
   402       handle Option.Option => primrec_error_eqn "malformed discriminator equation" imp_rhs;
       
   403 
   416 
   404     val discs = #ctr_specs corec_spec |> map #disc;
   417     val discs = #ctr_specs corec_spec |> map #disc;
   405     val ctrs = #ctr_specs corec_spec |> map #ctr;
   418     val ctrs = #ctr_specs corec_spec |> map #ctr;
   406     val not_disc = head_of imp_rhs = @{term Not};
   419     val not_disc = head_of concl = @{term Not};
   407     val _ = not_disc andalso length ctrs <> 2 andalso
   420     val _ = not_disc andalso length ctrs <> 2 andalso
   408       primrec_error_eqn "\<not>ed discriminator for a type with \<noteq> 2 constructors" imp_rhs;
   421       primrec_error_eqn "\<not>ed discriminator for a type with \<noteq> 2 constructors" concl;
   409     val disc = find_subterm (member (op =) discs o head_of) imp_rhs;
   422     val disc = find_subterm (member (op =) discs o head_of) concl;
   410     val eq_ctr0 = imp_rhs |> perhaps (try (HOLogic.dest_not)) |> try (HOLogic.dest_eq #> snd)
   423     val eq_ctr0 = concl |> perhaps (try (HOLogic.dest_not)) |> try (HOLogic.dest_eq #> snd)
   411         |> (fn SOME t => let val n = find_index (equal t) ctrs in
   424         |> (fn SOME t => let val n = find_index (equal t) ctrs in
   412           if n >= 0 then SOME n else NONE end | _ => NONE);
   425           if n >= 0 then SOME n else NONE end | _ => NONE);
   413     val _ = is_some disc orelse is_some eq_ctr0 orelse
   426     val _ = is_some disc orelse is_some eq_ctr0 orelse
   414       primrec_error_eqn "no discriminator in equation" imp_rhs;
   427       primrec_error_eqn "no discriminator in equation" concl;
   415     val ctr_no' =
   428     val ctr_no' =
   416       if is_none disc then the eq_ctr0 else find_index (equal (head_of (the disc))) discs;
   429       if is_none disc then the eq_ctr0 else find_index (equal (head_of (the disc))) discs;
   417     val ctr_no = if not_disc then 1 - ctr_no' else ctr_no';
   430     val ctr_no = if not_disc then 1 - ctr_no' else ctr_no';
   418     val fun_args = if is_none disc
   431 
   419       then imp_rhs |> perhaps (try HOLogic.dest_not) |> HOLogic.dest_eq |> fst |> strip_comb |> snd
   432     val catch_all = try (fst o dest_Free o the_single) prems' = SOME Name.uu_;
   420       else the disc |> the_single o snd o strip_comb
   433     val matcheds = AList.lookup (op =) matchedss fun_name |> the_default [];
   421         |> (fn t => if free_name (head_of t) = SOME fun_name
   434     val prems = map (abstract 0 (List.rev fun_args)) prems';
   422           then snd (strip_comb t) else []);
   435     val real_prems = (if catch_all orelse sequential then invert_prems matcheds else []) @
   423 
   436       (if catch_all then [] else prems);
   424     val mk_conjs = try (foldr1 HOLogic.mk_conj) #> the_default @{const True};
   437 
   425     val mk_disjs = try (foldr1 HOLogic.mk_disj) #> the_default @{const False};
   438     val matchedss' = AList.delete (op =) fun_name matchedss
   426     val catch_all = try (fst o dest_Free o the_single) imp_lhs' = SOME Name.uu_;
   439       |> cons (fun_name, if sequential then prems @ matcheds else real_prems @ matcheds);
   427     val matched_cond = filter (equal fun_name o fst) matched_conds |> map snd |> mk_disjs;
   440 
   428     val imp_lhs = mk_conjs imp_lhs'
   441     val user_eqn =
   429       |> incr_boundvars (length fun_args)
   442       (real_prems, betapply (#disc (nth (#ctr_specs corec_spec) ctr_no), applied_fun))
   430       |> subst_atomic (fun_args ~~ map Bound (length fun_args - 1 downto 0))
   443       |>> map HOLogic.mk_Trueprop ||> HOLogic.mk_Trueprop
   431     val cond =
   444       |> Logic.list_implies;
   432       if catch_all then
       
   433         matched_cond |> HOLogic.mk_not
       
   434       else if sequential then
       
   435         HOLogic.mk_conj (HOLogic.mk_not matched_cond, imp_lhs)
       
   436       else
       
   437         imp_lhs;
       
   438 
       
   439     val matched_conds' =
       
   440       (fun_name, if catch_all orelse not sequential then cond else imp_lhs) :: matched_conds;
       
   441   in
   445   in
   442     (Disc {
   446     (Disc {
   443       fun_name = fun_name,
   447       fun_name = fun_name,
   444       fun_args = fun_args,
   448       fun_args = fun_args,
   445       ctr_no = ctr_no,
   449       ctr_no = ctr_no,
   446       cond = cond,
   450       prems = real_prems,
   447       user_eqn = eqn'
   451       user_eqn = user_eqn
   448     }, matched_conds')
   452     }, matchedss')
   449   end;
   453   end;
   450 
   454 
   451 fun co_dissect_eqn_sel fun_name_corec_spec_list eqn' eqn =
   455 fun co_dissect_eqn_sel fun_names corec_specs eqn' eqn =
   452   let
   456   let
   453     val (lhs, rhs) = HOLogic.dest_eq eqn
   457     val (lhs, rhs) = HOLogic.dest_eq eqn
   454       handle TERM _ =>
   458       handle TERM _ =>
   455         primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn;
   459         primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn;
   456     val sel = head_of lhs;
   460     val sel = head_of lhs;
   457     val (fun_name, fun_args) = dest_comb lhs |> snd |> strip_comb |> apfst (fst o dest_Free)
   461     val (fun_name, fun_args) = dest_comb lhs |> snd |> strip_comb |> apfst (fst o dest_Free)
   458       handle TERM _ =>
   462       handle TERM _ =>
   459         primrec_error_eqn "malformed selector argument in left-hand side" eqn;
   463         primrec_error_eqn "malformed selector argument in left-hand side" eqn;
   460     val corec_spec = the (AList.lookup (op =) fun_name_corec_spec_list fun_name)
   464     val corec_spec = the (AList.lookup (op =) (fun_names ~~ corec_specs) fun_name)
   461       handle Option.Option => primrec_error_eqn "malformed selector argument in left-hand side" eqn;
   465       handle Option.Option => primrec_error_eqn "malformed selector argument in left-hand side" eqn;
   462     val (ctr_spec, sel) = #ctr_specs corec_spec
   466     val (ctr_spec, sel) = #ctr_specs corec_spec
   463       |> the o get_index (try (the o find_first (equal sel) o #sels))
   467       |> the o get_index (try (the o find_first (equal sel) o #sels))
   464       |>> nth (#ctr_specs corec_spec);
   468       |>> nth (#ctr_specs corec_spec);
       
   469     val user_eqn = drop_All eqn';
   465   in
   470   in
   466     Sel {
   471     Sel {
   467       fun_name = fun_name,
   472       fun_name = fun_name,
   468       fun_args = fun_args,
   473       fun_args = fun_args,
   469       ctr = #ctr ctr_spec,
   474       ctr = #ctr ctr_spec,
   470       sel = sel,
   475       sel = sel,
   471       rhs_term = rhs,
   476       rhs_term = rhs,
   472       user_eqn = eqn'
   477       user_eqn = user_eqn
   473     }
   478     }
   474   end;
   479   end;
   475 
   480 
   476 fun co_dissect_eqn_ctr sequential fun_name_corec_spec_list eqn' imp_lhs' imp_rhs matched_conds =
   481 fun co_dissect_eqn_ctr sequential fun_names corec_specs eqn' imp_prems imp_rhs matchedss =
   477   let 
   482   let 
   478     val (lhs, rhs) = HOLogic.dest_eq imp_rhs;
   483     val (lhs, rhs) = HOLogic.dest_eq imp_rhs;
   479     val fun_name = head_of lhs |> fst o dest_Free;
   484     val fun_name = head_of lhs |> fst o dest_Free;
   480     val corec_spec = the (AList.lookup (op =) fun_name_corec_spec_list fun_name);
   485     val corec_spec = the (AList.lookup (op =) (fun_names ~~ corec_specs) fun_name);
   481     val (ctr, ctr_args) = strip_comb rhs;
   486     val (ctr, ctr_args) = strip_comb rhs;
   482     val ctr_spec = the (find_first (equal ctr o #ctr) (#ctr_specs corec_spec))
   487     val ctr_spec = the (find_first (equal ctr o #ctr) (#ctr_specs corec_spec))
   483       handle Option.Option => primrec_error_eqn "not a constructor" ctr;
   488       handle Option.Option => primrec_error_eqn "not a constructor" ctr;
   484 
   489 
   485     val disc_imp_rhs = betapply (#disc ctr_spec, lhs);
   490     val disc_imp_rhs = betapply (#disc ctr_spec, lhs);
   486     val (maybe_eqn_data_disc, matched_conds') = if length (#ctr_specs corec_spec) = 1
   491     val (maybe_eqn_data_disc, matchedss') = if length (#ctr_specs corec_spec) = 1
   487       then (NONE, matched_conds)
   492       then (NONE, matchedss)
   488       else apfst SOME (co_dissect_eqn_disc
   493       else apfst SOME (co_dissect_eqn_disc
   489           sequential fun_name_corec_spec_list eqn' imp_lhs' disc_imp_rhs matched_conds);
   494           sequential fun_names corec_specs imp_prems disc_imp_rhs matchedss);
   490 
   495 
   491     val sel_imp_rhss = (#sels ctr_spec ~~ ctr_args)
   496     val sel_imp_rhss = (#sels ctr_spec ~~ ctr_args)
   492       |> map (fn (sel, ctr_arg) => HOLogic.mk_eq (betapply (sel, lhs), ctr_arg));
   497       |> map (fn (sel, ctr_arg) => HOLogic.mk_eq (betapply (sel, lhs), ctr_arg));
   493 
   498 
   494 val _ = tracing ("reduced\n    " ^ Syntax.string_of_term @{context} imp_rhs ^ "\nto\n    \<cdot> " ^
   499 val _ = tracing ("reduced\n    " ^ Syntax.string_of_term @{context} imp_rhs ^ "\nto\n    \<cdot> " ^
   495  (is_some maybe_eqn_data_disc ? K (Syntax.string_of_term @{context} disc_imp_rhs ^ "\n    \<cdot> ")) "" ^
   500  (is_some maybe_eqn_data_disc ? K (Syntax.string_of_term @{context} disc_imp_rhs ^ "\n    \<cdot> ")) "" ^
   496  space_implode "\n    \<cdot> " (map (Syntax.string_of_term @{context}) sel_imp_rhss));
   501  space_implode "\n    \<cdot> " (map (Syntax.string_of_term @{context}) sel_imp_rhss));
   497 
   502 
   498     val eqns_data_sel =
   503     val eqns_data_sel =
   499       map (co_dissect_eqn_sel fun_name_corec_spec_list eqn') sel_imp_rhss;
   504       map (co_dissect_eqn_sel fun_names corec_specs eqn') sel_imp_rhss;
   500   in
   505   in
   501     (map_filter I [maybe_eqn_data_disc] @ eqns_data_sel, matched_conds')
   506     (the_list maybe_eqn_data_disc @ eqns_data_sel, matchedss')
   502   end;
   507   end;
   503 
   508 
   504 fun co_dissect_eqn sequential fun_name_corec_spec_list eqn' matched_conds =
   509 fun co_dissect_eqn sequential fun_names corec_specs eqn' matchedss =
   505   let
   510   let
   506     val eqn = subst_bounds (strip_qnt_vars @{const_name all} eqn' |> map Free |> rev,
   511     val eqn = drop_All eqn'
   507         strip_qnt_body @{const_name all} eqn')
   512       handle TERM _ => primrec_error_eqn "malformed function equation" eqn';
   508         handle TERM _ => primrec_error_eqn "malformed function equation" eqn';
   513     val (imp_prems, imp_rhs) = Logic.strip_horn eqn
   509     val (imp_lhs', imp_rhs) = Logic.strip_horn eqn
       
   510       |> apfst (map HOLogic.dest_Trueprop) o apsnd HOLogic.dest_Trueprop;
   514       |> apfst (map HOLogic.dest_Trueprop) o apsnd HOLogic.dest_Trueprop;
   511 
   515 
   512     val head = imp_rhs
   516     val head = imp_rhs
   513       |> perhaps (try HOLogic.dest_not) |> perhaps (try (fst o HOLogic.dest_eq))
   517       |> perhaps (try HOLogic.dest_not) |> perhaps (try (fst o HOLogic.dest_eq))
   514       |> head_of;
   518       |> head_of;
   515 
   519 
   516     val maybe_rhs = imp_rhs |> perhaps (try (HOLogic.dest_not)) |> try (snd o HOLogic.dest_eq);
   520     val maybe_rhs = imp_rhs |> perhaps (try (HOLogic.dest_not)) |> try (snd o HOLogic.dest_eq);
   517 
   521 
   518     val fun_names = map fst fun_name_corec_spec_list;
   522     val discs = maps #ctr_specs corec_specs |> map #disc;
   519     val discs = maps (#ctr_specs o snd) fun_name_corec_spec_list |> map #disc;
   523     val sels = maps #ctr_specs corec_specs |> maps #sels;
   520     val sels = maps (#ctr_specs o snd) fun_name_corec_spec_list |> maps #sels;
   524     val ctrs = maps #ctr_specs corec_specs |> map #ctr;
   521     val ctrs = maps (#ctr_specs o snd) fun_name_corec_spec_list |> map #ctr;
       
   522   in
   525   in
   523     if member (op =) discs head orelse
   526     if member (op =) discs head orelse
   524       is_some maybe_rhs andalso
   527       is_some maybe_rhs andalso
   525         member (op =) (filter (null o binder_types o fastype_of) ctrs) (the maybe_rhs) then
   528         member (op =) (filter (null o binder_types o fastype_of) ctrs) (the maybe_rhs) then
   526       co_dissect_eqn_disc sequential fun_name_corec_spec_list eqn' imp_lhs' imp_rhs matched_conds
   529       co_dissect_eqn_disc sequential fun_names corec_specs imp_prems imp_rhs matchedss
   527       |>> single
   530       |>> single
   528     else if member (op =) sels head then
   531     else if member (op =) sels head then
   529       ([co_dissect_eqn_sel fun_name_corec_spec_list eqn' imp_rhs], matched_conds)
   532       ([co_dissect_eqn_sel fun_names corec_specs eqn' imp_rhs], matchedss)
   530     else if is_Free head andalso member (op =) fun_names (fst (dest_Free head)) then
   533     else if is_Free head andalso member (op =) fun_names (fst (dest_Free head)) then
   531       co_dissect_eqn_ctr sequential fun_name_corec_spec_list eqn' imp_lhs' imp_rhs matched_conds
   534       co_dissect_eqn_ctr sequential fun_names corec_specs eqn' imp_prems imp_rhs matchedss
   532     else
   535     else
   533       primrec_error_eqn "malformed function equation" eqn
   536       primrec_error_eqn "malformed function equation" eqn
   534   end;
   537   end;
   535 
   538 
   536 fun build_corec_args_discs disc_eqns ctr_specs =
   539 fun build_corec_arg_disc ctr_specs {fun_args, ctr_no, prems, ...} =
   537   if null disc_eqns then I else
   540   if is_none (#pred (nth ctr_specs ctr_no)) then I else
   538     let
   541     mk_conjs prems
   539 (*val _ = tracing ("d/p:\<cdot> " ^ space_implode "\n    \<cdot> " (map ((op ^) o
   542     |> curry subst_bounds (List.rev fun_args)
   540  apfst (Syntax.string_of_term @{context}) o apsnd (curry (op ^) " / " o @{make_string}))
   543     |> HOLogic.tupled_lambda (HOLogic.mk_tuple fun_args)
   541   (ctr_specs |> map_filter (fn {disc, pred = SOME pred, ...} => SOME (disc, pred) | _ => NONE))));*)
   544     |> K |> nth_map (the (#pred (nth ctr_specs ctr_no)));
   542       val conds = map #cond disc_eqns;
       
   543       val fun_args = #fun_args (hd disc_eqns);
       
   544       val args =
       
   545         if length ctr_specs = 1 then []
       
   546         else if length disc_eqns = length ctr_specs then
       
   547           fst (split_last conds)
       
   548         else if length disc_eqns = length ctr_specs - 1 then
       
   549           let val n = 0 upto length ctr_specs - 1
       
   550               |> the(*###*) o find_first (fn idx => not (exists (equal idx o #ctr_no) disc_eqns)) in
       
   551             if n = length ctr_specs - 1 then
       
   552               conds
       
   553             else
       
   554               split_last conds
       
   555               ||> HOLogic.mk_not
       
   556               |> `(uncurry (fold (curry HOLogic.mk_conj o HOLogic.mk_not)))
       
   557               ||> chop n o fst
       
   558               |> (fn (x, (l, r)) => l @ (x :: r))
       
   559           end
       
   560         else
       
   561           0 upto length ctr_specs - 1
       
   562           |> map (fn idx => find_first (equal idx o #ctr_no) disc_eqns
       
   563             |> Option.map #cond
       
   564             |> the_default undef_const)
       
   565           |> fst o split_last;
       
   566     in
       
   567       (* FIXME deal with #preds above *)
       
   568       (map_filter #pred ctr_specs, args)
       
   569       |-> fold2 (fn idx => fn t => nth_map idx
       
   570         (K (subst_bounds (List.rev fun_args, t)
       
   571           |> HOLogic.tupled_lambda (HOLogic.mk_tuple fun_args))))
       
   572     end;
       
   573 
   545 
   574 fun build_corec_arg_no_call sel_eqns sel = find_first (equal sel o #sel) sel_eqns
   546 fun build_corec_arg_no_call sel_eqns sel = find_first (equal sel o #sel) sel_eqns
   575   |> try (fn SOME sel_eqn => (#fun_args sel_eqn, #rhs_term sel_eqn))
   547   |> try (fn SOME sel_eqn => (#fun_args sel_eqn, #rhs_term sel_eqn))
   576   |> the_default ([], undef_const)
   548   |> the_default ([], undef_const)
   577   |-> abs_tuple
   549   |-> abs_tuple
   578   |> K;
   550   |> K;
   579 
   551 
   580 fun build_corec_arg_direct_call lthy has_call sel_eqns sel =
   552 fun build_corec_arg_direct_call lthy has_call sel_eqns sel =
   581   let
   553   let
   582     val maybe_sel_eqn = find_first (equal sel o #sel) sel_eqns;
   554     val maybe_sel_eqn = find_first (equal sel o #sel) sel_eqns;
   583     fun rewrite is_end U T t =
   555     fun rewrite is_end U _ t =
   584       if U = @{typ bool} then @{term True} |> has_call t ? K @{term False} (* stop? *)
   556       if U = @{typ bool} then @{term True} |> has_call t ? K @{term False} (* stop? *)
   585       else if is_end = has_call t then undef_const
   557       else if is_end = has_call t then undef_const
   586       else if is_end then t (* end *)
   558       else if is_end then t (* end *)
   587       else HOLogic.mk_tuple (snd (strip_comb t)); (* continue *)
   559       else HOLogic.mk_tuple (snd (strip_comb t)); (* continue *)
   588     fun massage rhs_term is_end t = massage_direct_corec_call
   560     fun massage rhs_term is_end t = massage_direct_corec_call
   624   let val sel_eqns = filter (equal (#ctr ctr_spec) o #ctr) all_sel_eqns in
   596   let val sel_eqns = filter (equal (#ctr ctr_spec) o #ctr) all_sel_eqns in
   625     if null sel_eqns then I else
   597     if null sel_eqns then I else
   626       let
   598       let
   627         val sel_call_list = #sels ctr_spec ~~ #calls ctr_spec;
   599         val sel_call_list = #sels ctr_spec ~~ #calls ctr_spec;
   628 
   600 
   629 (*val _ = tracing ("s/c:\<cdot> " ^ space_implode "\n    \<cdot> " (map ((op ^) o
       
   630  apfst (Syntax.string_of_term lthy) o apsnd (curry (op ^) " / " o @{make_string}))
       
   631   sel_call_list));*)
       
   632 
       
   633         val no_calls' = map_filter (try (apsnd (fn No_Corec n => n))) sel_call_list;
   601         val no_calls' = map_filter (try (apsnd (fn No_Corec n => n))) sel_call_list;
   634         val direct_calls' = map_filter (try (apsnd (fn Direct_Corec n => n))) sel_call_list;
   602         val direct_calls' = map_filter (try (apsnd (fn Direct_Corec n => n))) sel_call_list;
   635         val indirect_calls' = map_filter (try (apsnd (fn Indirect_Corec n => n))) sel_call_list;
   603         val indirect_calls' = map_filter (try (apsnd (fn Indirect_Corec n => n))) sel_call_list;
   636       in
   604       in
   637         I
   605         I
   643         #> fold (fn (sel, n) => nth_map n
   611         #> fold (fn (sel, n) => nth_map n
   644           (build_corec_arg_indirect_call lthy has_call sel_eqns sel)) indirect_calls'
   612           (build_corec_arg_indirect_call lthy has_call sel_eqns sel)) indirect_calls'
   645       end
   613       end
   646   end;
   614   end;
   647 
   615 
   648 fun co_build_defs lthy sequential bs mxs has_call arg_Tss fun_name_corec_spec_list eqns_data =
   616 fun mk_real_disc_eqns ctr_specs disc_eqns =
   649   let
   617   let
   650     val fun_names = map Binding.name_of bs;
   618     val real_disc_eqns =
   651 
   619       if length disc_eqns = 0 then disc_eqns
   652     val disc_eqnss = map_filter (try (fn Disc x => x)) eqns_data
   620       else if length disc_eqns = length ctr_specs - 1 then
   653       |> partition_eq ((op =) o pairself #fun_name)
   621         let
   654       |> finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names |> fst
   622           val n = 0 upto length ctr_specs
   655       |> map (sort ((op <) o pairself #ctr_no |> make_ord) o flat o snd);
   623             |> the(*###*) o find_first (fn idx => not (exists (equal idx o #ctr_no) disc_eqns));
   656 
   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 =
       
   639   let
   657     val _ = disc_eqnss |> map (fn x =>
   640     val _ = disc_eqnss |> map (fn x =>
   658       let val d = duplicates ((op =) o pairself #ctr_no) x in null d orelse
   641       let val d = duplicates ((op =) o pairself #ctr_no) x in null d orelse
   659         primrec_error_eqns "excess discriminator equations in definition"
   642         primrec_error_eqns "excess discriminator equations in definition"
   660           (maps (fn t => filter (equal (#ctr_no t) o #ctr_no) x) d |> map #user_eqn) end);
   643           (maps (fn t => filter (equal (#ctr_no t) o #ctr_no) x) d |> map #user_eqn) end);
   661 
   644     val corec_specs' = take (length bs) corec_specs;
   662 (*val _ = tracing ("disc_eqnss:\n    \<cdot> " ^ space_implode "\n    \<cdot> " (map @{make_string} disc_eqnss));*)
   645     val corecs = map #corec corec_specs';
   663 
   646     val ctr_specss = map #ctr_specs corec_specs';
   664     val sel_eqnss = map_filter (try (fn Sel x => x)) eqns_data
   647     val real_disc_eqnss = map2 mk_real_disc_eqns ctr_specss disc_eqnss;
   665       |> partition_eq ((op =) o pairself #fun_name)
       
   666       |> finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names |> fst
       
   667       |> map (flat o snd);
       
   668 
       
   669 (*val _ = tracing ("sel_eqnss:\n    \<cdot> " ^ space_implode "\n    \<cdot> " (map @{make_string} sel_eqnss));*)
       
   670 
       
   671     val corecs = map (#corec o snd) fun_name_corec_spec_list;
       
   672     val ctr_specss = map (#ctr_specs o snd) fun_name_corec_spec_list;
       
   673     val corec_args = hd corecs
   648     val corec_args = hd corecs
   674       |> fst o split_last o binder_types o fastype_of
   649       |> fst o split_last o binder_types o fastype_of
   675       |> map (Const o pair @{const_name undefined})
   650       |> map (Const o pair @{const_name undefined})
   676       |> fold2 build_corec_args_discs disc_eqnss ctr_specss
   651       |> fold2 (fold o build_corec_arg_disc) ctr_specss real_disc_eqnss
   677       |> fold2 (fold o build_corec_args_sel lthy has_call) sel_eqnss ctr_specss;
   652       |> fold2 (fold o build_corec_args_sel lthy has_call) sel_eqnss ctr_specss;
   678 
       
   679     fun currys Ts t = if length Ts <= 1 then t else
   653     fun currys Ts t = if length Ts <= 1 then t else
   680       t $ foldr1 (fn (u, v) => HOLogic.pair_const dummyT dummyT $ u $ v)
   654       t $ foldr1 (fn (u, v) => HOLogic.pair_const dummyT dummyT $ u $ v)
   681         (length Ts - 1 downto 0 |> map Bound)
   655         (length Ts - 1 downto 0 |> map Bound)
   682       |> fold_rev (fn T => fn u => Abs (Name.uu, T, u)) Ts;
   656       |> fold_rev (fn T => fn u => Abs (Name.uu, T, u)) Ts;
   683 
   657 
   684 val _ = tracing ("corecursor arguments:\n    \<cdot> " ^
   658 val _ = tracing ("corecursor arguments:\n    \<cdot> " ^
   685  space_implode "\n    \<cdot> " (map (Syntax.string_of_term lthy) corec_args));
   659  space_implode "\n    \<cdot> " (map (Syntax.string_of_term lthy) corec_args));
   686 
   660 
   687     fun uneq_pairs_rev xs = xs (* FIXME \<exists>? *)
   661     val exclss' =
   688       |> these o try (split_last #> (fn (ys, y) => uneq_pairs_rev ys @ map (pair y) ys));
   662       real_disc_eqnss
   689     val proof_obligations = if sequential then [] else
   663       |> map (map (fn {fun_args, ctr_no, prems, ...} => (fun_args, ctr_no, prems))
   690       disc_eqnss
   664         #> fst o (fn xs => fold_map (fn x => fn ys => ((x, ys), ys @ [x])) xs [])
   691       |> maps (uneq_pairs_rev o map (fn {fun_args, cond, ...} => (fun_args, cond)))
   665         #> maps (uncurry (map o pair)
   692       |> map (fn ((fun_args, x), (_, y)) => [x, HOLogic.mk_not y]
   666           #> map (fn ((fun_args, c, x), (_, c', y)) => ((c, c'), (x, mk_not (mk_conjs y)))
   693         |> map (HOLogic.mk_Trueprop o curry subst_bounds (List.rev fun_args))
   667             ||> apfst (map HOLogic.mk_Trueprop) o apsnd HOLogic.mk_Trueprop
   694         |> curry list_comb @{const ==>});
   668             ||> Logic.list_implies
   695 
   669             ||> curry Logic.list_all (map dest_Free fun_args))))
   696 val _ = tracing ("proof obligations:\n    \<cdot> " ^
       
   697  space_implode "\n    \<cdot> " (map (Syntax.string_of_term lthy) proof_obligations));
       
   698 
       
   699   in
   670   in
   700     map (list_comb o rpair corec_args) corecs
   671     map (list_comb o rpair corec_args) corecs
   701     |> map2 (fn Ts => fn t => if length Ts = 0 then t $ HOLogic.unit else t) arg_Tss
   672     |> map2 (fn Ts => fn t => if length Ts = 0 then t $ HOLogic.unit else t) arg_Tss
   702     |> map2 currys arg_Tss
   673     |> map2 currys arg_Tss
   703     |> Syntax.check_terms lthy
   674     |> Syntax.check_terms lthy
   704     |> map3 (fn b => fn mx => fn t => ((b, mx), ((Binding.map_name Thm.def_name b, []), t))) bs mxs
   675     |> map3 (fn b => fn mx => fn t => ((b, mx), ((Binding.map_name Thm.def_name b, []), t))) bs mxs
   705     |> rpair proof_obligations
   676     |> rpair exclss'
   706   end;
   677   end;
   707 
   678 
   708 fun add_primcorec sequential fixes specs lthy =
   679 fun add_primcorec sequential fixes specs lthy =
   709   let
   680   let
   710     val (bs, mxs) = map_split (apfst fst) fixes;
   681     val (bs, mxs) = map_split (apfst fst) fixes;
   715       |> map_index (fn (i, v) => if exists_subterm (equal v) t then SOME i else NONE)
   686       |> map_index (fn (i, v) => if exists_subterm (equal v) t then SOME i else NONE)
   716       |> map_filter I;
   687       |> map_filter I;
   717 
   688 
   718     val callssss = []; (* FIXME *)
   689     val callssss = []; (* FIXME *)
   719 
   690 
   720     val ((nontriv, corec_specs, _, coinduct_thm, strong_co_induct_thm, coinduct_thmss,
   691     val ((nontriv, corec_specs', _, coinduct_thm, strong_co_induct_thm, coinduct_thmss,
   721           strong_coinduct_thmss), lthy') =
   692           strong_coinduct_thmss), lthy') =
   722       corec_specs_of bs arg_Ts res_Ts get_indices callssss lthy;
   693       corec_specs_of bs arg_Ts res_Ts get_indices callssss lthy;
   723 
   694 
   724     val fun_names = map Binding.name_of bs;
   695     val fun_names = map Binding.name_of bs;
   725 
   696     val corec_specs = take (length fun_names) corec_specs'; (*###*)
   726     val fun_name_corec_spec_list = (fun_names ~~ res_Ts, corec_specs)
       
   727       |> uncurry (finds (fn ((_, T), {corec, ...}) => T = body_type (fastype_of corec))) |> fst
       
   728       |> map (apfst fst #> apsnd the_single); (*###*)
       
   729 
   697 
   730     val (eqns_data, _) =
   698     val (eqns_data, _) =
   731       fold_map (co_dissect_eqn sequential fun_name_corec_spec_list) (map snd specs) []
   699       fold_map (co_dissect_eqn sequential fun_names corec_specs) (map snd specs) []
   732       |>> flat;
   700       |>> flat;
       
   701 
       
   702     val disc_eqnss = map_filter (try (fn Disc x => x)) eqns_data
       
   703       |> partition_eq ((op =) o pairself #fun_name)
       
   704       |> finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names |> fst
       
   705       |> map (sort ((op <) o pairself #ctr_no |> make_ord) o flat o snd);
       
   706 
       
   707     val sel_eqnss = map_filter (try (fn Sel x => x)) eqns_data
       
   708       |> partition_eq ((op =) o pairself #fun_name)
       
   709       |> finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names |> fst
       
   710       |> map (flat o snd);
   733 
   711 
   734     val has_call = exists_subterm (map (fst #>> Binding.name_of #> Free) fixes |> member (op =));
   712     val has_call = exists_subterm (map (fst #>> Binding.name_of #> Free) fixes |> member (op =));
   735     val arg_Tss = map (binder_types o snd o fst) fixes;
   713     val arg_Tss = map (binder_types o snd o fst) fixes;
   736     val (defs, proof_obligations) =
   714     val (defs, exclss') =
   737       co_build_defs lthy' sequential bs mxs has_call arg_Tss fun_name_corec_spec_list eqns_data;
   715       co_build_defs lthy' bs mxs has_call arg_Tss corec_specs disc_eqnss sel_eqnss;
       
   716 
       
   717     (* try to prove (automatically generated) tautologies by ourselves *)
       
   718     val exclss'' = exclss'
       
   719       |> map (map (apsnd
       
   720         (`(try (fn t => Goal.prove lthy [] [] t (mk_primcorec_taut_tac lthy |> K))))));
       
   721     val taut_thmss = map (map (apsnd (the o fst)) o filter (is_some o fst o snd)) exclss'';
       
   722     val (obligation_idxss, obligationss) = exclss''
       
   723       |> map (map (apsnd (rpair [] o snd)) o filter (is_none o fst o snd))
       
   724       |> split_list o map split_list;
       
   725 
       
   726     fun prove thmss' def_thms' lthy =
       
   727       let
       
   728         val def_thms = map (snd o snd) def_thms';
       
   729 
       
   730         val exclss' = map (op ~~) (obligation_idxss ~~ thmss');
       
   731         fun mk_exclsss excls n =
       
   732           (excls, map (fn k => replicate k [TrueI] @ replicate (n - k) []) (0 upto n - 1))
       
   733           |-> 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)
       
   735           |-> map2 (fn excls => fn (_, {ctr_specs, ...}) => mk_exclsss excls (length ctr_specs));
       
   736 
       
   737         fun prove_disc {ctr_specs, ...} exclsss
       
   738             {fun_name, fun_args, ctr_no, prems, ...} =
       
   739           let
       
   740             val disc_corec = nth ctr_specs ctr_no |> #disc_corec;
       
   741             val k = 1 + ctr_no;
       
   742             val m = length prems;
       
   743             val t =
       
   744               (* FIXME use applied_fun from dissect_\<dots> instead? *)
       
   745               list_comb (Free (fun_name, dummyT), map Bound (length fun_args - 1 downto 0))
       
   746               |> curry betapply (#disc (nth ctr_specs ctr_no)) (*###*)
       
   747               |> HOLogic.mk_Trueprop
       
   748               |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
       
   749               |> 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
       
   775             mk_primcorec_eq_tac lthy def_thms sel_corec k m exclsss
       
   776               nested_maps nested_map_idents nested_map_comps
       
   777             |> K |> Goal.prove lthy [] [] t
       
   778           end;
       
   779 
       
   780         val disc_notes =
       
   781           fun_names ~~
       
   782             map3 (map oo prove_disc) (take (length disc_eqnss) corec_specs) exclssss disc_eqnss
       
   783           |> map (fn (fun_name, thms) =>
       
   784             ((Binding.qualify true fun_name (@{binding disc}), 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
       
   791         lthy |> snd o Local_Theory.notes (disc_notes @ sel_notes)
       
   792       end;
   738   in
   793   in
   739     lthy'
   794     lthy'
   740     |> fold_map Local_Theory.define defs |> snd
   795     |> Proof.theorem NONE (curry (op #->) (fold_map Local_Theory.define defs) o prove) obligationss
   741     |> Proof.theorem NONE (K I) [map (rpair []) proof_obligations]
       
   742     |> Proof.refine (Method.primitive_text I)
   796     |> Proof.refine (Method.primitive_text I)
   743     |> Seq.hd
   797     |> Seq.hd
   744   end
   798   end
   745 
   799 
   746 fun add_primcorec_cmd seq (raw_fixes, raw_specs) lthy =
   800 fun add_primcorec_cmd seq (raw_fixes, raw_specs) lthy =