377 val thy = Local_Theory.exit_global lthy12c; |
377 val thy = Local_Theory.exit_global lthy12c; |
378 val perm_names = map (fn x => "permute_" ^ x) qty_names |
378 val perm_names = map (fn x => "permute_" ^ x) qty_names |
379 val thy' = define_lifted_perms qty_full_names (perm_names ~~ perms) raw_perm_simps thy; |
379 val thy' = define_lifted_perms qty_full_names (perm_names ~~ perms) raw_perm_simps thy; |
380 val lthy13 = Theory_Target.init NONE thy'; |
380 val lthy13 = Theory_Target.init NONE thy'; |
381 val q_name = space_implode "_" qty_names; |
381 val q_name = space_implode "_" qty_names; |
|
382 fun suffix_bind s = Binding.qualify true q_name (Binding.name s); |
382 val _ = tracing "Lifting induction"; |
383 val _ = tracing "Lifting induction"; |
383 val constr_names = map (Long_Name.base_name o fst o dest_Const) consts; |
384 val constr_names = map (Long_Name.base_name o fst o dest_Const) consts; |
384 val q_induct = Rule_Cases.name constr_names (lift_thm lthy13 induct); |
385 val q_induct = Rule_Cases.name constr_names (lift_thm lthy13 induct); |
385 fun note_simp_suffix s th ctxt = |
386 fun note_simp_suffix s th ctxt = |
386 snd (Local_Theory.note ((Binding.name (q_name ^ "_" ^ s), |
387 snd (Local_Theory.note ((suffix_bind s, [Attrib.internal (K Simplifier.simp_add)]), th) ctxt); |
387 [Attrib.internal (K Simplifier.simp_add)]), th) ctxt); |
388 val (_, lthy14) = Local_Theory.note ((suffix_bind "induct", |
388 val (_, lthy14) = Local_Theory.note ((Binding.name (q_name ^ "_induct"), |
|
389 [Attrib.internal (K (Rule_Cases.case_names constr_names))]), [Rule_Cases.name constr_names q_induct]) lthy13; |
389 [Attrib.internal (K (Rule_Cases.case_names constr_names))]), [Rule_Cases.name constr_names q_induct]) lthy13; |
390 val q_inducts = Project_Rule.projects lthy13 (1 upto (length alpha_inducts)) q_induct |
390 val q_inducts = Project_Rule.projects lthy13 (1 upto (length alpha_inducts)) q_induct |
391 val (_, lthy14a) = Local_Theory.note ((Binding.name (q_name ^ "_inducts"), []), q_inducts) lthy14; |
391 val (_, lthy14a) = Local_Theory.note ((suffix_bind "inducts", []), q_inducts) lthy14; |
392 val q_perm = map (lift_thm lthy14) raw_perm_def; |
392 val q_perm = map (lift_thm lthy14) raw_perm_def; |
393 val lthy15 = note_simp_suffix "perm" q_perm lthy14a; |
393 val lthy15 = note_simp_suffix "perm" q_perm lthy14a; |
394 val q_fv = map (lift_thm lthy15) fv_def; |
394 val q_fv = map (lift_thm lthy15) fv_def; |
395 val lthy16 = note_simp_suffix "fv" q_fv lthy15; |
395 val lthy16 = note_simp_suffix "fv" q_fv lthy15; |
396 val q_bn = map (lift_thm lthy16) raw_bn_eqs; |
396 val q_bn = map (lift_thm lthy16) raw_bn_eqs; |
399 val eq_iff_unfolded1 = map (Local_Defs.unfold lthy17 @{thms alpha_gen2}) alpha_eq_iff |
399 val eq_iff_unfolded1 = map (Local_Defs.unfold lthy17 @{thms alpha_gen2}) alpha_eq_iff |
400 val eq_iff_unfolded2 = map (Local_Defs.unfold lthy17 @{thms alpha_gen}) eq_iff_unfolded1 |
400 val eq_iff_unfolded2 = map (Local_Defs.unfold lthy17 @{thms alpha_gen}) eq_iff_unfolded1 |
401 val q_eq_iff_pre1 = map (lift_thm lthy17) eq_iff_unfolded2; |
401 val q_eq_iff_pre1 = map (lift_thm lthy17) eq_iff_unfolded2; |
402 val q_eq_iff_pre2 = map (Local_Defs.fold lthy17 @{thms alpha_gen2}) q_eq_iff_pre1 |
402 val q_eq_iff_pre2 = map (Local_Defs.fold lthy17 @{thms alpha_gen2}) q_eq_iff_pre1 |
403 val q_eq_iff = map (Local_Defs.fold lthy17 @{thms alpha_gen}) q_eq_iff_pre2 |
403 val q_eq_iff = map (Local_Defs.fold lthy17 @{thms alpha_gen}) q_eq_iff_pre2 |
404 val (_, lthy18) = Local_Theory.note ((Binding.name (q_name ^ "_eq_iff"), []), q_eq_iff) lthy17; |
404 val (_, lthy18) = Local_Theory.note ((suffix_bind "eq_iff", []), q_eq_iff) lthy17; |
405 val rel_dists = flat (map (distinct_rel lthy18 alpha_cases) |
405 val rel_dists = flat (map (distinct_rel lthy18 alpha_cases) |
406 (rel_distinct ~~ (List.take (alpha_ts, (length dts))))) |
406 (rel_distinct ~~ (List.take (alpha_ts, (length dts))))) |
407 val q_dis = map (lift_thm lthy18) rel_dists; |
407 val q_dis = map (lift_thm lthy18) rel_dists; |
408 val lthy19 = note_simp_suffix "distinct" q_dis lthy18; |
408 val lthy19 = note_simp_suffix "distinct" q_dis lthy18; |
409 val q_eqvt = map (lift_thm lthy19) raw_fv_bv_eqvt; |
409 val q_eqvt = map (lift_thm lthy19) raw_fv_bv_eqvt; |