1(*  Title:      HOL/Tools/BNF/bnf_gfp_grec_sugar.ML
2    Author:     Aymeric Bouzy, Ecole polytechnique
3    Author:     Jasmin Blanchette, Inria, LORIA, MPII
4    Author:     Dmitriy Traytel, ETH Z��rich
5    Copyright   2015, 2016
6
7Generalized corecursor sugar ("corec" and friends).
8*)
9
10signature BNF_GFP_GREC_SUGAR =
11sig
12  datatype corec_option =
13    Plugins_Option of Proof.context -> Plugin_Name.filter |
14    Friend_Option |
15    Transfer_Option
16
17  val parse_corec_equation: Proof.context -> term list -> term -> term list * term
18  val explore_corec_equation: Proof.context -> bool -> bool -> string -> term ->
19    BNF_GFP_Grec_Sugar_Util.s_parse_info -> typ -> term list * term -> term list * term
20  val build_corecUU_arg_and_goals: bool -> term -> term list * term -> local_theory ->
21    (((thm list * thm list * thm list) * term list) * term) * local_theory
22  val derive_eq_corecUU: Proof.context -> BNF_GFP_Grec.corec_info -> term -> term -> thm -> thm
23  val derive_unique: Proof.context -> morphism -> term -> BNF_GFP_Grec.corec_info -> string ->
24    thm -> thm
25
26  val corec_cmd: bool -> corec_option list -> (binding * string option * mixfix) list * string ->
27    local_theory -> local_theory
28  val corecursive_cmd: bool -> corec_option list ->
29    (binding * string option * mixfix) list * string -> local_theory -> Proof.state
30  val friend_of_corec_cmd: (string * string option) * string -> local_theory -> Proof.state
31  val coinduction_upto_cmd: string * string -> local_theory -> local_theory
32end;
33
34structure BNF_GFP_Grec_Sugar : BNF_GFP_GREC_SUGAR =
35struct
36
37open Ctr_Sugar
38open BNF_Util
39open BNF_Tactics
40open BNF_Def
41open BNF_Comp
42open BNF_FP_Util
43open BNF_FP_Def_Sugar
44open BNF_FP_N2M_Sugar
45open BNF_GFP_Util
46open BNF_GFP_Rec_Sugar
47open BNF_FP_Rec_Sugar_Transfer
48open BNF_GFP_Grec
49open BNF_GFP_Grec_Sugar_Util
50open BNF_GFP_Grec_Sugar_Tactics
51
52val cong_N = "cong_";
53val baseN = "base";
54val reflN = "refl";
55val symN = "sym";
56val transN = "trans";
57val cong_introsN = prefix cong_N "intros";
58val codeN = "code";
59val coinductN = "coinduct";
60val coinduct_uptoN = "coinduct_upto";
61val corecN = "corec";
62val ctrN = "ctr";
63val discN = "disc";
64val disc_iffN = "disc_iff";
65val eq_algrhoN = "eq_algrho";
66val eq_corecUUN = "eq_corecUU";
67val friendN = "friend";
68val inner_elimN = "inner_elim";
69val inner_inductN = "inner_induct";
70val inner_simpN = "inner_simp";
71val rhoN = "rho";
72val selN = "sel";
73val uniqueN = "unique";
74
75val inner_fp_suffix = "_inner_fp";
76
77val nitpicksimp_attrs = @{attributes [nitpick_simp]};
78val simp_attrs = @{attributes [simp]};
79
80val unfold_id_thms1 =
81  map (fn thm => thm RS eq_reflection) @{thms id_bnf_o o_id_bnf id_apply o_apply} @
82  @{thms fst_def[abs_def, symmetric] snd_def[abs_def, symmetric]};
83
84fun unfold_id_bnf_etc lthy =
85  let val thy = Proof_Context.theory_of lthy in
86    Raw_Simplifier.rewrite_term thy unfold_id_thms1 []
87    #> Raw_Simplifier.rewrite_term thy @{thms BNF_Composition.id_bnf_def} []
88  end;
89
90datatype corec_option =
91  Plugins_Option of Proof.context -> Plugin_Name.filter |
92  Friend_Option |
93  Transfer_Option;
94
95val corec_option_parser = Parse.group (K "option")
96  (Plugin_Name.parse_filter >> Plugins_Option
97   || Parse.reserved "friend" >> K Friend_Option
98   || Parse.reserved "transfer" >> K Transfer_Option);
99
100type codatatype_extra =
101  {case_dtor: thm,
102   case_trivial: thm,
103   abs_rep_transfers: thm list};
104
105fun morph_codatatype_extra phi ({case_dtor, case_trivial, abs_rep_transfers} : codatatype_extra) =
106  {case_dtor = Morphism.thm phi case_dtor, case_trivial = Morphism.thm phi case_trivial,
107   abs_rep_transfers = map (Morphism.thm phi) abs_rep_transfers};
108
109val transfer_codatatype_extra = morph_codatatype_extra o Morphism.transfer_morphism;
110
111type coinduct_extra =
112  {coinduct: thm,
113   coinduct_attrs: Token.src list,
114   cong_intro_pairs: (string * thm) list};
115
116fun morph_coinduct_extra phi ({coinduct, coinduct_attrs, cong_intro_pairs} : coinduct_extra) =
117  {coinduct = Morphism.thm phi coinduct, coinduct_attrs = coinduct_attrs,
118   cong_intro_pairs = map (apsnd (Morphism.thm phi)) cong_intro_pairs};
119
120val transfer_coinduct_extra = morph_coinduct_extra o Morphism.transfer_morphism;
121
122type friend_extra =
123  {eq_algrhos: thm list,
124   algrho_eqs: thm list};
125
126val empty_friend_extra = {eq_algrhos = [], algrho_eqs = []};
127
128fun merge_friend_extras ({eq_algrhos = eq_algrhos1, algrho_eqs = algrho_eqs1},
129    {eq_algrhos = eq_algrhos2, algrho_eqs = algrho_eqs2}) =
130  {eq_algrhos = union Thm.eq_thm_prop eq_algrhos1 eq_algrhos2,
131   algrho_eqs = union Thm.eq_thm_prop algrho_eqs1 algrho_eqs2};
132
133type corec_sugar_data =
134  codatatype_extra Symtab.table * coinduct_extra Symtab.table * friend_extra Symtab.table;
135
136structure Data = Generic_Data
137(
138  type T = corec_sugar_data;
139  val empty = (Symtab.empty, Symtab.empty, Symtab.empty);
140  val extend = I;
141  fun merge data : T =
142    (Symtab.merge (K true) (apply2 #1 data), Symtab.merge (K true) (apply2 #2 data),
143     Symtab.join (K merge_friend_extras) (apply2 #3 data));
144);
145
146fun register_codatatype_extra fpT_name extra =
147  Local_Theory.declaration {syntax = false, pervasive = true} (fn phi =>
148    Data.map (@{apply 3(1)} (Symtab.update (fpT_name, morph_codatatype_extra phi extra))));
149
150fun codatatype_extra_of ctxt =
151  Symtab.lookup (#1 (Data.get (Context.Proof ctxt)))
152  #> Option.map (transfer_codatatype_extra (Proof_Context.theory_of ctxt));
153
154fun all_codatatype_extras_of ctxt =
155  Symtab.dest (#1 (Data.get (Context.Proof ctxt)));
156
157fun register_coinduct_extra fpT_name extra =
158  Local_Theory.declaration {syntax = false, pervasive = true} (fn phi =>
159    Data.map (@{apply 3(2)} (Symtab.update (fpT_name, morph_coinduct_extra phi extra))));
160
161fun coinduct_extra_of ctxt =
162  Symtab.lookup (#2 (Data.get (Context.Proof ctxt)))
163  #> Option.map (transfer_coinduct_extra (Proof_Context.theory_of ctxt));
164
165fun register_friend_extra fun_name eq_algrho algrho_eq =
166  Local_Theory.declaration {syntax = false, pervasive = true} (fn phi =>
167    Data.map (@{apply 3(3)} (Symtab.map_default (fun_name, empty_friend_extra)
168      (fn {eq_algrhos, algrho_eqs} =>
169        {eq_algrhos = Morphism.thm phi eq_algrho :: eq_algrhos,
170         algrho_eqs = Morphism.thm phi algrho_eq :: algrho_eqs}))));
171
172fun all_friend_extras_of ctxt =
173  Symtab.dest (#3 (Data.get (Context.Proof ctxt)));
174
175fun coinduct_extras_of_generic context =
176  corec_infos_of_generic context
177  #> map (#corecUU #> dest_Const #> fst #> Symtab.lookup (#2 (Data.get context)) #> the
178    #> transfer_coinduct_extra (Context.theory_of context));
179
180fun get_coinduct_uptos fpT_name context =
181  coinduct_extras_of_generic context fpT_name |> map #coinduct;
182fun get_cong_all_intros fpT_name context =
183  coinduct_extras_of_generic context fpT_name |> maps (#cong_intro_pairs #> map snd);
184fun get_cong_intros fpT_name name context =
185  coinduct_extras_of_generic context fpT_name
186  |> map_filter (#cong_intro_pairs #> (fn ps => AList.lookup (op =) ps name));
187
188fun ctr_names_of_fp_name lthy fpT_name =
189  fpT_name |> fp_sugar_of lthy |> the |> #fp_ctr_sugar |> #ctr_sugar |> #ctrs
190  |> map (Long_Name.base_name o name_of_ctr);
191
192fun register_coinduct_dynamic_base fpT_name lthy =
193  let val fp_b = Binding.name (Long_Name.base_name fpT_name) in
194    lthy
195    |> fold Local_Theory.add_thms_dynamic
196      ((mk_fp_binding fp_b coinduct_uptoN, get_coinduct_uptos fpT_name) ::
197        map (fn N =>
198          let val N = cong_N ^ N in
199            (mk_fp_binding fp_b N, get_cong_intros fpT_name N)
200          end)
201        ([baseN, reflN, symN, transN] @ ctr_names_of_fp_name lthy fpT_name))
202    |> Local_Theory.add_thms_dynamic
203      (mk_fp_binding fp_b cong_introsN, get_cong_all_intros fpT_name)
204  end;
205
206fun register_coinduct_dynamic_friend fpT_name friend_name =
207  let
208    val fp_b = Binding.name (Long_Name.base_name fpT_name);
209    val friend_base_name = cong_N ^ Long_Name.base_name friend_name;
210  in
211    Local_Theory.add_thms_dynamic
212      (mk_fp_binding fp_b friend_base_name, get_cong_intros fpT_name friend_base_name)
213  end;
214
215fun derive_case_dtor ctxt fpT_name =
216  let
217    val thy = Proof_Context.theory_of ctxt;
218
219    val SOME ({fp_res_index, fp_res = {dtors = dtors0, dtor_ctors, ...},
220        absT_info = {rep = rep0, abs_inverse, ...},
221        fp_ctr_sugar = {ctr_defs, ctr_sugar = {casex, exhaust, case_thms, ...}, ...}, ...}) =
222      fp_sugar_of ctxt fpT_name;
223
224    val (f_Ts, Type (_, [fpT as Type (_, As), _])) = strip_fun_type (fastype_of casex);
225    val x_Tss = map binder_types f_Ts;
226
227    val (((u, fs), xss), _) = ctxt
228      |> yield_singleton (mk_Frees "y") fpT
229      ||>> mk_Frees "f" f_Ts
230      ||>> mk_Freess "x" x_Tss;
231
232    val dtor0 = nth dtors0 fp_res_index;
233    val dtor = mk_dtor As dtor0;
234
235    val u' = dtor $ u;
236
237    val absT = fastype_of u';
238
239    val rep = mk_rep absT rep0;
240
241    val goal = mk_Trueprop_eq (list_comb (casex, fs) $ u,
242        mk_case_absumprod absT rep fs xss xss $ u')
243      |> Raw_Simplifier.rewrite_term thy @{thms comp_def[THEN eq_reflection]} [];
244    val vars = map (fst o dest_Free) (u :: fs);
245
246    val dtor_ctor = nth dtor_ctors fp_res_index;
247  in
248    Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} =>
249      mk_case_dtor_tac ctxt u abs_inverse dtor_ctor ctr_defs exhaust case_thms)
250    |> Thm.close_derivation \<^here>
251  end;
252
253fun derive_case_trivial ctxt fpT_name =
254  let
255    val SOME {casex, exhaust, case_thms, ...} = ctr_sugar_of ctxt fpT_name;
256
257    val Type (_, As0) = domain_type (body_fun_type (fastype_of casex));
258
259    val (As, _) = ctxt
260      |> mk_TFrees' (map Type.sort_of_atyp As0);
261    val fpT = Type (fpT_name, As);
262
263    val (var_name, ()) = singleton (Variable.variant_frees ctxt []) ("x", ());
264    val var = Free (var_name, fpT);
265    val goal = mk_Trueprop_eq (expand_to_ctr_term ctxt fpT var, var);
266
267    val exhaust' = infer_instantiate' ctxt [SOME (Thm.cterm_of ctxt var)] exhaust;
268  in
269    Goal.prove_sorry ctxt [var_name] [] goal (fn {context = ctxt, prems = _} =>
270      HEADGOAL (rtac ctxt exhaust') THEN ALLGOALS (hyp_subst_tac ctxt) THEN
271      unfold_thms_tac ctxt case_thms THEN ALLGOALS (rtac ctxt refl))
272    |> Thm.close_derivation \<^here>
273  end;
274
275fun mk_abs_rep_transfers ctxt fpT_name =
276  [mk_abs_transfer ctxt fpT_name, mk_rep_transfer ctxt fpT_name]
277  handle Fail _ => [];
278
279fun ensure_codatatype_extra fpT_name ctxt =
280  (case codatatype_extra_of ctxt fpT_name of
281    NONE =>
282    let val abs_rep_transfers = mk_abs_rep_transfers ctxt fpT_name in
283      ctxt
284      |> register_codatatype_extra fpT_name
285        {case_dtor = derive_case_dtor ctxt fpT_name,
286         case_trivial = derive_case_trivial ctxt fpT_name,
287         abs_rep_transfers = abs_rep_transfers}
288      |> set_transfer_rule_attrs abs_rep_transfers
289    end
290  | SOME {abs_rep_transfers, ...} => ctxt |> set_transfer_rule_attrs abs_rep_transfers);
291
292fun setup_base fpT_name =
293  register_coinduct_dynamic_base fpT_name
294  #> ensure_codatatype_extra fpT_name;
295
296fun is_set ctxt (const_name, T) =
297  (case T of
298    Type (\<^type_name>\<open>fun\<close>, [Type (fpT_name, _), Type (\<^type_name>\<open>set\<close>, [_])]) =>
299    (case bnf_of ctxt fpT_name of
300      SOME bnf => exists (fn Const (s, _) => s = const_name | _ => false) (sets_of_bnf bnf)
301    | NONE => false)
302  | _ => false);
303
304fun case_eq_if_thms_of_term ctxt t =
305  let val ctr_sugars = map_filter (ctr_sugar_of_case ctxt o fst) (Term.add_consts t []) in
306    maps #case_eq_ifs ctr_sugars
307  end;
308
309fun all_algrho_eqs_of ctxt =
310  maps (#algrho_eqs o snd) (all_friend_extras_of ctxt);
311
312fun derive_code ctxt inner_fp_simps goal
313    {sig_fp_sugars, ssig_fp_sugar, eval, eval_simps, all_algLam_algs, corecUU_thm, ...} fun_t
314    fun_def =
315  let
316    val fun_T = fastype_of fun_t;
317    val (arg_Ts, Type (fpT_name, _)) = strip_type fun_T;
318    val num_args = length arg_Ts;
319
320    val SOME {pre_bnf, fp_bnf, absT_info, fp_nesting_bnfs, live_nesting_bnfs, fp_ctr_sugar, ...} =
321      fp_sugar_of ctxt fpT_name;
322    val SOME {case_trivial, ...} = codatatype_extra_of ctxt fpT_name;
323
324    val ctr_sugar = #ctr_sugar fp_ctr_sugar;
325    val pre_map_def = map_def_of_bnf pre_bnf;
326    val abs_inverse = #abs_inverse absT_info;
327    val ctr_defs = #ctr_defs fp_ctr_sugar;
328    val case_eq_ifs = #case_eq_ifs ctr_sugar @ case_eq_if_thms_of_term ctxt goal;
329    val all_sig_map_thms = maps (#map_thms o #fp_bnf_sugar) sig_fp_sugars;
330
331    val fp_map_ident = map_ident_of_bnf fp_bnf;
332    val fpsig_nesting_bnfs = fp_nesting_bnfs @ maps #live_nesting_bnfs sig_fp_sugars;
333    val fpsig_nesting_T_names = map (fst o dest_Type o T_of_bnf) fpsig_nesting_bnfs;
334    val fpsig_nesting_fp_sugars = map_filter (fp_sugar_of ctxt) fpsig_nesting_T_names;
335    val fpsig_nesting_fp_bnf_sugars = map #fp_bnf_sugar fpsig_nesting_fp_sugars;
336    val ssig_fp_bnf_sugar = #fp_bnf_sugar ssig_fp_sugar;
337    val ssig_bnf = #fp_bnf ssig_fp_sugar;
338    val ssig_map = map_of_bnf ssig_bnf;
339    val fpsig_nesting_maps = map map_of_bnf fpsig_nesting_bnfs;
340    val fpsig_nesting_map_ident0s = map map_ident0_of_bnf fpsig_nesting_bnfs;
341    val fpsig_nesting_map_comps = map map_comp_of_bnf fpsig_nesting_bnfs;
342    val fpsig_nesting_map_thms = maps #map_thms fpsig_nesting_fp_bnf_sugars;
343    val live_nesting_map_ident0s = map map_ident0_of_bnf live_nesting_bnfs;
344    val ssig_map_thms = #map_thms ssig_fp_bnf_sugar;
345    val all_algLam_alg_pointfuls = map (mk_pointful ctxt) all_algLam_algs;
346  in
347    Variable.add_free_names ctxt goal []
348    |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} =>
349      mk_code_tac ctxt num_args fpsig_nesting_maps ssig_map eval pre_map_def abs_inverse
350        fpsig_nesting_map_ident0s fpsig_nesting_map_comps fpsig_nesting_map_thms
351        live_nesting_map_ident0s fp_map_ident case_trivial ctr_defs case_eq_ifs corecUU_thm
352        all_sig_map_thms ssig_map_thms all_algLam_alg_pointfuls (all_algrho_eqs_of ctxt) eval_simps
353        inner_fp_simps fun_def))
354    |> Thm.close_derivation \<^here>
355  end;
356
357fun derive_unique ctxt phi code_goal
358    {sig_fp_sugars, ssig_fp_sugar, eval, eval_simps, all_algLam_algs, corecUU_unique, ...} fpT_name
359    eq_corecUU =
360  let
361    val SOME {pre_bnf, fp_bnf, absT_info, fp_nesting_bnfs, live_nesting_bnfs, fp_ctr_sugar, ...} =
362      fp_sugar_of ctxt fpT_name;
363    val SOME {case_trivial, ...} = codatatype_extra_of ctxt fpT_name;
364
365    val ctr_sugar = #ctr_sugar fp_ctr_sugar;
366    val pre_map_def = map_def_of_bnf pre_bnf;
367    val abs_inverse = #abs_inverse absT_info;
368    val ctr_defs = #ctr_defs fp_ctr_sugar;
369    val case_eq_ifs = #case_eq_ifs ctr_sugar @ case_eq_if_thms_of_term ctxt code_goal;
370    val all_sig_map_thms = maps (#map_thms o #fp_bnf_sugar) sig_fp_sugars;
371
372    val fp_map_ident = map_ident_of_bnf fp_bnf;
373    val fpsig_nesting_bnfs = fp_nesting_bnfs @ maps #live_nesting_bnfs sig_fp_sugars;
374    val fpsig_nesting_T_names = map (fst o dest_Type o T_of_bnf) fpsig_nesting_bnfs;
375    val fpsig_nesting_fp_sugars = map_filter (fp_sugar_of ctxt) fpsig_nesting_T_names;
376    val fpsig_nesting_fp_bnf_sugars = map #fp_bnf_sugar fpsig_nesting_fp_sugars;
377    val ssig_fp_bnf_sugar = #fp_bnf_sugar ssig_fp_sugar;
378    val ssig_bnf = #fp_bnf ssig_fp_sugar;
379    val ssig_map = map_of_bnf ssig_bnf;
380    val fpsig_nesting_maps = map map_of_bnf fpsig_nesting_bnfs;
381    val fpsig_nesting_map_ident0s = map map_ident0_of_bnf fpsig_nesting_bnfs;
382    val fpsig_nesting_map_comps = map map_comp_of_bnf fpsig_nesting_bnfs;
383    val fpsig_nesting_map_thms = maps #map_thms fpsig_nesting_fp_bnf_sugars;
384    val live_nesting_map_ident0s = map map_ident0_of_bnf live_nesting_bnfs;
385    val ssig_map_thms = #map_thms ssig_fp_bnf_sugar;
386    val all_algLam_alg_pointfuls = map (mk_pointful ctxt) all_algLam_algs;
387
388    val \<^const>\<open>Trueprop\<close> $ (Const (\<^const_name>\<open>HOL.eq\<close>, _) $ lhs $ rhs) = code_goal;
389    val (fun_t, args) = strip_comb lhs;
390    val closed_rhs = fold_rev lambda args rhs;
391
392    val fun_T = fastype_of fun_t;
393    val num_args = num_binder_types fun_T;
394
395    val f = Free (singleton (Variable.variant_frees ctxt []) ("f", fun_T));
396
397    val is_self_call = curry (op aconv) fun_t;
398    val has_self_call = exists_subterm is_self_call;
399
400    fun fify args (t $ u) = fify (u :: args) t $ fify [] u
401      | fify _ (Abs (s, T, t)) = Abs (s, T, fify [] t)
402      | fify args t = if t = fun_t andalso not (exists has_self_call args) then f else t;
403
404    val goal = Logic.mk_implies (mk_Trueprop_eq (f, fify [] closed_rhs), mk_Trueprop_eq (f, fun_t))
405      |> Morphism.term phi;
406  in
407    Goal.prove_sorry ctxt [fst (dest_Free f)] [] goal (fn {context = ctxt, prems = _} =>
408      mk_unique_tac ctxt num_args fpsig_nesting_maps ssig_map eval pre_map_def abs_inverse
409        fpsig_nesting_map_ident0s fpsig_nesting_map_comps fpsig_nesting_map_thms
410        live_nesting_map_ident0s fp_map_ident case_trivial ctr_defs case_eq_ifs all_sig_map_thms
411        ssig_map_thms all_algLam_alg_pointfuls (all_algrho_eqs_of ctxt) eval_simps corecUU_unique
412        eq_corecUU)
413    |> Thm.close_derivation \<^here>
414  end;
415
416fun derive_last_disc ctxt fcT_name =
417  let
418    val SOME {T = fcT, discs, exhaust, disc_thmss, ...} = ctr_sugar_of ctxt fcT_name;
419
420    val (u, _) = ctxt
421      |> yield_singleton (mk_Frees "x") fcT;
422
423    val udiscs = map (rapp u) discs;
424    val (not_udiscs, last_udisc) = split_last udiscs
425      |>> map HOLogic.mk_not;
426
427    val goal = mk_Trueprop_eq (last_udisc, foldr1 HOLogic.mk_conj not_udiscs);
428  in
429    Goal.prove_sorry ctxt [fst (dest_Free u)] [] goal (fn {context = ctxt, prems = _} =>
430      mk_last_disc_tac ctxt u exhaust (flat disc_thmss))
431    |> Thm.close_derivation \<^here>
432  end;
433
434fun derive_eq_algrho ctxt {sig_fp_sugars, ssig_fp_sugar, eval, eval_simps, all_algLam_algs,
435      corecUU_unique, ...}
436    ({algrho = algrho0, dtor_algrho, ...} : friend_info) fun_t k_T code_goal const_transfers rho_def
437    eq_corecUU =
438  let
439    val fun_T = fastype_of fun_t;
440    val (arg_Ts, Type (fpT_name, Ts)) = strip_type fun_T;
441    val num_args = length arg_Ts;
442
443    val SOME {fp_res_index, fp_res, pre_bnf, fp_bnf, absT_info, fp_nesting_bnfs, live_nesting_bnfs,
444        fp_ctr_sugar, ...} =
445      fp_sugar_of ctxt fpT_name;
446    val SOME {case_dtor, ...} = codatatype_extra_of ctxt fpT_name;
447
448    val fp_nesting_Ts = map T_of_bnf fp_nesting_bnfs;
449
450    fun is_nullary_disc_def (\<^const>\<open>Trueprop\<close> $ (Const (\<^const_name>\<open>HOL.eq\<close>, _) $ _
451          $ (Const (\<^const_name>\<open>HOL.eq\<close>, _) $ _ $ _))) = true
452      | is_nullary_disc_def (Const (\<^const_name>\<open>Pure.eq\<close>, _) $ _
453          $ (Const (\<^const_name>\<open>HOL.eq\<close>, _) $ _ $ _)) = true
454      | is_nullary_disc_def _ = false;
455
456    val dtor_ctor = nth (#dtor_ctors fp_res) fp_res_index;
457    val ctor_iff_dtor = #ctor_iff_dtor fp_ctr_sugar;
458    val ctr_sugar = #ctr_sugar fp_ctr_sugar;
459    val pre_map_def = map_def_of_bnf pre_bnf;
460    val abs_inverse = #abs_inverse absT_info;
461    val ctr_defs = #ctr_defs fp_ctr_sugar;
462    val nullary_disc_defs = filter (is_nullary_disc_def o Thm.prop_of) (#disc_defs ctr_sugar);
463    val disc_sel_eq_cases = #disc_eq_cases ctr_sugar @ #sel_defs ctr_sugar;
464    val case_eq_ifs = #case_eq_ifs ctr_sugar @ case_eq_if_thms_of_term ctxt code_goal;
465    val all_sig_map_thms = maps (#map_thms o #fp_bnf_sugar) sig_fp_sugars;
466
467    fun add_tnameT (Type (s, Ts)) = insert (op =) s #> fold add_tnameT Ts
468      | add_tnameT _ = I;
469
470    fun map_disc_sels'_of s =
471      (case fp_sugar_of ctxt s of
472        SOME {fp_bnf_sugar = {map_disc_iffs, map_selss, ...}, ...} =>
473        let
474          val map_selss' =
475            if length map_selss <= 1 then map_selss
476            else map (map (unfold_thms ctxt (no_refl [derive_last_disc ctxt s]))) map_selss;
477        in
478          map_disc_iffs @ flat map_selss'
479        end
480      | NONE => []);
481
482    fun mk_const_pointful_natural const_transfer =
483      SOME (mk_pointful_natural_from_transfer ctxt const_transfer)
484      handle UNNATURAL () => NONE;
485
486    val const_pointful_natural_opts = map mk_const_pointful_natural const_transfers;
487    val const_pointful_naturals = map_filter I const_pointful_natural_opts;
488    val fp_nesting_k_T_names = fold add_tnameT (k_T :: fp_nesting_Ts) [];
489    val fp_nesting_k_map_disc_sels' = maps map_disc_sels'_of fp_nesting_k_T_names;
490
491    val fp_map_ident = map_ident_of_bnf fp_bnf;
492    val fpsig_nesting_bnfs = fp_nesting_bnfs @ maps #live_nesting_bnfs sig_fp_sugars;
493    val fpsig_nesting_T_names = map (fst o dest_Type o T_of_bnf) fpsig_nesting_bnfs;
494    val fpsig_nesting_fp_sugars = map_filter (fp_sugar_of ctxt) fpsig_nesting_T_names;
495    val fpsig_nesting_fp_bnf_sugars = map #fp_bnf_sugar fpsig_nesting_fp_sugars;
496    val ssig_fp_bnf_sugar = #fp_bnf_sugar ssig_fp_sugar;
497    val ssig_bnf = #fp_bnf ssig_fp_sugar;
498    val ssig_map = map_of_bnf ssig_bnf;
499    val fpsig_nesting_maps = map map_of_bnf fpsig_nesting_bnfs;
500    val fpsig_nesting_map_ident0s = map map_ident0_of_bnf fpsig_nesting_bnfs;
501    val fpsig_nesting_map_comps = map map_comp_of_bnf fpsig_nesting_bnfs;
502    val fpsig_nesting_map_thms = maps #map_thms fpsig_nesting_fp_bnf_sugars;
503    val live_nesting_map_ident0s = map map_ident0_of_bnf live_nesting_bnfs;
504    val ssig_map_thms = #map_thms ssig_fp_bnf_sugar;
505    val all_algLam_alg_pointfuls = map (mk_pointful ctxt) all_algLam_algs;
506
507    val ctor = nth (#ctors fp_res) fp_res_index;
508    val abs = #abs absT_info;
509    val rep = #rep absT_info;
510    val algrho = mk_ctr Ts algrho0;
511
512    val goal = mk_Trueprop_eq (fun_t, abs_curried_balanced arg_Ts algrho);
513
514    fun const_of_transfer thm =
515      (case Thm.prop_of thm of \<^const>\<open>Trueprop\<close> $ (_ $ cst $ _) => cst);
516
517    val eq_algrho =
518      Goal.prove (*no sorry*) ctxt [] [] goal (fn {context = ctxt, prems = _} =>
519        mk_eq_algrho_tac ctxt fpsig_nesting_maps abs rep ctor ssig_map eval pre_map_def abs_inverse
520          fpsig_nesting_map_ident0s fpsig_nesting_map_comps fpsig_nesting_map_thms
521          live_nesting_map_ident0s fp_map_ident dtor_ctor ctor_iff_dtor ctr_defs nullary_disc_defs
522          disc_sel_eq_cases case_dtor case_eq_ifs const_pointful_naturals
523          fp_nesting_k_map_disc_sels' rho_def dtor_algrho corecUU_unique eq_corecUU all_sig_map_thms
524          ssig_map_thms all_algLam_alg_pointfuls (all_algrho_eqs_of ctxt) eval_simps)
525      |> Thm.close_derivation \<^here>
526      handle e as ERROR _ =>
527        (case filter (is_none o snd) (const_transfers ~~ const_pointful_natural_opts) of
528          [] => Exn.reraise e
529        | thm_nones =>
530          error ("Failed to state naturality property for " ^
531            commas (map (Syntax.string_of_term ctxt o const_of_transfer o fst) thm_nones)));
532
533    val algrho_eq = eq_algrho RS (mk_curry_uncurryN_balanced ctxt num_args RS iffD2) RS sym;
534  in
535    (eq_algrho, algrho_eq)
536  end;
537
538fun prime_rho_transfer_goal ctxt fpT_name rho_def goal =
539  let
540    val thy = Proof_Context.theory_of ctxt;
541
542    val SOME {pre_bnf, ...} = fp_sugar_of ctxt fpT_name;
543    val SOME {abs_rep_transfers, ...} = codatatype_extra_of ctxt fpT_name;
544
545    val simps = rel_def_of_bnf pre_bnf :: rho_transfer_simps;
546    val fold_rho = unfold_thms ctxt [rho_def RS @{thm symmetric}];
547
548    fun derive_unprimed rho_transfer' =
549      Variable.add_free_names ctxt goal []
550      |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} =>
551        unfold_thms_tac ctxt simps THEN HEADGOAL (rtac ctxt rho_transfer')))
552      |> Thm.close_derivation \<^here>;
553
554    val goal' = Raw_Simplifier.rewrite_term thy simps [] goal;
555  in
556    if null abs_rep_transfers then (goal', derive_unprimed #> fold_rho)
557    else (goal, fold_rho)
558  end;
559
560fun derive_rho_transfer_folded ctxt fpT_name const_transfers rho_def goal =
561  let
562    val SOME {pre_bnf, ...} = fp_sugar_of ctxt fpT_name;
563    val SOME {abs_rep_transfers, ...} = codatatype_extra_of ctxt fpT_name;
564  in
565    Variable.add_free_names ctxt goal []
566    |> (fn vars => Goal.prove (*no sorry*) ctxt vars [] goal (fn {context = ctxt, prems = _} =>
567      mk_rho_transfer_tac ctxt (null abs_rep_transfers) (rel_def_of_bnf pre_bnf)
568      const_transfers))
569    |> unfold_thms ctxt [rho_def RS @{thm symmetric}]
570    |> Thm.close_derivation \<^here>
571  end;
572
573fun mk_cong_intro_ctr_or_friend_goal ctxt fpT Rcong alg =
574  let
575    val xy_Ts = binder_types (fastype_of alg);
576
577    val ((xs, ys), _) = ctxt
578      |> mk_Frees "x" xy_Ts
579      ||>> mk_Frees "y" xy_Ts;
580
581    fun mk_prem xy_T x y =
582      build_rel [] ctxt [fpT] [] (fn (T, _) => if T = fpT then Rcong else HOLogic.eq_const T)
583        (xy_T, xy_T) $ x $ y;
584
585    val prems = @{map 3} mk_prem xy_Ts xs ys;
586    val concl = Rcong $ list_comb (alg, xs) $ list_comb (alg, ys);
587  in
588    Logic.list_implies (map HOLogic.mk_Trueprop prems, HOLogic.mk_Trueprop concl)
589  end;
590
591fun derive_cong_ctr_intros ctxt cong_ctor_intro =
592  let
593    val \<^const>\<open>Pure.imp\<close> $ _ $ (\<^const>\<open>Trueprop\<close> $ ((Rcong as _ $ _) $ _ $ (ctor $ _))) =
594      Thm.prop_of cong_ctor_intro;
595
596    val fpT as Type (fpT_name, fp_argTs) = range_type (fastype_of ctor);
597
598    val SOME {pre_bnf, absT_info = {abs_inverse, ...},
599        fp_ctr_sugar = {ctr_defs, ctr_sugar = {ctrs = ctrs0, ...}, ...}, ...} =
600      fp_sugar_of ctxt fpT_name;
601
602    val ctrs = map (mk_ctr fp_argTs) ctrs0;
603    val pre_rel_def = rel_def_of_bnf pre_bnf;
604
605    fun prove ctr_def goal =
606      Variable.add_free_names ctxt goal []
607      |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} =>
608        mk_cong_intro_ctr_or_friend_tac ctxt ctr_def [pre_rel_def, abs_inverse] cong_ctor_intro))
609      |> Thm.close_derivation \<^here>;
610
611    val goals = map (mk_cong_intro_ctr_or_friend_goal ctxt fpT Rcong) ctrs;
612  in
613    map2 prove ctr_defs goals
614  end;
615
616fun derive_cong_friend_intro ctxt cong_algrho_intro =
617  let
618    val \<^const>\<open>Pure.imp\<close> $ _ $ (\<^const>\<open>Trueprop\<close> $ ((Rcong as _ $ _) $ _
619        $ ((algrho as Const (algrho_name, _)) $ _))) =
620      Thm.prop_of cong_algrho_intro;
621
622    val fpT as Type (_, fp_argTs) = range_type (fastype_of algrho);
623
624    fun has_algrho (\<^const>\<open>Trueprop\<close> $ (Const (\<^const_name>\<open>HOL.eq\<close>, _) $ _ $ rhs)) =
625      fst (dest_Const (head_of (strip_abs_body rhs))) = algrho_name;
626
627    val eq_algrho :: _ =
628      maps (filter (has_algrho o Thm.prop_of) o #eq_algrhos o snd) (all_friend_extras_of ctxt);
629
630    val \<^const>\<open>Trueprop\<close> $ (Const (\<^const_name>\<open>HOL.eq\<close>, _) $ friend0 $ _) = Thm.prop_of eq_algrho;
631    val friend = mk_ctr fp_argTs friend0;
632
633    val goal = mk_cong_intro_ctr_or_friend_goal ctxt fpT Rcong friend;
634  in
635    Variable.add_free_names ctxt goal []
636    |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} =>
637      mk_cong_intro_ctr_or_friend_tac ctxt eq_algrho [] cong_algrho_intro))
638    |> Thm.close_derivation \<^here>
639  end;
640
641fun derive_cong_intros lthy ctr_names friend_names
642    ({cong_base, cong_refl, cong_sym, cong_trans, cong_alg_intros, ...} : dtor_coinduct_info) =
643  let
644    val cong_ctor_intro :: cong_algrho_intros = rev cong_alg_intros;
645    val names = map (prefix cong_N) ([baseN, reflN, symN, transN] @ ctr_names @ friend_names);
646    val thms = [cong_base, cong_refl, cong_sym, cong_trans] @
647      derive_cong_ctr_intros lthy cong_ctor_intro @
648      map (derive_cong_friend_intro lthy) cong_algrho_intros;
649  in
650    names ~~ thms
651  end;
652
653fun derive_coinduct ctxt (fpT as Type (fpT_name, fpT_args)) dtor_coinduct =
654  let
655    val thy = Proof_Context.theory_of ctxt;
656
657    val \<^const>\<open>Pure.imp\<close> $ (\<^const>\<open>Trueprop\<close> $ (_ $ Abs (_, _, _ $
658        Abs (_, _, \<^const>\<open>implies\<close> $ _ $ (_ $ (cong0 $ _) $ _ $ _))))) $ _ =
659      Thm.prop_of dtor_coinduct;
660
661    val SOME {X as TVar ((X_s, _), _), fp_res = {dtor_ctors, ...}, pre_bnf,
662        absT_info = {abs_inverse, ...}, live_nesting_bnfs,
663        fp_ctr_sugar = {ctrXs_Tss, ctr_defs,
664          ctr_sugar = ctr_sugar0 as {T = Type (_, T0_args), ctrs = ctrs0, discs = discs0, ...},
665            ...}, ...} =
666      fp_sugar_of ctxt fpT_name;
667
668    val n = length ctrXs_Tss;
669    val ms = map length ctrXs_Tss;
670
671    val X' = TVar ((X_s, maxidx_of_typ fpT + 1), \<^sort>\<open>type\<close>);
672    val As_rho = tvar_subst thy T0_args fpT_args;
673    val substXAT = Term.typ_subst_TVars As_rho o Tsubst X X';
674    val substXA = Term.subst_TVars As_rho o substT X X';
675    val phi = Morphism.typ_morphism "BNF" substXAT $> Morphism.term_morphism "BNF" substXA;
676
677    fun mk_applied_cong arg =
678      enforce_type ctxt domain_type (fastype_of arg) cong0 $ arg;
679
680    val thm = derive_coinduct_thms_for_types ctxt false mk_applied_cong [pre_bnf] dtor_coinduct
681        dtor_ctors live_nesting_bnfs [fpT] [substXAT X] [map (map substXAT) ctrXs_Tss] [n]
682        [abs_inverse] [abs_inverse] I [ctr_defs] [morph_ctr_sugar phi ctr_sugar0]
683      |> map snd |> the_single;
684    val (attrs, _) = mk_coinduct_attrs [fpT] [ctrs0] [discs0] [ms];
685  in
686    (thm, attrs)
687  end;
688
689type explore_parameters =
690  {bound_Us: typ list,
691   bound_Ts: typ list,
692   U: typ,
693   T: typ};
694
695fun update_UT {bound_Us, bound_Ts, ...} U T =
696  {bound_Us = bound_Us, bound_Ts = bound_Ts, U = U, T = T};
697
698fun explore_nested lthy explore {bound_Us, bound_Ts, U, T} t =
699  let
700    fun build_simple (T, U) =
701      if T = U then
702        \<^term>\<open>%y. y\<close>
703      else
704        Bound 0
705        |> explore {bound_Us = T :: bound_Us, bound_Ts = T :: bound_Ts, U = U, T = T}
706        |> (fn t => Abs (Name.uu, T, t));
707  in
708    betapply (build_map lthy [] [] build_simple (T, U), t)
709  end;
710
711fun add_boundvar t = betapply (incr_boundvars 1 t, Bound 0);
712
713fun explore_fun (arg_U :: arg_Us) explore {bound_Us, bound_Ts, U, T} t =
714    let val arg_name = the_default Name.uu (try (fn (Abs (s, _, _)) => s) t) in
715      add_boundvar t
716      |> explore_fun arg_Us explore
717        {bound_Us = arg_U :: bound_Us, bound_Ts = domain_type T :: bound_Ts, U = range_type U,
718         T = range_type T}
719      |> (fn t => Abs (arg_name, arg_U, t))
720    end
721  | explore_fun [] explore params t = explore params t;
722
723fun massage_fun explore (params as {T, U, ...}) =
724  if can dest_funT T then explore_fun [domain_type U] explore params else explore params;
725
726fun massage_star massages explore =
727  let
728    fun after_massage massages' t params t' =
729      if t aconv t' then massage_any massages' params t else massage_any massages params t'
730    and massage_any [] params t = explore params t
731      | massage_any (massage :: massages') params t =
732        massage (after_massage massages' t) params t;
733  in
734    massage_any massages
735  end;
736
737fun massage_let explore params t =
738  (case strip_comb t of
739    (Const (\<^const_name>\<open>Let\<close>, _), [_, _]) => unfold_lets_splits t
740  | _ => t)
741  |> explore params;
742
743fun check_corec_equation ctxt fun_frees (lhs, rhs) =
744  let
745    val (fun_t, arg_ts) = strip_comb lhs;
746
747    fun check_fun_name () =
748      null fun_frees orelse member (op aconv) fun_frees fun_t orelse
749      ill_formed_equation_head ctxt [] fun_t;
750
751    fun check_no_other_frees () =
752      (case Term.add_frees rhs [] |> map Free |> subtract (op =) (fun_frees @ arg_ts)
753          |> find_first (not o Variable.is_fixed ctxt o fst o dest_Free) of
754        NONE => ()
755      | SOME t => extra_variable_in_rhs ctxt [] t);
756  in
757    check_duplicate_variables_in_lhs ctxt [] arg_ts;
758    check_fun_name ();
759    check_all_fun_arg_frees ctxt [] (filter_out is_Var arg_ts);
760    check_no_other_frees ()
761  end;
762
763fun parse_corec_equation ctxt fun_frees eq =
764  let
765    val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop (drop_all eq))
766      handle TERM _ => ill_formed_equation_lhs_rhs ctxt [eq];
767
768    val _ = check_corec_equation ctxt fun_frees (lhs, rhs);
769
770    val (fun_t, arg_ts) = strip_comb lhs;
771    val (arg_Ts, _) = strip_type (fastype_of fun_t);
772    val added_Ts = drop (length arg_ts) arg_Ts;
773    val free_names = mk_names (length added_Ts) "x" ~~ added_Ts;
774    val free_args = Variable.variant_frees ctxt [rhs, lhs] free_names |> map Free;
775  in
776    (arg_ts @ free_args, list_comb (rhs, free_args))
777  end;
778
779fun morph_views phi (code, ctrs, discs, disc_iffs, sels) =
780  (Morphism.term phi code, map (Morphism.term phi) ctrs, map (Morphism.term phi) discs,
781   map (Morphism.term phi) disc_iffs, map (Morphism.term phi) sels);
782
783fun generate_views ctxt eq fun_t (lhs_free_args, rhs) =
784  let
785    val lhs = list_comb (fun_t, lhs_free_args);
786    val T as Type (T_name, Ts) = fastype_of rhs;
787    val SOME {fp_ctr_sugar = {ctr_sugar = {ctrs = ctrs0, discs = discs0, selss = selss0, ...}, ...},
788        ...} =
789      fp_sugar_of ctxt T_name;
790    val ctrs = map (mk_ctr Ts) ctrs0;
791    val discs = map (mk_disc_or_sel Ts) discs0;
792    val selss = map (map (mk_disc_or_sel Ts)) selss0;
793
794    val code_view = drop_all eq;
795
796    fun can_case_expand t = not (can (dest_ctr ctxt T_name) t);
797
798    fun generate_raw_views conds t raw_views =
799      let
800        fun analyse (ctr :: ctrs) (disc :: discs) ctr' =
801          if ctr = ctr' then
802            (conds, disc, ctr)
803          else
804            analyse ctrs discs ctr';
805      in
806        (analyse ctrs discs (fst (strip_comb t))) :: raw_views
807      end;
808
809    fun generate_disc_views raw_views =
810      if length discs = 1 then
811        ([], [])
812      else
813        let
814          fun collect_condss_disc condss [] _ = condss
815            | collect_condss_disc condss ((conds, disc', _) :: raw_views) disc =
816              collect_condss_disc (condss |> disc = disc' ? cons conds) raw_views disc;
817
818          val grouped_disc_views = discs
819            |> map (collect_condss_disc [] raw_views)
820            |> curry (op ~~) (map (fn disc => disc $ lhs) discs);
821
822          fun mk_disc_iff_props props [] = props
823            | mk_disc_iff_props _ ((lhs, \<^const>\<open>HOL.True\<close>) :: _) = [lhs]
824            | mk_disc_iff_props props ((lhs, rhs) :: views) =
825              mk_disc_iff_props ((HOLogic.mk_eq (lhs, rhs)) :: props) views;
826        in
827          (grouped_disc_views
828           |> map swap,
829           grouped_disc_views
830           |> map (apsnd (s_dnf #> mk_conjs))
831           |> mk_disc_iff_props []
832           |> map (fn eq => ([[]], eq)))
833        end;
834
835    fun generate_ctr_views raw_views =
836      let
837        fun collect_condss_ctr condss [] _ = condss
838          | collect_condss_ctr condss ((conds, _, ctr') :: raw_views) ctr =
839            collect_condss_ctr (condss |> ctr = ctr' ? cons conds) raw_views ctr;
840
841        fun mk_ctr_eq ctr_sels ctr =
842          let
843            fun extract_arg n sel _(*bound_Ts*) fun_t arg_ts =
844              if ctr = fun_t then
845                nth arg_ts n
846              else
847                let val t = list_comb (fun_t, arg_ts) in
848                  if can_case_expand t then
849                    sel $ t
850                  else
851                    Term.dummy_pattern (range_type (fastype_of sel))
852                end;
853          in
854            ctr_sels
855            |> map_index (uncurry extract_arg)
856            |> map (fn extract => massage_corec_code_rhs ctxt extract [] rhs)
857            |> curry list_comb ctr
858            |> curry HOLogic.mk_eq lhs
859          end;
860
861          fun remove_condss_if_alone [(_, concl)] = [([[]], concl)]
862            | remove_condss_if_alone views = views;
863      in
864        ctrs
865        |> `(map (collect_condss_ctr [] raw_views))
866        ||> map2 mk_ctr_eq selss
867        |> op ~~
868        |> filter_out (null o fst)
869        |> remove_condss_if_alone
870      end;
871
872    fun generate_sel_views raw_views only_one_ctr =
873      let
874        fun mk_sel_positions sel =
875          let
876            fun get_sel_position _ [] = NONE
877              | get_sel_position i (sel' :: sels) =
878                if sel = sel' then SOME i else get_sel_position (i + 1) sels;
879          in
880            ctrs ~~ map (get_sel_position 0) selss
881            |> map_filter (fn (ctr, pos_opt) =>
882              if is_some pos_opt then SOME (ctr, the pos_opt) else NONE)
883          end;
884
885        fun collect_sel_condss0 condss [] _ = condss
886          | collect_sel_condss0 condss ((conds, _, ctr) :: raw_views) sel_positions =
887            let val condss' = condss |> is_some (AList.lookup (op =) sel_positions ctr) ? cons conds
888            in
889              collect_sel_condss0 condss' raw_views sel_positions
890            end;
891
892        val collect_sel_condss =
893          if only_one_ctr then K [[]] else collect_sel_condss0 [] raw_views;
894
895        fun mk_sel_rhs sel_positions sel =
896          let
897            val sel_T = range_type (fastype_of sel);
898
899            fun extract_sel_value _(*bound_Ts*) fun_t arg_ts =
900              (case AList.lookup (op =) sel_positions fun_t of
901                SOME n => nth arg_ts n
902              | NONE =>
903                let val t = list_comb (fun_t, arg_ts) in
904                  if can_case_expand t then
905                    sel $ t
906                  else
907                    Term.dummy_pattern sel_T
908                end);
909          in
910            massage_corec_code_rhs ctxt extract_sel_value [] rhs
911          end;
912
913        val ordered_sels = distinct (op =) (flat selss);
914        val sel_positionss = map mk_sel_positions ordered_sels;
915        val sel_rhss = map2 mk_sel_rhs sel_positionss ordered_sels;
916        val sel_lhss = map (rapp lhs o mk_disc_or_sel Ts) ordered_sels;
917        val sel_condss = map collect_sel_condss sel_positionss;
918
919        fun is_undefined (Const (\<^const_name>\<open>undefined\<close>, _)) = true
920          | is_undefined _ = false;
921      in
922        sel_condss ~~ (sel_lhss ~~ sel_rhss)
923        |> filter_out (is_undefined o snd o snd)
924        |> map (apsnd HOLogic.mk_eq)
925      end;
926
927    fun mk_atomic_prop fun_args (condss, concl) =
928      (Logic.list_all (map dest_Free fun_args, abstract_over_list fun_args
929        (Logic.list_implies (map HOLogic.mk_Trueprop (s_dnf condss), HOLogic.mk_Trueprop concl))));
930
931    val raw_views = rhs
932      |> massage_let_if_case ctxt (K false) (fn _(*bound_Ts*) => fn t => t
933          |> can_case_expand t ? expand_to_ctr_term ctxt T) (K (K ())) (K I) []
934      |> (fn expanded_rhs => fold_rev_let_if_case ctxt generate_raw_views [] expanded_rhs [])
935      |> rev;
936    val (disc_views, disc_iff_views) = generate_disc_views raw_views;
937    val ctr_views = generate_ctr_views raw_views;
938    val sel_views = generate_sel_views raw_views (length ctr_views = 1);
939
940    val mk_props = filter_out (null o fst) #> map (mk_atomic_prop lhs_free_args);
941  in
942    (code_view, mk_props ctr_views, mk_props disc_views, mk_props disc_iff_views,
943     mk_props sel_views)
944  end;
945
946fun find_all_associated_types [] _ = []
947  | find_all_associated_types ((Type (_, Ts1), Type (_, Ts2)) :: TTs) T =
948    find_all_associated_types ((Ts1 ~~ Ts2) @ TTs) T
949  | find_all_associated_types ((T1, T2) :: TTs) T =
950    find_all_associated_types TTs T |> T1 = T ? cons T2;
951
952fun as_member_of tab = try dest_Const #> Option.mapPartial (fst #> Symtab.lookup tab);
953
954fun extract_rho_from_equation
955    ({ctr_guards, inner_buffer = {Oper, VLeaf, CLeaf, ctr_wrapper, friends}, ...},
956     {pattern_ctrs, discs, sels, it, mk_case})
957    b version Y preT ssig_T friend_tm (lhs_args, rhs) lthy =
958  let
959    val thy = Proof_Context.theory_of lthy;
960
961    val res_T = fastype_of rhs;
962    val YpreT = HOLogic.mk_prodT (Y, preT);
963
964    fun fpT_to new_T T =
965      if T = res_T then
966        new_T
967      else
968        (case T of
969          Type (s, Ts) => Type (s, map (fpT_to new_T) Ts)
970        | _ => T);
971
972    fun build_params bound_Us bound_Ts T =
973      {bound_Us = bound_Us, bound_Ts = bound_Ts, U = T, T = T};
974
975    fun typ_before explore {bound_Us, bound_Ts, ...} t =
976      explore (build_params bound_Us bound_Ts (fastype_of1 (bound_Ts, t))) t;
977
978    val is_self_call = curry (op aconv) friend_tm;
979    val has_self_call = Term.exists_subterm is_self_call;
980
981    fun has_res_T bound_Ts t = fastype_of1 (bound_Ts, t) = res_T;
982
983    fun contains_res_T (Type (s, Ts)) = s = fst (dest_Type res_T) orelse exists contains_res_T Ts
984      | contains_res_T _ = false;
985
986    val is_lhs_arg = member (op =) lhs_args;
987
988    fun is_constant t =
989      not (Term.exists_subterm is_lhs_arg t orelse has_self_call t orelse loose_bvar (t, 0));
990    fun is_nested_type T = T <> res_T andalso T <> YpreT andalso T <> ssig_T;
991
992    val is_valid_case_argumentT = not o member (op =) [Y, ssig_T];
993
994    fun is_same_type_constr (Type (s, _)) (Type (s', _)) = (s = s')
995      | is_same_type_constr _ _ = false;
996
997    exception NO_ENCAPSULATION of unit;
998
999    val parametric_consts = Unsynchronized.ref [];
1000
1001    (* We are assuming that set functions are marked with "[transfer_rule]" (cf. the "transfer"
1002       plugin). Otherwise, the "eq_algrho" tactic might fail. *)
1003    fun is_special_parametric_const (x as (s, _)) =
1004      s = \<^const_name>\<open>id\<close> orelse is_set lthy x;
1005
1006    fun add_parametric_const s general_T T U =
1007      let
1008        fun tupleT_of_funT T =
1009          let val (Ts, T) = strip_type T in
1010            mk_tupleT_balanced (Ts @ [T])
1011          end;
1012
1013        fun funT_of_tupleT n =
1014          dest_tupleT_balanced (n + 1)
1015          #> split_last
1016          #> op --->;
1017
1018        val m = num_binder_types general_T;
1019        val param1_T = Type_Infer.paramify_vars general_T;
1020        val param2_T = Type_Infer.paramify_vars param1_T;
1021
1022        val deadfixed_T =
1023          build_map lthy [] [] (mk_undefined o op -->) (apply2 tupleT_of_funT (param1_T, param2_T))
1024          |> singleton (Type_Infer_Context.infer_types lthy)
1025          |> singleton (Type_Infer.fixate lthy false)
1026          |> type_of
1027          |> dest_funT
1028          |-> generalize_types 1
1029          |> funT_of_tupleT m;
1030
1031        val j = maxidx_of_typ deadfixed_T + 1;
1032
1033        fun varifyT (Type (s, Ts)) = Type (s, map varifyT Ts)
1034          | varifyT (TFree (s, T)) = TVar ((s, j), T)
1035          | varifyT T = T;
1036
1037        val dedvarified_T = varifyT deadfixed_T;
1038
1039        val new_vars = Sign.typ_match thy (dedvarified_T, T) Vartab.empty
1040          |> Vartab.dest
1041          |> filter (curry (op =) j o snd o fst)
1042          |> Vartab.make;
1043
1044        val deadinstantiated_T = map_atyps (Type.devar new_vars) dedvarified_T;
1045
1046        val final_T =
1047          if Sign.typ_instance thy (U, deadinstantiated_T) then deadfixed_T else general_T;
1048      in
1049        parametric_consts := insert (op =) (s, final_T) (!parametric_consts)
1050      end;
1051
1052    fun encapsulate (params as {U, T, ...}) t =
1053      if U = T then
1054        t
1055      else if T = Y then
1056        VLeaf $ t
1057      else if T = res_T then
1058        CLeaf $ t
1059      else if T = YpreT then
1060        it $ t
1061      else if is_nested_type T andalso is_same_type_constr T U then
1062        explore_nested lthy encapsulate params t
1063      else
1064        raise NO_ENCAPSULATION ();
1065
1066    fun build_function_after_encapsulation fun_t fun_t' (params as {bound_Us, ...}) arg_ts arg_ts' =
1067      let
1068        val arg_Us' = fst (strip_typeN (length arg_ts) (fastype_of1 (bound_Us, fun_t')));
1069
1070        fun the_or_error arg NONE =
1071            error ("Illegal argument " ^ quote (Syntax.string_of_term lthy arg) ^
1072              " to " ^ quote (Syntax.string_of_term lthy fun_t))
1073          | the_or_error _ (SOME arg') = arg';
1074      in
1075        arg_ts'
1076        |> `(map (curry fastype_of1 bound_Us))
1077        |>> map2 (update_UT params) arg_Us'
1078        |> op ~~
1079        |> map (try (uncurry encapsulate))
1080        |> map2 the_or_error arg_ts
1081        |> curry list_comb fun_t'
1082      end;
1083
1084    fun rebuild_function_after_exploration old_fn new_fn explore params arg_ts =
1085      arg_ts
1086      |> map (typ_before explore params)
1087      |> build_function_after_encapsulation old_fn new_fn params arg_ts;
1088
1089    fun update_case Us U casex =
1090      let
1091        val Type (T_name, _) = domain_type (snd (strip_fun_type (fastype_of casex)));
1092        val SOME {fp_ctr_sugar = {ctr_sugar = {T = Type (_, Ts), casex, ...}, ...}, ...} =
1093          fp_sugar_of lthy T_name;
1094        val T = body_type (fastype_of casex);
1095      in
1096        Term.subst_atomic_types ((T :: Ts) ~~ (U :: Us)) casex
1097      end;
1098
1099    fun deduce_according_type default_T [] = default_T
1100      | deduce_according_type default_T Ts = (case distinct (op =) Ts of
1101          U :: [] => U
1102        | _ => fpT_to ssig_T default_T);
1103
1104    fun massage_if explore_cond explore (params as {bound_Us, bound_Ts, ...}) t =
1105      (case strip_comb t of
1106        (const as Const (\<^const_name>\<open>If\<close>, _), obj :: (branches as [_, _])) =>
1107        (case List.partition Term.is_dummy_pattern (map (explore params) branches) of
1108          (dummy_branch' :: _, []) => dummy_branch'
1109        | (_, [branch']) => branch'
1110        | (_, branches') =>
1111          let
1112            val brancheUs = map (curry fastype_of1 bound_Us) branches';
1113            val U = deduce_according_type (fastype_of1 (bound_Ts, hd branches)) brancheUs;
1114            val const_obj' = (If_const U, obj)
1115              ||> explore_cond (update_UT params \<^typ>\<open>bool\<close> \<^typ>\<open>bool\<close>)
1116              |> op $;
1117          in
1118            build_function_after_encapsulation (const $ obj) const_obj' params branches branches'
1119          end)
1120      | _ => explore params t);
1121
1122    fun massage_map explore (params as {bound_Us, bound_Ts, T = Type (T_name, Ts), ...})
1123          (t as func $ mapped_arg) =
1124        if is_self_call (head_of func) then
1125          explore params t
1126        else
1127          (case try (dest_map lthy T_name) func of
1128            SOME (map_tm, fs) =>
1129            let
1130              val n = length fs;
1131              val mapped_arg' = mapped_arg
1132                |> `(curry fastype_of1 bound_Ts)
1133                |>> build_params bound_Us bound_Ts
1134                |-> explore;
1135            in
1136              (case fastype_of1 (bound_Us, mapped_arg') of
1137                Type (U_name, Us0) =>
1138                if U_name = T_name then
1139                  let
1140                    val Us = map (fpT_to ssig_T) Us0;
1141                    val temporary_map = map_tm
1142                      |> mk_map n Us Ts;
1143                    val map_fn_Ts = fastype_of #> strip_fun_type #> fst;
1144                    val binder_Uss = map_fn_Ts temporary_map
1145                      |> map (map (fpT_to ssig_T) o binder_types);
1146                    val fun_paramss = map_fn_Ts (head_of func)
1147                      |> map (build_params bound_Us bound_Ts);
1148                    val fs' = fs
1149                      |> @{map 4} explore_fun binder_Uss (replicate n explore) fun_paramss;
1150                    val SOME bnf = bnf_of lthy T_name;
1151                    val Type (_, bnf_Ts) = T_of_bnf bnf;
1152                    val typ_alist =
1153                      lives_of_bnf bnf ~~ map (curry fastype_of1 bound_Us #> range_type) fs';
1154                    val Us' = map2 the_default Us (map (AList.lookup (op =) typ_alist) bnf_Ts);
1155                    val map_tm' = map_tm |> mk_map n Us Us';
1156                  in
1157                    build_function_after_encapsulation func (list_comb (map_tm', fs')) params
1158                      [mapped_arg] [mapped_arg']
1159                  end
1160                else
1161                  explore params t
1162              | _ => explore params t)
1163            end
1164          | NONE => explore params t)
1165      | massage_map explore params t = explore params t;
1166
1167    fun massage_comp explore (params as {bound_Us, ...}) t =
1168      (case strip_comb t of
1169        (Const (\<^const_name>\<open>comp\<close>, _), f1 :: f2 :: args) =>
1170        let
1171          val args' = map (typ_before explore params) args;
1172          val f2' = typ_before (explore_fun (map (curry fastype_of1 bound_Us) args') explore) params
1173            f2;
1174          val f1' = typ_before (explore_fun [range_type (fastype_of1 (bound_Us, f2'))] explore)
1175            params f1;
1176        in
1177          betapply (f1', list_comb (f2', args'))
1178        end
1179      | _ => explore params t);
1180
1181    fun massage_ctr explore (params as {T = T as Type (s, Ts), bound_Us, ...}) t =
1182        if T <> res_T then
1183          (case try (dest_ctr lthy s) t of
1184            SOME (ctr, args) =>
1185            let
1186              val args' = map (typ_before explore params) args;
1187              val SOME {T = Type (_, ctr_Ts), ...} = ctr_sugar_of lthy s;
1188              val temp_ctr = mk_ctr ctr_Ts ctr;
1189              val argUs = map (curry fastype_of1 bound_Us) args';
1190              val typ_alist = binder_types (fastype_of temp_ctr) ~~ argUs;
1191              val Us = ctr_Ts
1192                |> map (find_all_associated_types typ_alist)
1193                |> map2 deduce_according_type Ts;
1194              val ctr' = mk_ctr Us ctr;
1195            in
1196              build_function_after_encapsulation ctr ctr' params args args'
1197            end
1198          | NONE => explore params t)
1199        else
1200          explore params t
1201      | massage_ctr explore params t = explore params t;
1202
1203    fun const_of [] _ = NONE
1204      | const_of ((sel as Const (s1, _)) :: r) (const as Const (s2, _)) =
1205        if s1 = s2 then SOME sel else const_of r const
1206      | const_of _ _ = NONE;
1207
1208    fun massage_disc explore (params as {T, bound_Us, bound_Ts, ...}) t =
1209      (case (strip_comb t, T = \<^typ>\<open>bool\<close>) of
1210        ((fun_t, arg :: []), true) =>
1211        let val arg_T = fastype_of1 (bound_Ts, arg) in
1212          if arg_T <> res_T then
1213            (case arg_T |> try (fst o dest_Type) |> Option.mapPartial (ctr_sugar_of lthy) of
1214              SOME {discs, T = Type (_, Ts), ...} =>
1215              (case const_of discs fun_t of
1216                SOME disc =>
1217                let
1218                  val arg' = arg |> typ_before explore params;
1219                  val Type (_, Us) = fastype_of1 (bound_Us, arg');
1220                  val disc' = disc |> Term.subst_TVars (map (fst o dest_TVar) Ts ~~ Us);
1221                in
1222                  disc' $ arg'
1223                end
1224              | NONE => explore params t)
1225            | NONE => explore params t)
1226          else
1227            explore params t
1228        end
1229      | _ => explore params t);
1230
1231    fun massage_sel explore (params as {bound_Us, bound_Ts, ...}) t =
1232      let val (fun_t, args) = strip_comb t in
1233        if args = [] then
1234          explore params t
1235        else
1236          let val T = fastype_of1 (bound_Ts, hd args) in
1237            (case (Option.mapPartial (ctr_sugar_of lthy) (try (fst o dest_Type) T), T <> res_T) of
1238              (SOME {selss, T = Type (_, Ts), ...}, true) =>
1239              (case const_of (flat selss) fun_t of
1240                SOME sel =>
1241                let
1242                  val args' = args |> map (typ_before explore params);
1243                  val Type (_, Us) = fastype_of1 (bound_Us, hd args');
1244                  val sel' = sel |> Term.subst_TVars (map (fst o dest_TVar) Ts ~~ Us);
1245                in
1246                  build_function_after_encapsulation sel sel' params args args'
1247                end
1248              | NONE => explore params t)
1249            | _ => explore params t)
1250          end
1251      end;
1252
1253    fun massage_equality explore (params as {bound_Us, bound_Ts, ...})
1254          (t as Const (\<^const_name>\<open>HOL.eq\<close>, _) $ t1 $ t2) =
1255        let
1256          val check_is_VLeaf =
1257            not o (Term.exists_subterm (fn t => t aconv CLeaf orelse t aconv Oper));
1258
1259          fun try_pattern_matching (fun_t, arg_ts) t =
1260            (case as_member_of pattern_ctrs fun_t of
1261              SOME (disc, sels) =>
1262              let val t' = typ_before explore params t in
1263                if fastype_of1 (bound_Us, t') = YpreT then
1264                  let
1265                    val arg_ts' = map (typ_before explore params) arg_ts;
1266                    val sels_t' = map (fn sel => betapply (sel, t')) sels;
1267                    val Ts = map (curry fastype_of1 bound_Us) arg_ts';
1268                    val Us = map (curry fastype_of1 bound_Us) sels_t';
1269                    val arg_ts' = map2 encapsulate (map2 (update_UT params) Us Ts) arg_ts';
1270                  in
1271                    if forall check_is_VLeaf arg_ts' then
1272                      SOME (Library.foldl1 HOLogic.mk_conj
1273                        (betapply (disc, t') :: (map HOLogic.mk_eq (arg_ts' ~~ sels_t'))))
1274                    else
1275                      NONE
1276                  end
1277                else
1278                  NONE
1279              end
1280            | NONE => NONE);
1281        in
1282          (case try_pattern_matching (strip_comb t1) t2 of
1283            SOME cond => cond
1284          | NONE => (case try_pattern_matching (strip_comb t2) t1 of
1285              SOME cond => cond
1286            | NONE =>
1287              let
1288                val T = fastype_of1 (bound_Ts, t1);
1289                val params' = build_params bound_Us bound_Ts T;
1290                val t1' = explore params' t1;
1291                val t2' = explore params' t2;
1292              in
1293                if fastype_of1 (bound_Us, t1') = T andalso fastype_of1 (bound_Us, t2') = T then
1294                  HOLogic.mk_eq (t1', t2')
1295                else
1296                  error ("Unsupported condition: " ^ quote (Syntax.string_of_term lthy t))
1297              end))
1298        end
1299      | massage_equality explore params t = explore params t;
1300
1301    fun infer_types (TVar _) (TVar _) = []
1302      | infer_types (U as TVar _) T = [(U, T)]
1303      | infer_types (Type (s', Us)) (Type (s, Ts)) =
1304        if s' = s then flat (map2 infer_types Us Ts) else []
1305      | infer_types _ _ = [];
1306
1307    fun group_by_fst associations [] = associations
1308      | group_by_fst associations ((a, b) :: r) = group_by_fst (add_association a b associations) r
1309    and add_association a b [] = [(a, [b])]
1310      | add_association a b ((c, d) :: r) =
1311        if a = c then (c, b :: d) :: r
1312        else (c, d) :: (add_association a b r);
1313
1314    fun new_TVar known_TVars =
1315      Name.invent_list (map (fst o fst o dest_TVar) known_TVars) "x" 1
1316      |> (fn [s] => TVar ((s, 0), []));
1317
1318    fun instantiate_type inferred_types =
1319      Term.typ_subst_TVars (map (apfst (fst o dest_TVar)) inferred_types);
1320
1321    fun chose_unknown_TVar (T as TVar _) = SOME T
1322      | chose_unknown_TVar (Type (_, Ts)) =
1323        fold (curry merge_options) (map chose_unknown_TVar Ts) NONE
1324      | chose_unknown_TVar _ = NONE;
1325
1326    (* The function under definition might not be defined yet when this is queried. *)
1327    fun maybe_const_type ctxt (s, T) =
1328      Sign.const_type (Proof_Context.theory_of ctxt) s |> the_default T;
1329
1330    fun massage_const polymorphic explore (params as {bound_Us, ...}) t =
1331      let val (fun_t, arg_ts) = strip_comb t in
1332        (case fun_t of
1333          Const (fun_x as (s, fun_T)) =>
1334          let val general_T = if polymorphic then maybe_const_type lthy fun_x else fun_T in
1335            if fun_t aconv friend_tm orelse contains_res_T (body_type general_T) orelse
1336               is_constant t then
1337              explore params t
1338            else
1339              let
1340                val inferred_types = infer_types general_T fun_T;
1341
1342                fun prepare_skeleton [] _ = []
1343                  | prepare_skeleton ((T, U) :: inferred_types) As =
1344                    let
1345                      fun schematize_res_T U As =
1346                        if U = res_T then
1347                          let val A = new_TVar As in
1348                            (A, A :: As)
1349                          end
1350                        else
1351                          (case U of
1352                            Type (s, Us) => fold_map schematize_res_T Us As |>> curry Type s
1353                          | _ => (U, As));
1354
1355                      val (U', As') = schematize_res_T U As;
1356                    in
1357                      (T, U') :: (prepare_skeleton inferred_types As')
1358                    end;
1359
1360                val inferred_types' = prepare_skeleton inferred_types (map fst inferred_types);
1361                val skeleton_T = instantiate_type inferred_types' general_T;
1362
1363                fun explore_if_possible (exp_arg as (_, true)) _ = exp_arg
1364                  | explore_if_possible (exp_arg as (arg, false)) arg_T =
1365                    if exists (exists_subtype is_TVar) (binder_types arg_T) then exp_arg
1366                    else (typ_before (explore_fun (binder_types arg_T) explore) params arg, true);
1367
1368                fun collect_inferred_types [] _ = []
1369                  | collect_inferred_types ((arg, explored) :: exp_arg_ts) (arg_T :: arg_Ts) =
1370                    (if explored then infer_types arg_T (fastype_of1 (bound_Us, arg)) else []) @
1371                    collect_inferred_types exp_arg_ts arg_Ts;
1372
1373                fun propagate exp_arg_ts skeleton_T =
1374                  let
1375                    val arg_gen_Ts = binder_types skeleton_T;
1376                    val exp_arg_ts = map2 explore_if_possible exp_arg_ts arg_gen_Ts;
1377                    val inferred_types = collect_inferred_types exp_arg_ts arg_gen_Ts
1378                      |> group_by_fst []
1379                      |> map (apsnd (deduce_according_type ssig_T));
1380                  in
1381                    (exp_arg_ts, instantiate_type inferred_types skeleton_T)
1382                  end;
1383
1384                val remaining_to_be_explored = filter_out snd #> length;
1385
1386                fun try_exploring_args exp_arg_ts skeleton_T =
1387                  let
1388                    val n = remaining_to_be_explored exp_arg_ts;
1389                    val (exp_arg_ts', skeleton_T') = propagate exp_arg_ts skeleton_T;
1390                    val n' = remaining_to_be_explored exp_arg_ts';
1391
1392                    fun try_instantiating A T =
1393                      try (try_exploring_args exp_arg_ts') (instantiate_type [(A, T)] skeleton_T');
1394                  in
1395                    if n' = 0 then
1396                      SOME (exp_arg_ts', skeleton_T')
1397                    else if n = n' then
1398                      if exists_subtype is_TVar skeleton_T' then
1399                        let val SOME A = chose_unknown_TVar skeleton_T' in
1400                          (case try_instantiating A ssig_T of
1401                            SOME result => result
1402                          | NONE => (case try_instantiating A YpreT of
1403                              SOME result => result
1404                            | NONE => (case try_instantiating A res_T of
1405                                SOME result => result
1406                              | NONE => NONE)))
1407                        end
1408                      else
1409                        NONE
1410                    else
1411                      try_exploring_args exp_arg_ts' skeleton_T'
1412                  end;
1413              in
1414                (case try_exploring_args (map (fn arg => (arg, false)) arg_ts) skeleton_T of
1415                  SOME (exp_arg_ts, fun_U) =>
1416                  let
1417                    val arg_ts' = map fst exp_arg_ts;
1418                    val fun_t' = Const (s, fun_U);
1419
1420                    fun finish_off () =
1421                      let
1422                        val t' =
1423                          build_function_after_encapsulation fun_t fun_t' params arg_ts arg_ts';
1424                      in
1425                        if can type_of1 (bound_Us, t') then
1426                          (if fun_T = fun_U orelse is_special_parametric_const (s, fun_T) then ()
1427                           else add_parametric_const s general_T fun_T fun_U;
1428                           t')
1429                        else
1430                          explore params t
1431                      end;
1432                  in
1433                    if polymorphic then
1434                      finish_off ()
1435                    else
1436                      (case try finish_off () of
1437                        SOME t' => t'
1438                      | NONE => explore params t)
1439                  end
1440                | NONE => explore params t)
1441              end
1442          end
1443        | _ => explore params t)
1444      end;
1445
1446    fun massage_rho explore =
1447      massage_star [massage_let, massage_if explore_cond, massage_case, massage_fun, massage_comp,
1448          massage_map, massage_ctr, massage_sel, massage_disc, massage_equality,
1449          massage_const false, massage_const true]
1450        explore
1451    and massage_case explore (params as {bound_Ts, bound_Us, ...}) t =
1452      (case strip_comb t of
1453        (casex as Const (case_x as (c, _)), args as _ :: _ :: _) =>
1454        (case try strip_fun_type (maybe_const_type lthy case_x) of
1455          SOME (gen_branch_Ts, gen_body_fun_T) =>
1456          let
1457            val gen_branch_ms = map num_binder_types gen_branch_Ts;
1458            val n = length gen_branch_ms;
1459            val (branches, obj_leftovers) = chop n args;
1460          in
1461            if n < length args then
1462              (case gen_body_fun_T of
1463                Type (_, [Type (T_name, _), _]) =>
1464                if case_of lthy T_name = SOME (c, true) then
1465                  let
1466                    val brancheTs = binder_fun_types (fastype_of1 (bound_Ts, casex));
1467                    val obj_leftover_Ts = map (curry fastype_of1 bound_Ts) obj_leftovers;
1468                    val obj_leftovers' =
1469                      if is_constant (hd obj_leftovers) then
1470                        obj_leftovers
1471                      else
1472                        (obj_leftover_Ts, obj_leftovers)
1473                        |>> map (build_params bound_Us bound_Ts)
1474                        |> op ~~
1475                        |> map (uncurry explore_inner);
1476                    val obj_leftoverUs = obj_leftovers' |> map (curry fastype_of1 bound_Us);
1477
1478                    val _ = is_valid_case_argumentT (hd obj_leftoverUs) orelse
1479                      error (quote (Syntax.string_of_term lthy (hd obj_leftovers)) ^
1480                        " is not a valid case argument");
1481
1482                    val Us = obj_leftoverUs |> hd |> dest_Type |> snd;
1483
1484                    val branche_binderUss =
1485                      (if hd obj_leftoverUs = YpreT then mk_case HOLogic.boolT
1486                       else update_case Us HOLogic.boolT casex)
1487                      |> fastype_of
1488                      |> binder_fun_types
1489                      |> map binder_types;
1490                    val b_params = map (build_params bound_Us bound_Ts) brancheTs;
1491
1492                    val branches' = branches
1493                      |> @{map 4} explore_fun branche_binderUss (replicate n explore) b_params;
1494                    val brancheUs = map (curry fastype_of1 bound_Us) branches';
1495                    val U = deduce_according_type (body_type (hd brancheTs))
1496                      (map body_type brancheUs);
1497                    val casex' =
1498                      if hd obj_leftoverUs = YpreT then mk_case U else update_case Us U casex;
1499                  in
1500                    build_function_after_encapsulation casex casex' params
1501                      (branches @ obj_leftovers) (branches' @ obj_leftovers')
1502                  end
1503                else
1504                  explore params t
1505              | _ => explore params t)
1506            else
1507              explore params t
1508          end
1509        | NONE => explore params t)
1510      | _ => explore params t)
1511    and explore_cond params t =
1512      if has_self_call t then unexpected_rec_call_in lthy [] t else explore_inner params t
1513    and explore_inner params t =
1514      massage_rho explore_inner_general params t
1515    and explore_inner_general (params as {bound_Us, bound_Ts, T, ...}) t =
1516      let val (fun_t, arg_ts) = strip_comb t in
1517        if is_constant t then
1518          t
1519        else
1520          (case (as_member_of discs fun_t,
1521              length arg_ts = 1 andalso has_res_T bound_Ts (the_single arg_ts)) of
1522            (SOME disc', true) =>
1523            let
1524              val arg' = explore_inner params (the_single arg_ts);
1525              val arg_U = fastype_of1 (bound_Us, arg');
1526            in
1527              if arg_U = res_T then
1528                fun_t $ arg'
1529              else if arg_U = YpreT then
1530                disc' $ arg'
1531              else
1532                error ("Discriminator " ^ quote (Syntax.string_of_term lthy fun_t) ^
1533                  " cannot be applied to non-variable " ^
1534                  quote (Syntax.string_of_term lthy (hd arg_ts)))
1535            end
1536          | _ =>
1537            (case as_member_of sels fun_t of
1538              SOME sel' =>
1539              let
1540                val arg_ts' = map (explore_inner params) arg_ts;
1541                val arg_U = fastype_of1 (bound_Us, hd arg_ts');
1542              in
1543                if arg_U = res_T then
1544                  build_function_after_encapsulation fun_t fun_t params arg_ts arg_ts'
1545                else if arg_U = YpreT then
1546                  build_function_after_encapsulation fun_t sel' params arg_ts arg_ts'
1547                else
1548                  error ("Selector " ^ quote (Syntax.string_of_term lthy fun_t) ^
1549                    " cannot be applied to non-variable " ^
1550                    quote (Syntax.string_of_term lthy (hd arg_ts)))
1551              end
1552            | NONE =>
1553              (case as_member_of friends fun_t of
1554                SOME (_, friend') =>
1555                rebuild_function_after_exploration fun_t friend' explore_inner params arg_ts
1556                |> curry (op $) Oper
1557              | NONE =>
1558                (case as_member_of ctr_guards fun_t of
1559                  SOME ctr_guard' =>
1560                  rebuild_function_after_exploration fun_t ctr_guard' explore_inner params arg_ts
1561                  |> curry (op $) ctr_wrapper
1562                  |> curry (op $) Oper
1563                | NONE =>
1564                  if is_Bound fun_t then
1565                    rebuild_function_after_exploration fun_t fun_t explore_inner params arg_ts
1566                  else if is_Free fun_t then
1567                    let val fun_t' = map_types (fpT_to YpreT) fun_t in
1568                      rebuild_function_after_exploration fun_t fun_t' explore_inner params arg_ts
1569                    end
1570                  else if T = res_T then
1571                    error (quote (Syntax.string_of_term lthy fun_t) ^
1572                      " not polymorphic enough to be applied like this and no friend")
1573                  else
1574                    error (quote (Syntax.string_of_term lthy fun_t) ^
1575                      " not polymorphic enough to be applied like this")))))
1576      end;
1577
1578    fun explore_ctr params t =
1579      massage_rho explore_ctr_general params t
1580    and explore_ctr_general params t =
1581      let
1582        val (fun_t, arg_ts) = strip_comb t;
1583        val ctr_opt = as_member_of ctr_guards fun_t;
1584      in
1585        if is_some ctr_opt then
1586          rebuild_function_after_exploration fun_t (the ctr_opt) explore_inner params arg_ts
1587        else
1588          not_constructor_in_rhs lthy [] fun_t
1589      end;
1590
1591    val rho_rhs = rhs
1592      |> explore_ctr (build_params [] [] (fastype_of rhs))
1593      |> abs_tuple_balanced (map (map_types (fpT_to YpreT)) lhs_args)
1594      |> unfold_id_bnf_etc lthy;
1595  in
1596    lthy
1597    |> define_const false b version rhoN rho_rhs
1598    |>> pair (!parametric_consts, rho_rhs)
1599  end;
1600
1601fun mk_rho_parametricity_goal ctxt Y Z preT ssig_T dead_pre_rel dead_k_rel dead_ssig_rel rho_rhs =
1602  let
1603    val YpreT = HOLogic.mk_prodT (Y, preT);
1604    val ZpreT = Tsubst Y Z YpreT;
1605    val ssigZ_T = Tsubst Y Z ssig_T;
1606
1607    val dead_pre_rel' = Term.subst_atomic_types [(Y, ssig_T), (Z, ssigZ_T)] dead_pre_rel;
1608    val dead_k_rel' = Term.subst_atomic_types [(Y, YpreT), (Z, ZpreT)] dead_k_rel;
1609
1610    val (R, _) = ctxt
1611      |> yield_singleton (mk_Frees "R") (mk_pred2T Y Z);
1612
1613    val rho_rel = mk_rel_fun (dead_k_rel' $ mk_rel_prod R (dead_pre_rel $ R))
1614      (dead_pre_rel' $ (dead_ssig_rel $ R));
1615    val rho_rhsZ = substT Y Z rho_rhs;
1616  in
1617    HOLogic.mk_Trueprop (rho_rel $ rho_rhs $ rho_rhsZ)
1618  end;
1619
1620fun extract_rho_return_transfer_goals fun_b version dead_pre_bnf dead_k_bnf Y Z preT k_T ssig_T
1621    ssig_fp_sugar friend_parse_info fun_t parsed_eq lthy =
1622  let
1623    fun mk_rel T bnf =
1624      let
1625        val ZT = Tsubst Y Z T;
1626        val rel_T = mk_predT [mk_pred2T Y Z, T, ZT];
1627      in
1628        enforce_type lthy I rel_T (rel_of_bnf bnf)
1629      end;
1630
1631    val ssig_bnf = #fp_bnf ssig_fp_sugar;
1632
1633    val (dead_ssig_bnf, lthy') = bnf_kill_all_but 1 ssig_bnf lthy;
1634
1635    val dead_pre_rel = mk_rel preT dead_pre_bnf;
1636    val dead_k_rel = mk_rel k_T dead_k_bnf;
1637    val dead_ssig_rel = mk_rel ssig_T dead_ssig_bnf;
1638
1639    val (((parametric_consts, rho_rhs), rho_data), lthy'') =
1640      extract_rho_from_equation friend_parse_info fun_b version Y preT ssig_T fun_t parsed_eq lthy';
1641
1642    val const_transfer_goals = map (mk_const_transfer_goal lthy'') parametric_consts;
1643
1644    val rho_transfer_goal =
1645      mk_rho_parametricity_goal lthy'' Y Z preT ssig_T dead_pre_rel dead_k_rel dead_ssig_rel rho_rhs;
1646  in
1647    ((rho_data, (const_transfer_goals, rho_transfer_goal)), lthy'')
1648  end;
1649
1650fun explore_corec_equation ctxt could_be_friend friend fun_name fun_free
1651    {outer_buffer, ctr_guards, inner_buffer} res_T (args, rhs) =
1652  let
1653    val is_self_call = curry (op aconv) fun_free;
1654    val has_self_call = Term.exists_subterm is_self_call;
1655
1656    val outer_ssig_T = body_type (fastype_of (#Oper outer_buffer));
1657
1658    fun inner_fp_of (Free (s, _)) =
1659      Free (s ^ inner_fp_suffix, mk_tupleT_balanced (map fastype_of args) --> outer_ssig_T);
1660
1661    fun build_params bound_Ts U T =
1662      {bound_Us = bound_Ts, bound_Ts = bound_Ts, U = U, T = T};
1663
1664    fun rebuild_function_after_exploration new_fn explore {bound_Ts, ...} arg_ts =
1665      let
1666        val binder_types_old_fn = map (curry fastype_of1 bound_Ts) arg_ts;
1667        val binder_types_new_fn = new_fn
1668          |> binder_types o (curry fastype_of1 bound_Ts)
1669          |> take (length binder_types_old_fn);
1670        val paramss =
1671          map2 (build_params bound_Ts) binder_types_new_fn binder_types_old_fn;
1672      in
1673        map2 explore paramss arg_ts
1674        |> curry list_comb new_fn
1675      end;
1676
1677    fun massage_map_corec explore {bound_Ts, U, T, ...} t =
1678      let val explore' = explore ooo build_params in
1679        massage_nested_corec_call ctxt has_self_call explore' explore' bound_Ts U T t
1680      end;
1681
1682    fun massage_comp explore params t =
1683      (case strip_comb t of
1684        (Const (\<^const_name>\<open>comp\<close>, _), f1 :: f2 :: args) =>
1685        explore params (betapply (f1, (betapplys (f2, args))))
1686      | _ => explore params t);
1687
1688    fun massage_fun explore (params as {bound_Us, bound_Ts, U, T}) t =
1689      if can dest_funT T then
1690        let
1691          val arg_T = domain_type T;
1692          val arg_name = the_default Name.uu (try (fn (Abs (s, _, _)) => s) t);
1693        in
1694          add_boundvar t
1695          |> explore {bound_Us = arg_T :: bound_Us, bound_Ts = arg_T :: bound_Ts,
1696             U = range_type U, T = range_type T}
1697          |> (fn t => Abs (arg_name, arg_T, t))
1698        end
1699      else
1700        explore params t
1701
1702    fun massage_let_if_case_corec explore {bound_Ts, U, T, ...} t =
1703      massage_let_if_case ctxt has_self_call (fn bound_Ts => explore (build_params bound_Ts U T))
1704        (K (unexpected_corec_call_in ctxt [t])) (K (unsupported_case_around_corec_call ctxt [t]))
1705        bound_Ts t;
1706
1707    val massage_map_let_if_case =
1708      massage_star [massage_map_corec, massage_fun, massage_comp, massage_let_if_case_corec];
1709
1710    fun explore_arg _ t =
1711      if has_self_call t then
1712        error (quote (Syntax.string_of_term ctxt t) ^ " contains a nested corecursive call" ^
1713          (if could_be_friend then " (try specifying \"(friend)\")" else ""))
1714      else
1715        t;
1716
1717    fun explore_inner params t =
1718      massage_map_let_if_case explore_inner_general params t
1719    and explore_inner_general (params as {bound_Ts, T, ...}) t =
1720      if T = res_T then
1721        let val (f_t, arg_ts) = strip_comb t in
1722          if has_self_call t then
1723            (case as_member_of (#friends inner_buffer) f_t of
1724              SOME (_, friend') =>
1725              rebuild_function_after_exploration friend' explore_inner params arg_ts
1726              |> curry (op $) (#Oper inner_buffer)
1727            | NONE =>
1728              (case as_member_of ctr_guards f_t of
1729                SOME ctr_guard' =>
1730                rebuild_function_after_exploration ctr_guard' explore_inner params arg_ts
1731                |> curry (op $) (#ctr_wrapper inner_buffer)
1732                |> curry (op $) (#Oper inner_buffer)
1733              | NONE =>
1734                if is_self_call f_t then
1735                  if friend andalso exists has_self_call arg_ts then
1736                    (case Symtab.lookup (#friends inner_buffer) fun_name of
1737                      SOME (_, friend') =>
1738                      rebuild_function_after_exploration friend' explore_inner params arg_ts
1739                      |> curry (op $) (#Oper inner_buffer))
1740                  else
1741                    let val arg_Ts = binder_types (fastype_of1 (bound_Ts, f_t)) in
1742                      map2 explore_arg (map2 (update_UT params) arg_Ts arg_Ts) arg_ts
1743                      |> mk_tuple1_balanced bound_Ts
1744                      |> curry (op $) (#VLeaf inner_buffer)
1745                    end
1746                else
1747                  error (quote (Syntax.string_of_term ctxt f_t) ^ " not registered as friend")))
1748          else
1749            #CLeaf inner_buffer $ t
1750        end
1751      else if has_self_call t then
1752        error (quote (Syntax.string_of_term ctxt t) ^ " contains a corecursive call but has type " ^
1753          quote (Syntax.string_of_typ ctxt T))
1754      else
1755        explore_nested ctxt explore_inner_general params t;
1756
1757    fun explore_outer params t =
1758      massage_map_let_if_case explore_outer_general params t
1759    and explore_outer_general (params as {bound_Ts, T, ...}) t =
1760      if T = res_T then
1761        let val (f_t, arg_ts) = strip_comb t in
1762          (case as_member_of ctr_guards f_t of
1763            SOME ctr_guard' =>
1764            rebuild_function_after_exploration ctr_guard' explore_inner params arg_ts
1765            |> curry (op $) (#VLeaf outer_buffer)
1766          | NONE =>
1767            if not (has_self_call t) then
1768              t
1769              |> expand_to_ctr_term ctxt T
1770              |> massage_let_if_case_corec explore_outer_general params
1771            else
1772              (case as_member_of (#friends outer_buffer) f_t of
1773                SOME (_, friend') =>
1774                rebuild_function_after_exploration friend' explore_outer params arg_ts
1775                |> curry (op $) (#Oper outer_buffer)
1776              | NONE =>
1777                if is_self_call f_t then
1778                  let val arg_Ts = binder_types (fastype_of1 (bound_Ts, f_t)) in
1779                    map2 explore_arg (map2 (update_UT params) arg_Ts arg_Ts) arg_ts
1780                    |> mk_tuple1_balanced bound_Ts
1781                    |> curry (op $) (inner_fp_of f_t)
1782                  end
1783                else
1784                  error (quote (Syntax.string_of_term ctxt f_t) ^ " not registered as friend")))
1785        end
1786      else if has_self_call t then
1787        error (quote (Syntax.string_of_term ctxt t) ^ " contains a corecursive call but has type " ^
1788          quote (Syntax.string_of_typ ctxt T))
1789      else
1790        explore_nested ctxt explore_outer_general params t;
1791  in
1792    (args, rhs
1793      |> explore_outer (build_params [] outer_ssig_T res_T)
1794      |> abs_tuple_balanced args)
1795  end;
1796
1797fun mk_corec_fun_def_rhs ctxt arg_Ts corecUU0 corecUU_arg =
1798  let val corecUU = enforce_type ctxt domain_type (fastype_of corecUU_arg) corecUU0 in
1799    abs_curried_balanced arg_Ts (corecUU $ unfold_id_bnf_etc ctxt corecUU_arg)
1800  end;
1801
1802fun get_options ctxt opts =
1803  let
1804    val plugins = get_first (fn Plugins_Option f => SOME (f ctxt) | _ => NONE) (rev opts)
1805      |> the_default Plugin_Name.default_filter;
1806    val friend = exists (can (fn Friend_Option => ())) opts;
1807    val transfer = exists (can (fn Transfer_Option => ())) opts;
1808  in
1809    (plugins, friend, transfer)
1810  end;
1811
1812fun add_function binding parsed_eq lthy =
1813  let
1814    fun pat_completeness_auto ctxt =
1815      Pat_Completeness.pat_completeness_tac ctxt 1 THEN auto_tac ctxt;
1816
1817    val ({defname, pelims = [[pelim]], pinducts = [pinduct], psimps = [psimp], ...}, lthy') =
1818      Function.add_function [(Binding.concealed binding, NONE, NoSyn)]
1819        [(((Binding.concealed Binding.empty, []), parsed_eq), [], [])]
1820        Function_Common.default_config pat_completeness_auto lthy;
1821  in
1822    ((defname, (pelim, pinduct, psimp)), lthy')
1823  end;
1824
1825fun build_corecUU_arg_and_goals prove_termin (Free (fun_base_name, _)) (arg_ts, explored_rhs) lthy =
1826  let
1827    val inner_fp_name0 = fun_base_name ^ inner_fp_suffix;
1828    val inner_fp_free = Free (inner_fp_name0, fastype_of explored_rhs);
1829  in
1830    if Term.exists_subterm (curry (op aconv) inner_fp_free) explored_rhs then
1831      let
1832        val arg = mk_tuple_balanced arg_ts;
1833        val inner_fp_eq =
1834          mk_Trueprop_eq (betapply (inner_fp_free, arg), betapply (explored_rhs, arg));
1835
1836        val ((inner_fp_name, (pelim, pinduct, psimp)), lthy') =
1837          add_function (Binding.name inner_fp_name0) inner_fp_eq lthy;
1838
1839        fun mk_triple elim induct simp = ([elim], [induct], [simp]);
1840
1841        fun prepare_termin () =
1842          let
1843            val {goal, ...} = Proof.goal (Function.termination NONE lthy');
1844            val termin_goal = goal |> Thm.concl_of |> Logic.unprotect |> Envir.beta_eta_contract;
1845          in
1846            (lthy', (mk_triple pelim pinduct psimp, [termin_goal]))
1847          end;
1848
1849        val (lthy'', (inner_fp_triple, termin_goals)) =
1850          if prove_termin then
1851            (case try (Function.prove_termination NONE
1852                (Function_Common.termination_prover_tac true lthy')) lthy' of
1853              NONE => prepare_termin ()
1854            | SOME ({elims = SOME [[elim]], inducts = SOME [induct], simps = SOME [simp], ...},
1855                lthy'') =>
1856              (lthy'', (mk_triple elim induct simp, [])))
1857          else
1858            prepare_termin ();
1859
1860        val inner_fp_const = (Binding.name_of inner_fp_name, fastype_of explored_rhs)
1861          |>> Proof_Context.read_const {proper = true, strict = false} lthy'
1862          |> (fn (Const (s, _), T) => Const (s, T));
1863      in
1864        (((inner_fp_triple, termin_goals), inner_fp_const), lthy'')
1865      end
1866    else
1867      (((([], [], []), []), explored_rhs), lthy)
1868  end;
1869
1870fun derive_eq_corecUU ctxt {sig_fp_sugars, ssig_fp_sugar, eval, corecUU, eval_simps,
1871      all_algLam_algs, corecUU_unique, ...}
1872    fun_t corecUU_arg fun_code =
1873  let
1874    val fun_T = fastype_of fun_t;
1875    val (arg_Ts, Type (fpT_name, _)) = strip_type fun_T;
1876    val num_args = length arg_Ts;
1877
1878    val SOME {pre_bnf, fp_bnf, absT_info, fp_nesting_bnfs, live_nesting_bnfs, fp_ctr_sugar, ...} =
1879      fp_sugar_of ctxt fpT_name;
1880    val SOME {case_trivial, ...} = codatatype_extra_of ctxt fpT_name;
1881
1882    val ctr_sugar = #ctr_sugar fp_ctr_sugar;
1883    val pre_map_def = map_def_of_bnf pre_bnf;
1884    val abs_inverse = #abs_inverse absT_info;
1885    val ctr_defs = #ctr_defs fp_ctr_sugar;
1886    val case_eq_ifs = #case_eq_ifs ctr_sugar @ case_eq_if_thms_of_term ctxt (Thm.prop_of fun_code);
1887    val all_sig_map_thms = maps (#map_thms o #fp_bnf_sugar) sig_fp_sugars;
1888
1889    val fp_map_ident = map_ident_of_bnf fp_bnf;
1890    val fpsig_nesting_bnfs = fp_nesting_bnfs @ maps #live_nesting_bnfs sig_fp_sugars;
1891    val fpsig_nesting_T_names = map (fst o dest_Type o T_of_bnf) fpsig_nesting_bnfs;
1892    val fpsig_nesting_fp_sugars = map_filter (fp_sugar_of ctxt) fpsig_nesting_T_names;
1893    val fpsig_nesting_fp_bnf_sugars = map #fp_bnf_sugar fpsig_nesting_fp_sugars;
1894    val ssig_fp_bnf_sugar = #fp_bnf_sugar ssig_fp_sugar;
1895    val ssig_bnf = #fp_bnf ssig_fp_sugar;
1896    val ssig_map = map_of_bnf ssig_bnf;
1897    val fpsig_nesting_maps = map map_of_bnf fpsig_nesting_bnfs;
1898    val fpsig_nesting_map_ident0s = map map_ident0_of_bnf fpsig_nesting_bnfs;
1899    val fpsig_nesting_map_comps = map map_comp_of_bnf fpsig_nesting_bnfs;
1900    val fpsig_nesting_map_thms = maps #map_thms fpsig_nesting_fp_bnf_sugars;
1901    val live_nesting_map_ident0s = map map_ident0_of_bnf live_nesting_bnfs;
1902    val ssig_map_thms = #map_thms ssig_fp_bnf_sugar;
1903    val all_algLam_alg_pointfuls = map (mk_pointful ctxt) all_algLam_algs;
1904
1905    val def_rhs = mk_corec_fun_def_rhs ctxt arg_Ts corecUU corecUU_arg;
1906
1907    val goal = mk_Trueprop_eq (fun_t, def_rhs);
1908  in
1909    Goal.prove_sorry ctxt [] [] goal (fn {context = ctxt, prems = _} =>
1910      mk_eq_corecUU_tac ctxt num_args fpsig_nesting_maps ssig_map eval pre_map_def abs_inverse
1911        fpsig_nesting_map_ident0s fpsig_nesting_map_comps fpsig_nesting_map_thms
1912        live_nesting_map_ident0s fp_map_ident case_trivial ctr_defs case_eq_ifs all_sig_map_thms
1913        ssig_map_thms all_algLam_alg_pointfuls (all_algrho_eqs_of ctxt) eval_simps corecUU_unique
1914        fun_code)
1915    |> Thm.close_derivation \<^here>
1916  end;
1917
1918fun derive_coinduct_cong_intros
1919    ({fpT = fpT0 as Type (fpT_name, _), friend_names = friend_names0,
1920      corecUU = Const (corecUU_name, _), dtor_coinduct_info as {dtor_coinduct, ...}, ...})
1921    lthy =
1922  let
1923    val thy = Proof_Context.theory_of lthy;
1924    val phi = Proof_Context.export_morphism lthy (Local_Theory.target_of lthy);
1925
1926    val fpT = Morphism.typ phi fpT0;
1927    val general_fpT = body_type (Sign.the_const_type thy corecUU_name);
1928    val most_general = Sign.typ_instance thy (general_fpT, fpT);
1929  in
1930    (case (most_general, coinduct_extra_of lthy corecUU_name) of
1931      (true, SOME extra) => ((false, extra), lthy)
1932    | _ =>
1933      let
1934        val ctr_names = ctr_names_of_fp_name lthy fpT_name;
1935        val friend_names = friend_names0 |> map Long_Name.base_name |> rev;
1936        val cong_intro_pairs = derive_cong_intros lthy ctr_names friend_names dtor_coinduct_info;
1937        val (coinduct, coinduct_attrs) = derive_coinduct lthy fpT0 dtor_coinduct;
1938        val ((_, [coinduct]), lthy) = (* TODO check: only if most_general?*)
1939          Local_Theory.note ((Binding.empty, coinduct_attrs), [coinduct]) lthy;
1940
1941        val extra = {coinduct = coinduct, coinduct_attrs = coinduct_attrs,
1942          cong_intro_pairs = cong_intro_pairs};
1943      in
1944        ((most_general, extra), lthy |> most_general ? register_coinduct_extra corecUU_name extra)
1945      end)
1946  end;
1947
1948fun update_coinduct_cong_intross_dynamic fpT_name lthy =
1949  let val all_corec_infos = corec_infos_of lthy fpT_name in
1950    lthy
1951    |> fold_map (apfst snd oo derive_coinduct_cong_intros) all_corec_infos
1952    |> snd
1953  end;
1954
1955fun derive_and_update_coinduct_cong_intross [] = pair (false, [])
1956  | derive_and_update_coinduct_cong_intross (corec_infos as {fpT = Type (fpT_name, _), ...} :: _) =
1957    fold_map derive_coinduct_cong_intros corec_infos
1958    #>> split_list
1959    #> (fn ((changeds, extras), lthy) =>
1960      if exists I changeds then
1961        ((true, extras), lthy |> update_coinduct_cong_intross_dynamic fpT_name)
1962      else
1963        ((false, extras), lthy));
1964
1965fun prepare_corec_ursive_cmd int long_cmd opts (raw_fixes, raw_eq) lthy =
1966  let
1967    val _ = can the_single raw_fixes orelse
1968      error "Mutually corecursive functions not supported";
1969
1970    val (plugins, friend, transfer) = get_options lthy opts;
1971    val ([((b, fun_T), mx)], [(_, eq)]) =
1972      fst (Specification.read_multi_specs raw_fixes [((Binding.empty_atts, raw_eq), [], [])] lthy);
1973
1974    val _ = check_top_sort lthy b fun_T;
1975
1976    val (arg_Ts, res_T) = strip_type fun_T;
1977    val fpT_name = (case res_T of Type (s, _) => s | _ => not_codatatype lthy res_T);
1978    val fun_free = Free (Binding.name_of b, fun_T);
1979    val parsed_eq = parse_corec_equation lthy [fun_free] eq;
1980
1981    val fun_name = Local_Theory.full_name lthy b;
1982    val fun_t = Const (fun_name, fun_T);
1983      (* FIXME: does this work with locales that fix variables? *)
1984
1985    val no_base = has_no_corec_info lthy fpT_name;
1986    val lthy1 = lthy |> no_base ? setup_base fpT_name;
1987
1988    fun extract_rho lthy' =
1989      let
1990        val lthy'' = lthy' |> Variable.declare_typ fun_T;
1991        val (prepared as (_, _, version, Y, Z, preT, k_T, ssig_T, dead_pre_bnf, dead_k_bnf, _,
1992               ssig_fp_sugar, buffer), lthy''') =
1993          prepare_friend_corec fun_name fun_T lthy'';
1994        val friend_parse_info = friend_parse_info_of lthy''' arg_Ts res_T buffer;
1995
1996        val parsed_eq' = parsed_eq ||> subst_atomic [(fun_free, fun_t)];
1997      in
1998        lthy'''
1999        |> extract_rho_return_transfer_goals b version dead_pre_bnf dead_k_bnf Y Z preT k_T ssig_T
2000          ssig_fp_sugar friend_parse_info fun_t parsed_eq'
2001        |>> pair prepared
2002      end;
2003
2004    val ((prepareds, (rho_datas, transfer_goal_datas)), lthy2) =
2005      if friend then extract_rho lthy1 |>> (apfst single ##> (apfst single #> apsnd single))
2006      else (([], ([], [])), lthy1);
2007
2008    val ((buffer, corec_infos), lthy3) =
2009      if friend then
2010        ((#13 (the_single prepareds), []), lthy2)
2011      else
2012        corec_info_of res_T lthy2
2013        ||> no_base ? update_coinduct_cong_intross_dynamic fpT_name
2014        |>> (fn info as {buffer, ...} => (buffer, [info]));
2015
2016    val corec_parse_info = corec_parse_info_of lthy3 arg_Ts res_T buffer;
2017
2018    val explored_eq =
2019      explore_corec_equation lthy3 true friend fun_name fun_free corec_parse_info res_T parsed_eq;
2020
2021    val (((inner_fp_triple, termin_goals), corecUU_arg), lthy4) =
2022      build_corecUU_arg_and_goals (not long_cmd) fun_free explored_eq lthy3;
2023
2024    fun def_fun (inner_fp_elims0, inner_fp_inducts0, inner_fp_simps0) const_transfers
2025        rho_transfers_foldeds lthy5 =
2026      let
2027        fun register_friend lthy' =
2028          let
2029            val [(old_corec_info, fp_b, version, Y, Z, _, k_T, _, _, dead_k_bnf, sig_fp_sugar,
2030                  ssig_fp_sugar, _)] = prepareds;
2031            val [(rho, rho_def)] = rho_datas;
2032            val [(_, rho_transfer_goal)] = transfer_goal_datas;
2033            val Type (fpT_name, _) = res_T;
2034
2035            val rho_transfer_folded =
2036              (case rho_transfers_foldeds of
2037                [] =>
2038                derive_rho_transfer_folded lthy' fpT_name const_transfers rho_def rho_transfer_goal
2039              | [thm] => thm);
2040          in
2041            lthy'
2042            |> register_coinduct_dynamic_friend fpT_name fun_name
2043            |> register_friend_corec fun_name fp_b version Y Z k_T dead_k_bnf sig_fp_sugar
2044              ssig_fp_sugar fun_t rho rho_transfer_folded old_corec_info
2045          end;
2046
2047        val (friend_infos, lthy6) = lthy5 |> (if friend then register_friend #>> single else pair []);
2048        val (corec_info as {corecUU = corecUU0, ...}, lthy7) =
2049          (case corec_infos of
2050            [] => corec_info_of res_T lthy6
2051          | [info] => (info, lthy6));
2052
2053        val def_rhs = mk_corec_fun_def_rhs lthy7 arg_Ts corecUU0 corecUU_arg;
2054        val def = ((b, mx), ((Binding.concealed (Thm.def_binding b), []), def_rhs));
2055
2056        val ((fun_lhs0, (_, fun_def0)), (lthy9, lthy8')) = lthy7
2057          |> Local_Theory.open_target |> snd
2058          |> Local_Theory.define def
2059          |> tap (fn (def, lthy') => print_def_consts int [def] lthy')
2060          ||> `Local_Theory.close_target;
2061
2062        val parsed_eq = parse_corec_equation lthy9 [fun_free] eq;
2063        val views0 = generate_views lthy9 eq fun_free parsed_eq;
2064
2065        val lthy9' = lthy9 |> fold Variable.declare_typ (res_T :: arg_Ts);
2066        val phi = Proof_Context.export_morphism lthy8' lthy9';
2067
2068        val fun_lhs = Morphism.term phi fun_lhs0;
2069        val fun_def = Morphism.thm phi fun_def0;
2070        val inner_fp_elims = map (Morphism.thm phi) inner_fp_elims0;
2071        val inner_fp_inducts = map (Morphism.thm phi) inner_fp_inducts0;
2072        val inner_fp_simps = map (Morphism.thm phi) inner_fp_simps0;
2073        val (code_goal, _, _, _, _) = morph_views phi views0;
2074
2075        fun derive_and_note_friend_extra_theorems lthy' =
2076          let
2077            val k_T = #7 (the_single prepareds);
2078            val rho_def = snd (the_single rho_datas);
2079
2080            val (eq_algrho, algrho_eq) = derive_eq_algrho lthy' corec_info (the_single friend_infos)
2081              fun_lhs k_T code_goal const_transfers rho_def fun_def;
2082
2083            val notes =
2084              (if Config.get lthy' bnf_internals then
2085                 [(eq_algrhoN, [eq_algrho])]
2086               else
2087                 [])
2088              |> map (fn (thmN, thms) =>
2089                ((Binding.qualify true (Binding.name_of b)
2090                    (Binding.qualify false friendN (Binding.name thmN)), []),
2091                 [(thms, [])]));
2092          in
2093            lthy'
2094            |> register_friend_extra fun_name eq_algrho algrho_eq
2095            |> Local_Theory.notes notes |> snd
2096          end;
2097
2098        val lthy10 = lthy9 |> friend ? derive_and_note_friend_extra_theorems;
2099
2100        val code_thm = derive_code lthy10 inner_fp_simps code_goal corec_info fun_lhs fun_def;
2101(* TODO:
2102        val ctr_thmss = map mk_thm (#2 views);
2103        val disc_thmss = map mk_thm (#3 views);
2104        val disc_iff_thmss = map mk_thm (#4 views);
2105        val sel_thmss = map mk_thm (#5 views);
2106*)
2107
2108        val uniques =
2109          if null inner_fp_simps then
2110            [derive_unique lthy10 phi (#1 views0) corec_info fpT_name fun_def]
2111          else
2112            [];
2113
2114(* TODO:
2115        val disc_iff_or_disc_thmss =
2116          map2 (fn [] => I | disc_iffs => K disc_iffs) disc_iff_thmss disc_thmss;
2117        val simp_thmss = map2 append disc_iff_or_disc_thmss sel_thmss;
2118*)
2119
2120        val ((_, [{cong_intro_pairs, coinduct, coinduct_attrs}]), lthy11) = lthy10
2121          |> derive_and_update_coinduct_cong_intross [corec_info];
2122        val cong_intros_pairs = AList.group (op =) cong_intro_pairs;
2123
2124        val anonymous_notes = [];
2125(* TODO:
2126          [(flat disc_iff_or_disc_thmss, simp_attrs)]
2127          |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])]));
2128*)
2129
2130        val notes =
2131          [(cong_introsN, maps snd cong_intros_pairs, []),
2132           (codeN, [code_thm], nitpicksimp_attrs),
2133           (coinductN, [coinduct], coinduct_attrs),
2134           (inner_inductN, inner_fp_inducts, []),
2135           (uniqueN, uniques, [])] @
2136           map (fn (thmN, thms) => (thmN, thms, [])) cong_intros_pairs @
2137          (if Config.get lthy11 bnf_internals then
2138             [(inner_elimN, inner_fp_elims, []),
2139              (inner_simpN, inner_fp_simps, [])]
2140           else
2141             [])
2142(* TODO:
2143           (ctrN, ctr_thms, []),
2144           (discN, disc_thms, []),
2145           (disc_iffN, disc_iff_thms, []),
2146           (selN, sel_thms, simp_attrs),
2147           (simpsN, simp_thms, []),
2148*)
2149          |> map (fn (thmN, thms, attrs) =>
2150              ((Binding.qualify true (Binding.name_of b)
2151                  (Binding.qualify false corecN (Binding.name thmN)), attrs),
2152               [(thms, [])]))
2153          |> filter_out (null o fst o hd o snd);
2154      in
2155        lthy11
2156(* TODO:
2157        |> Spec_Rules.add Spec_Rules.equational ([fun_lhs], flat sel_thmss)
2158        |> Spec_Rules.add Spec_Rules.equational ([fun_lhs], flat ctr_thmss)
2159*)
2160        |> Spec_Rules.add Binding.empty Spec_Rules.equational [fun_lhs] [code_thm]
2161        |> plugins code_plugin ? Code.declare_default_eqns [(code_thm, true)]
2162        |> Local_Theory.notes (anonymous_notes @ notes)
2163        |> snd
2164      end;
2165
2166    fun prove_transfer_goal ctxt goal =
2167      Variable.add_free_names ctxt goal []
2168      |> (fn vars => Goal.prove (*no sorry*) ctxt vars [] goal (fn {context = ctxt, prems = _} =>
2169        HEADGOAL (Transfer.transfer_prover_tac ctxt)))
2170      |> Thm.close_derivation \<^here>;
2171
2172    fun maybe_prove_transfer_goal ctxt goal =
2173      (case try (prove_transfer_goal ctxt) goal of
2174        SOME thm => apfst (cons thm)
2175      | NONE => apsnd (cons goal));
2176
2177    val const_transfer_goals = fold (union (op aconv) o fst) transfer_goal_datas [];
2178    val (const_transfers, const_transfer_goals') =
2179      if long_cmd then ([], const_transfer_goals)
2180      else fold (maybe_prove_transfer_goal lthy4) const_transfer_goals ([], []);
2181  in
2182    ((def_fun, (([res_T], prepareds, rho_datas, map snd transfer_goal_datas),
2183        (inner_fp_triple, termin_goals), (const_transfers, const_transfer_goals'))), lthy4)
2184  end;
2185
2186fun corec_cmd int opts (raw_fixes, raw_eq) lthy =
2187  let
2188    val ((def_fun, (_, (inner_fp_triple, termin_goals), (const_transfers, const_transfer_goals))),
2189         lthy') =
2190      prepare_corec_ursive_cmd int false opts (raw_fixes, raw_eq) lthy;
2191  in
2192    if not (null termin_goals) then
2193      error ("Termination prover failed (try " ^ quote (#1 \<^command_keyword>\<open>corecursive\<close>) ^
2194        " instead of " ^ quote (#1 \<^command_keyword>\<open>corec\<close>) ^ ")")
2195    else if not (null const_transfer_goals) then
2196      error ("Transfer prover failed (try " ^ quote (#1 \<^command_keyword>\<open>corecursive\<close>) ^
2197        " instead of " ^ quote (#1 \<^command_keyword>\<open>corec\<close>) ^ ")")
2198    else
2199      def_fun inner_fp_triple const_transfers [] lthy'
2200  end;
2201
2202fun corecursive_cmd int opts (raw_fixes, raw_eq) lthy =
2203  let
2204    val ((def_fun, (([Type (fpT_name, _)], prepareds, rho_datas, rho_transfer_goals),
2205            (inner_fp_triple, termin_goals), (const_transfers, const_transfer_goals))), lthy') =
2206      prepare_corec_ursive_cmd int true opts (raw_fixes, raw_eq) lthy;
2207
2208    val (rho_transfer_goals', unprime_rho_transfer_and_folds) =
2209      @{map 3} (fn (_, _, _, _, _, _, _, _, _, _, _, _, _) => fn (_, rho_def) =>
2210          prime_rho_transfer_goal lthy' fpT_name rho_def)
2211        prepareds rho_datas rho_transfer_goals
2212      |> split_list;
2213  in
2214    Proof.theorem NONE (fn [termin_thms, const_transfers', rho_transfers'] =>
2215      let
2216        val remove_domain_condition =
2217          full_simplify (put_simpset HOL_basic_ss lthy'
2218            addsimps (@{thm True_implies_equals} :: termin_thms));
2219      in
2220        def_fun (@{apply 3} (map remove_domain_condition) inner_fp_triple)
2221          (const_transfers @ const_transfers')
2222          (map2 (fn f => f) unprime_rho_transfer_and_folds rho_transfers')
2223      end)
2224      (map (map (rpair [])) [termin_goals, const_transfer_goals, rho_transfer_goals']) lthy'
2225  end;
2226
2227fun friend_of_corec_cmd ((raw_fun_name, raw_fun_T_opt), raw_eq) lthy =
2228  let
2229    val Const (fun_name, _) =
2230      Proof_Context.read_const {proper = true, strict = false} lthy raw_fun_name;
2231
2232    val fake_lthy = lthy
2233      |> (case raw_fun_T_opt of
2234           SOME raw_T =>
2235           Proof_Context.add_const_constraint (fun_name, SOME (Syntax.read_typ lthy raw_T))
2236         | NONE => I)
2237      handle TYPE (s, _, _) => error s;
2238
2239    val fun_b = Binding.name (Long_Name.base_name fun_name);
2240    val code_goal = Syntax.read_prop fake_lthy raw_eq;
2241
2242    val fun_T =
2243      (case code_goal of
2244        \<^const>\<open>Trueprop\<close> $ (Const (\<^const_name>\<open>HOL.eq\<close>, _) $ t $ _) => fastype_of (head_of t)
2245      | _ => ill_formed_equation_lhs_rhs lthy [code_goal]);
2246    val fun_t = Const (fun_name, fun_T);
2247
2248    val (arg_Ts, res_T as Type (fpT_name, _)) = strip_type fun_T;
2249
2250    val no_base = has_no_corec_info lthy fpT_name;
2251    val lthy1 = lthy |> no_base ? setup_base fpT_name;
2252
2253    val lthy2 = lthy1 |> Variable.declare_typ fun_T;
2254    val ((old_corec_info, fp_b, version, Y, Z, preT, k_T, ssig_T, dead_pre_bnf, dead_k_bnf,
2255          sig_fp_sugar, ssig_fp_sugar, buffer), lthy3) =
2256      prepare_friend_corec fun_name fun_T lthy2;
2257    val friend_parse_info = friend_parse_info_of lthy3 arg_Ts res_T buffer;
2258
2259    val parsed_eq = parse_corec_equation lthy3 [] code_goal;
2260
2261    val (((rho, rho_def), (const_transfer_goals, rho_transfer_goal)), lthy4) =
2262      extract_rho_return_transfer_goals fun_b version dead_pre_bnf dead_k_bnf Y Z preT k_T ssig_T
2263        ssig_fp_sugar friend_parse_info fun_t parsed_eq lthy3;
2264
2265    fun register_friend_extra_and_note_thms code_goal code_thm const_transfers k_T friend_info
2266        lthy5 =
2267      let
2268        val (corec_info, lthy6) = corec_info_of res_T lthy5;
2269
2270        val fun_free = Free (Binding.name_of fun_b, fun_T);
2271
2272        fun freeze_fun (t as Const (s, T)) = if s = fun_name andalso T = fun_T then fun_free else t
2273          | freeze_fun t = t;
2274
2275        val eq = Term.map_aterms freeze_fun code_goal;
2276        val parsed_eq = parse_corec_equation lthy6 [fun_free] eq;
2277
2278        val corec_parse_info = corec_parse_info_of lthy6 arg_Ts res_T buffer;
2279        val explored_eq = explore_corec_equation lthy6 false false fun_name fun_free corec_parse_info
2280          res_T parsed_eq;
2281
2282        val ((_, corecUU_arg), _) = build_corecUU_arg_and_goals false fun_free explored_eq lthy6;
2283
2284        val eq_corecUU = derive_eq_corecUU lthy6 corec_info fun_t corecUU_arg code_thm;
2285        val (eq_algrho, algrho_eq) = derive_eq_algrho lthy6 corec_info friend_info fun_t k_T
2286          code_goal const_transfers rho_def eq_corecUU;
2287
2288        val ((_, [{cong_intro_pairs, coinduct, coinduct_attrs}]), lthy7) = lthy6
2289          |> register_friend_extra fun_name eq_algrho algrho_eq
2290          |> register_coinduct_dynamic_friend fpT_name fun_name
2291          |> derive_and_update_coinduct_cong_intross [corec_info];
2292        val cong_intros_pairs = AList.group (op =) cong_intro_pairs;
2293
2294        val unique = derive_unique lthy7 Morphism.identity code_goal corec_info fpT_name eq_corecUU;
2295
2296        val notes =
2297          [(codeN, [code_thm], []),
2298           (coinductN, [coinduct], coinduct_attrs),
2299           (cong_introsN, maps snd cong_intros_pairs, []),
2300           (uniqueN, [unique], [])] @
2301           map (fn (thmN, thms) => (thmN, thms, [])) cong_intros_pairs @
2302          (if Config.get lthy7 bnf_internals then
2303             [(eq_algrhoN, [eq_algrho], []),
2304              (eq_corecUUN, [eq_corecUU], [])]
2305           else
2306             [])
2307          |> map (fn (thmN, thms, attrs) =>
2308            ((Binding.qualify true (Binding.name_of fun_b)
2309                (Binding.qualify false friendN (Binding.name thmN)), attrs),
2310             [(thms, [])]));
2311      in
2312        lthy7
2313        |> Local_Theory.notes notes |> snd
2314      end;
2315
2316    val (rho_transfer_goal', unprime_rho_transfer_and_fold) =
2317      prime_rho_transfer_goal lthy4 fpT_name rho_def rho_transfer_goal;
2318  in
2319    lthy4
2320    |> Proof.theorem NONE (fn [[code_thm], const_transfers, [rho_transfer']] =>
2321        register_friend_corec fun_name fp_b version Y Z k_T dead_k_bnf sig_fp_sugar ssig_fp_sugar
2322          fun_t rho (unprime_rho_transfer_and_fold rho_transfer') old_corec_info
2323        #-> register_friend_extra_and_note_thms code_goal code_thm const_transfers k_T)
2324      (map (map (rpair [])) [[code_goal], const_transfer_goals, [rho_transfer_goal']])
2325    |> Proof.refine_singleton (Method.primitive_text (K I))
2326  end;
2327
2328fun coinduction_upto_cmd (base_name, raw_fpT) lthy =
2329  let
2330    val fpT as Type (fpT_name, _) = Syntax.read_typ lthy raw_fpT;
2331
2332    val no_base = has_no_corec_info lthy fpT_name;
2333
2334    val (corec_info as {version, ...}, lthy1) = lthy
2335      |> corec_info_of fpT;
2336    val lthy2 = lthy1 |> no_base ? setup_base fpT_name;
2337
2338    val ((changed, [{cong_intro_pairs, coinduct, coinduct_attrs}]), lthy3) = lthy2
2339      |> derive_and_update_coinduct_cong_intross [corec_info];
2340    val lthy4 = lthy3 |> (changed orelse no_base) ? update_coinduct_cong_intross_dynamic fpT_name;
2341    val cong_intros_pairs = AList.group (op =) cong_intro_pairs;
2342
2343    val notes =
2344      [(cong_introsN, maps snd cong_intros_pairs, []),
2345       (coinduct_uptoN, [coinduct], coinduct_attrs)] @
2346      map (fn (thmN, thms) => (thmN, thms, [])) cong_intros_pairs
2347      |> map (fn (thmN, thms, attrs) =>
2348        (((Binding.qualify true base_name
2349            (Binding.qualify false ("v" ^ string_of_int version) (Binding.name thmN))), attrs),
2350         [(thms, [])]));
2351  in
2352    lthy4 |> Local_Theory.notes notes |> snd
2353  end;
2354
2355fun consolidate lthy =
2356  let
2357    val corec_infoss = map (corec_infos_of lthy o fst) (all_codatatype_extras_of lthy);
2358    val (changeds, lthy') = lthy
2359      |> fold_map (apfst fst oo derive_and_update_coinduct_cong_intross) corec_infoss;
2360  in
2361    if exists I changeds then lthy' else raise Same.SAME
2362  end;
2363
2364fun consolidate_global thy =
2365  SOME (Named_Target.theory_map consolidate thy)
2366  handle Same.SAME => NONE;
2367
2368val _ = Outer_Syntax.local_theory \<^command_keyword>\<open>corec\<close>
2369  "define nonprimitive corecursive functions"
2370  ((Scan.optional (\<^keyword>\<open>(\<close> |-- Parse.!!! (Parse.list1 corec_option_parser)
2371      --| \<^keyword>\<open>)\<close>) []) -- (Parse.vars --| Parse.where_ -- Parse.prop)
2372   >> uncurry (corec_cmd true));
2373
2374val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\<open>corecursive\<close>
2375  "define nonprimitive corecursive functions"
2376  ((Scan.optional (\<^keyword>\<open>(\<close> |-- Parse.!!! (Parse.list1 corec_option_parser)
2377      --| \<^keyword>\<open>)\<close>) []) -- (Parse.vars --| Parse.where_ -- Parse.prop)
2378   >> uncurry (corecursive_cmd true));
2379
2380val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\<open>friend_of_corec\<close>
2381  "register a function as a legal context for nonprimitive corecursion"
2382  (Parse.const -- Scan.option (\<^keyword>\<open>::\<close> |-- Parse.typ) --| Parse.where_ -- Parse.prop
2383   >> friend_of_corec_cmd);
2384
2385val _ = Outer_Syntax.local_theory \<^command_keyword>\<open>coinduction_upto\<close>
2386  "derive a coinduction up-to principle and a corresponding congruence closure"
2387  (Parse.name --| \<^keyword>\<open>:\<close> -- Parse.typ >> coinduction_upto_cmd);
2388
2389val _ = Theory.setup (Theory.at_begin consolidate_global);
2390
2391end;
2392