src/HOL/Tools/BNF/bnf_fp_n2m_sugar.ML
changeset 57397 5004aca20821
parent 56650 1f9ab71d43a5
child 57489 8f0ba9f2d10f
equal deleted inserted replaced
57396:42eede5610a9 57397:5004aca20821
   235         val reps = map #rep absT_infos;
   235         val reps = map #rep absT_infos;
   236         val absTs = map #absT absT_infos;
   236         val absTs = map #absT absT_infos;
   237         val repTs = map #repT absT_infos;
   237         val repTs = map #repT absT_infos;
   238         val abs_inverses = map #abs_inverse absT_infos;
   238         val abs_inverses = map #abs_inverse absT_infos;
   239 
   239 
   240         val nesting_bnfs = nesty_bnfs lthy ctrXs_Tsss As;
   240         val fp_nesting_bnfs = nesting_bnfs lthy ctrXs_Tsss Xs;
   241         val nested_bnfs = nesty_bnfs lthy ctrXs_Tsss Xs;
   241         val live_nesting_bnfs = nesting_bnfs lthy ctrXs_Tsss As;
   242 
   242 
   243         val ((xtor_co_recs, recs_args_types, corecs_args_types), _) =
   243         val ((xtor_co_recs, recs_args_types, corecs_args_types), _) =
   244           mk_co_recs_prelims fp ctr_Tsss fpTs Cs absTs repTs ns mss xtor_co_recs0 lthy;
   244           mk_co_recs_prelims fp ctr_Tsss fpTs Cs absTs repTs ns mss xtor_co_recs0 lthy;
   245 
   245 
   246         fun mk_binding b suf = Binding.suffix_name ("_" ^ suf) b;
   246         fun mk_binding b suf = Binding.suffix_name ("_" ^ suf) b;
   254 
   254 
   255         val ((common_co_inducts, co_inductss, co_rec_thmss, disc_corec_thmss, sel_corec_thmsss),
   255         val ((common_co_inducts, co_inductss, co_rec_thmss, disc_corec_thmss, sel_corec_thmsss),
   256              fp_sugar_thms) =
   256              fp_sugar_thms) =
   257           if fp = Least_FP then
   257           if fp = Least_FP then
   258             derive_induct_recs_thms_for_types pre_bnfs recs_args_types xtor_co_induct
   258             derive_induct_recs_thms_for_types pre_bnfs recs_args_types xtor_co_induct
   259               xtor_co_rec_thms nesting_bnfs nested_bnfs fpTs Cs Xs ctrXs_Tsss fp_abs_inverses
   259               xtor_co_rec_thms live_nesting_bnfs fp_nesting_bnfs fpTs Cs Xs ctrXs_Tsss
   260               fp_type_definitions abs_inverses ctrss ctr_defss co_recs co_rec_defs lthy
   260               fp_abs_inverses fp_type_definitions abs_inverses ctrss ctr_defss co_recs co_rec_defs
       
   261               lthy
   261             |> `(fn ((inducts, induct, _), (rec_thmss, _)) =>
   262             |> `(fn ((inducts, induct, _), (rec_thmss, _)) =>
   262               ([induct], [inducts], rec_thmss, replicate nn [], replicate nn []))
   263               ([induct], [inducts], rec_thmss, replicate nn [], replicate nn []))
   263             ||> (fn info => (SOME info, NONE))
   264             ||> (fn info => (SOME info, NONE))
   264           else
   265           else
   265             derive_coinduct_corecs_thms_for_types pre_bnfs (the corecs_args_types) xtor_co_induct
   266             derive_coinduct_corecs_thms_for_types pre_bnfs (the corecs_args_types) xtor_co_induct
   266               dtor_injects dtor_ctors xtor_co_rec_thms nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss mss ns
   267               dtor_injects dtor_ctors xtor_co_rec_thms live_nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss
   267               fp_abs_inverses abs_inverses (fn thm => thm RS @{thm vimage2p_refl}) ctr_defss
   268               mss ns fp_abs_inverses abs_inverses (fn thm => thm RS @{thm vimage2p_refl}) ctr_defss
   268               ctr_sugars co_recs co_rec_defs (Proof_Context.export lthy no_defs_lthy) lthy
   269               ctr_sugars co_recs co_rec_defs (Proof_Context.export lthy no_defs_lthy) lthy
   269             |> `(fn ((coinduct_thms_pairs, _), (corec_thmss, _), (disc_corec_thmss, _), _,
   270             |> `(fn ((coinduct_thms_pairs, _), (corec_thmss, _), (disc_corec_thmss, _), _,
   270                      (sel_corec_thmsss, _)) =>
   271                      (sel_corec_thmsss, _)) =>
   271               (map snd coinduct_thms_pairs, map fst coinduct_thms_pairs, corec_thmss,
   272               (map snd coinduct_thms_pairs, map fst coinduct_thms_pairs, corec_thmss,
   272                disc_corec_thmss, sel_corec_thmsss))
   273                disc_corec_thmss, sel_corec_thmsss))
   276 
   277 
   277         fun mk_target_fp_sugar T X kk pre_bnf absT_info ctrXs_Tss ctr_defs ctr_sugar co_rec
   278         fun mk_target_fp_sugar T X kk pre_bnf absT_info ctrXs_Tss ctr_defs ctr_sugar co_rec
   278             co_rec_def maps co_inducts co_rec_thms disc_corec_thms sel_corec_thmss
   279             co_rec_def maps co_inducts co_rec_thms disc_corec_thms sel_corec_thmss
   279             ({rel_injects, rel_distincts, ...} : fp_sugar) =
   280             ({rel_injects, rel_distincts, ...} : fp_sugar) =
   280           {T = T, BT = Term.dummyT (*wrong but harmless*), X = X, fp = fp, fp_res = fp_res,
   281           {T = T, BT = Term.dummyT (*wrong but harmless*), X = X, fp = fp, fp_res = fp_res,
   281            fp_res_index = kk, pre_bnf = pre_bnf, absT_info = absT_info, nested_bnfs = nested_bnfs,
   282            fp_res_index = kk, pre_bnf = pre_bnf, absT_info = absT_info,
   282            nesting_bnfs = nesting_bnfs, ctrXs_Tss = ctrXs_Tss, ctr_defs = ctr_defs,
   283            fp_nesting_bnfs = fp_nesting_bnfs, live_nesting_bnfs = live_nesting_bnfs,
   283            ctr_sugar = ctr_sugar, co_rec = co_rec, co_rec_def = co_rec_def, maps = maps,
   284            ctrXs_Tss = ctrXs_Tss, ctr_defs = ctr_defs, ctr_sugar = ctr_sugar, co_rec = co_rec,
   284            common_co_inducts = common_co_inducts, co_inducts = co_inducts,
   285            co_rec_def = co_rec_def, maps = maps, common_co_inducts = common_co_inducts,
   285            co_rec_thms = co_rec_thms, disc_co_recs = disc_corec_thms,
   286            co_inducts = co_inducts, co_rec_thms = co_rec_thms, disc_co_recs = disc_corec_thms,
   286            sel_co_recss = sel_corec_thmss, rel_injects = rel_injects, rel_distincts = rel_distincts}
   287            sel_co_recss = sel_corec_thmss, rel_injects = rel_injects, rel_distincts = rel_distincts}
   287           |> morph_fp_sugar phi;
   288           |> morph_fp_sugar phi;
   288 
   289 
   289         val target_fp_sugars =
   290         val target_fp_sugars =
   290           map16 mk_target_fp_sugar fpTs Xs kks pre_bnfs absT_infos ctrXs_Tsss ctr_defss ctr_sugars
   291           map16 mk_target_fp_sugar fpTs Xs kks pre_bnfs absT_infos ctrXs_Tsss ctr_defss ctr_sugars