# HG changeset patch # User wenzelm # Date 1439466952 -7200 # Node ID 6584c0f3f0e06424de0039dd0c389e86d691e2eb # Parent 61a7f9bb9e6bc5d6dda70d8322f73742e2fe4dfb# Parent 0ccb5fb83c24ef774981d9af2210a266553d69d1 merged diff -r 0ccb5fb83c24 -r 6584c0f3f0e0 CONTRIBUTORS --- a/CONTRIBUTORS Thu Aug 13 13:55:48 2015 +0200 +++ b/CONTRIBUTORS Thu Aug 13 13:55:52 2015 +0200 @@ -21,6 +21,10 @@ * Summer 2015: Florian Haftmann, TUM Fundamentals of abstract type class for factorial rings. +* Summer 2015: Julian Biendarra, TUM and Dmitriy Traytel, ETH Zurich + Command to lift a BNF structure on the raw type to the abstract type + for typedefs. + Contributions to Isabelle2015 ----------------------------- diff -r 0ccb5fb83c24 -r 6584c0f3f0e0 NEWS --- a/NEWS Thu Aug 13 13:55:48 2015 +0200 +++ b/NEWS Thu Aug 13 13:55:52 2015 +0200 @@ -188,6 +188,9 @@ * Nitpick: - Removed "check_potential" and "check_genuine" options. +* New commands lift_bnf and copy_bnf for lifting (copying) a BNF structure + on the raw type to an abstract type defined using typedef. + * Division on integers is bootstrapped directly from division on naturals and uses generic numeral algorithm for computations. Slight INCOMPATIBILITY, simproc numeral_divmod replaces and generalizes diff -r 0ccb5fb83c24 -r 6584c0f3f0e0 src/Doc/Datatypes/Datatypes.thy --- a/src/Doc/Datatypes/Datatypes.thy Thu Aug 13 13:55:48 2015 +0200 +++ b/src/Doc/Datatypes/Datatypes.thy Thu Aug 13 13:55:52 2015 +0200 @@ -2771,6 +2771,65 @@ for further examples of BNF registration, some of which feature custom witnesses. +For many typedefs (in particular for type copies) lifting the BNF structure +from the raw type to the abstract type can be done uniformly. This is the task +of the @{command lift_bnf} command. Using @{command lift_bnf} the above +registration of @{typ "('d, 'a) fn"} as a BNF becomes much shorter. +*} + + (*<*) context early begin (*>*) + lift_bnf (*<*)(no_warn_wits)(*>*) ('d, 'a) fn by auto + (*<*) end (*>*) + +text {* +Indeed, for type copies the proof obligations are so simple that they can be +discharged automatically, yielding another command @{command copy_bnf} which +does not issue proof obligations. +*} + + (*<*) context late begin (*>*) + copy_bnf ('d, 'a) fn + (*<*) end (*>*) + +text {* +Since records (or rather record schemes) are particular type copies, +the @{command copy_bnf} can be used to register records as BNFs. +*} + + record 'a point = + xval :: 'a + yval :: 'a + + copy_bnf ('a, 'z) point_ext + +text {* +The proof obligations handed over to the user by @{command lift_bnf} in +the general case are simpler than the acual BNF properties (in particular, +no cardinality reasoning is required). They are best illustrated on an +example. Suppose we define the type of nonempty lists. +*} + + typedef 'a nonempty_list = "{xs :: 'a list. xs \ []}" by auto + +text {* +Then, @{command lift_bnf} requires us to prove that the set of nonempty lists +is closed under the map function and the zip function (where the latter only +occurs implicitly in the goal, in form of the variable +@{term "zs :: ('a \ 'b) list"}). +*} + + lift_bnf (*<*)(no_warn_wits)(*>*) 'a nonempty_list + proof - + fix f and xs :: "'a list" + assume "xs \ {xs. xs \ []}" + then show "map f xs \ {xs. xs \ []}" by (cases xs(*<*)rule: list.exhaust(*>*)) auto + next + fix zs :: "('a \ 'b) list" + assume "map fst zs \ {xs. xs \ []}" "map snd zs \ {xs. xs \ []}" + then show "zs \ {xs. xs \ []}" by (cases zs(*<*)rule: list.exhaust(*>*)) auto + qed + +text {* The next example declares a BNF axiomatically. This can be convenient for reasoning abstractly about an arbitrary BNF. The @{command bnf_axiomatization} command below introduces a type @{text "('a, 'b, 'c) F"}, three set constants, @@ -2825,6 +2884,48 @@ %%% TODO: elaborate on proof obligations *} +subsubsection {* \keyw{lift_bnf} and \keyw{copy_bnf} + \label{sssec:lift-bnf} *} + +text {* +\begin{matharray}{rcl} + @{command_def "lift_bnf"} & : & @{text "local_theory \ proof(prove)"} \\ + @{command_def "copy_bnf"} & : & @{text "local_theory \ local_theory"} +\end{matharray} + +@{rail \ + @@{command lift_bnf} target? lb_options? \ + @{syntax tyargs} name wit_terms? \ + ('via' @{syntax thmref})? @{syntax map_rel}? + ; + @{syntax_def lb_options}: '(' ((@{syntax plugins} | 'no_warn_wits') + ',') ')' + ; + @{syntax_def wit_terms}: '[' 'wits' ':' terms ']' + ; + @@{command copy_bnf} target? ('(' @{syntax plugins} ')')? \ + @{syntax tyargs} name ('via' @{syntax thmref})? @{syntax map_rel}? +\} +\medskip + +\noindent +The @{command lift_bnf} command registers an existing type (abstract type), +defined as a subtype of a BNF (raw type) using the @{command typedef} command, +as a BNF. To do so, it lifts the BNF structure on the raw type to the abstract +type following a @{term type_definition} theorem (the theorem is usually inferred +from the type, but can also be explicitly supplied by the user using the +@{text via} slot). In addition, custom names for map, set functions, and the relator, +as well as nonemptiness witnesses can be specified; otherwise, default versions are used. +Nonemptiness witnesses are not lifted from the raw type's BNF (this would be +inherently incomplete), but must be given as terms (on the raw type) and proved +to be witnesses. The command warns aggresively about witness types that a present +in the raw type's BNF but not supplied by the user. The warning can be turned off +by passing the @{text no_warn_wits} option. + +The @{command copy_bnf} command, performes the same lifting for type copies +(@{command typedef}s with @{term UNIV} as the representing set) without bothering +the user with trivial proof obligations. (Here, all nonemptiness witnesses from the raw +type's BNF can also be lifted without problems.) +*} subsubsection {* \keyw{bnf_axiomatization} \label{sssec:bnf-axiomatization} *} diff -r 0ccb5fb83c24 -r 6584c0f3f0e0 src/HOL/BNF_Composition.thy --- a/src/HOL/BNF_Composition.thy Thu Aug 13 13:55:48 2015 +0200 +++ b/src/HOL/BNF_Composition.thy Thu Aug 13 13:55:52 2015 +0200 @@ -10,8 +10,14 @@ theory BNF_Composition imports BNF_Def +keywords + "copy_bnf" :: thy_decl and + "lift_bnf" :: thy_goal begin +lemma ssubst_mem: "\t = s; s \ X\ \ t \ X" + by simp + lemma empty_natural: "(\_. {}) o f = image g o (\_. {})" by (rule ext) simp @@ -168,6 +174,7 @@ ML_file "Tools/BNF/bnf_comp_tactics.ML" ML_file "Tools/BNF/bnf_comp.ML" +ML_file "Tools/BNF/bnf_lift.ML" hide_fact DEADID.inj_map DEADID.inj_map_strong DEADID.map_comp DEADID.map_cong DEADID.map_cong0 diff -r 0ccb5fb83c24 -r 6584c0f3f0e0 src/HOL/BNF_Fixpoint_Base.thy --- a/src/HOL/BNF_Fixpoint_Base.thy Thu Aug 13 13:55:48 2015 +0200 +++ b/src/HOL/BNF_Fixpoint_Base.thy Thu Aug 13 13:55:52 2015 +0200 @@ -55,9 +55,6 @@ thus "EX a. b = f a" by blast qed -lemma ssubst_mem: "\t = s; s \ X\ \ t \ X" - by simp - lemma case_sum_step: "case_sum (case_sum f' g') g (Inl p) = case_sum f' g' p" "case_sum f (case_sum f' g') (Inr p) = case_sum f' g' p" diff -r 0ccb5fb83c24 -r 6584c0f3f0e0 src/HOL/Datatype_Examples/Lift_BNF.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Datatype_Examples/Lift_BNF.thy Thu Aug 13 13:55:52 2015 +0200 @@ -0,0 +1,75 @@ +theory Lift_BNF +imports Main +begin + +typedef 'a nonempty_list = "{xs :: 'a list. xs \ []}" + by blast + +lift_bnf (no_warn_wits) (neset: 'a) nonempty_list + for map: nemap rel: nerel + by simp_all + +typedef ('a :: finite, 'b) fin_nonempty_list = "{(xs :: 'a set, ys :: 'b list). ys \ []}" + by blast + +lift_bnf (dead 'a :: finite, 'b) fin_nonempty_list + by auto + +datatype 'a tree = Leaf | Node 'a "'a tree nonempty_list" + +record 'a point = + xCoord :: 'a + yCoord :: 'a + +copy_bnf ('a, 's) point_ext + +typedef 'a it = "UNIV :: 'a set" + by blast + +copy_bnf (plugins del: size) 'a it + +typedef ('a, 'b) T_prod = "UNIV :: ('a \ 'b) set" + by blast + +copy_bnf ('a, 'b) T_prod + +typedef ('a, 'b, 'c) T_func = "UNIV :: ('a \ 'b * 'c) set" + by blast + +copy_bnf ('a, 'b, 'c) T_func + +typedef ('a, 'b) sum_copy = "UNIV :: ('a + 'b) set" + by blast + +copy_bnf ('a, 'b) sum_copy + +typedef ('a, 'b) T_sum = "{Inl x | x. True} :: ('a + 'b) set" + by blast + +lift_bnf (no_warn_wits) ('a, 'b) T_sum [wits: "Inl :: 'a \ 'a + 'b"] + by (auto simp: map_sum_def sum_set_defs split: sum.splits) + +typedef ('key, 'value) alist = "{xs :: ('key \ 'value) list. (distinct \ map fst) xs}" + morphisms impl_of Alist +proof + show "[] \ {xs. (distinct o map fst) xs}" + by simp +qed + +lift_bnf (dead 'k, 'v) alist [wits: "Nil :: ('k \ 'v) list"] + by simp_all + +typedef 'a myopt = "{X :: 'a set. finite X \ card X \ 1}" by (rule exI[of _ "{}"]) auto +lemma myopt_type_def: "type_definition + (\X. if card (Rep_myopt X) = 1 then Some (the_elem (Rep_myopt X)) else None) + (\x. Abs_myopt (case x of Some x \ {x} | _ \ {})) + (UNIV :: 'a option set)" + apply unfold_locales + apply (auto simp: Abs_myopt_inverse dest!: card_eq_SucD split: option.splits) + apply (metis Rep_myopt_inverse) + apply (metis One_nat_def Rep_myopt Rep_myopt_inverse Suc_le_mono card_0_eq le0 le_antisym mem_Collect_eq nat.exhaust) + done + +copy_bnf 'a myopt via myopt_type_def + +end diff -r 0ccb5fb83c24 -r 6584c0f3f0e0 src/HOL/Library/DAList.thy --- a/src/HOL/Library/DAList.thy Thu Aug 13 13:55:48 2015 +0200 +++ b/src/HOL/Library/DAList.thy Thu Aug 13 13:55:52 2015 +0200 @@ -198,54 +198,8 @@ section \alist is a BNF\ -lift_definition map :: "('a \ 'b) \ ('k, 'a) alist \ ('k, 'b) alist" - is "\f xs. List.map (map_prod id f) xs" by simp - -lift_definition set :: "('k, 'v) alist => 'v set" - is "\xs. snd ` List.set xs" . - -lift_definition rel :: "('a \ 'b \ bool) \ ('k, 'a) alist \ ('k, 'b) alist \ bool" - is "\R xs ys. list_all2 (rel_prod op = R) xs ys" . - -bnf "('k, 'v) alist" - map: map - sets: set - bd: natLeq - wits: empty - rel: rel -proof (unfold OO_Grp_alt) - show "map id = id" by (rule ext, transfer) (simp add: prod.map_id0) -next - fix f g - show "map (g \ f) = map g \ map f" - by (rule ext, transfer) (simp add: prod.map_comp) -next - fix x f g - assume "(\z. z \ set x \ f z = g z)" - then show "map f x = map g x" by transfer force -next - fix f - show "set \ map f = op ` f \ set" - by (rule ext, transfer) (simp add: image_image) -next - fix x - show "ordLeq3 (card_of (set x)) natLeq" - by transfer (auto simp: finite_iff_ordLess_natLeq[symmetric] intro: ordLess_imp_ordLeq) -next - fix R S - show "rel R OO rel S \ rel (R OO S)" - by (rule predicate2I, transfer) - (auto simp: list.rel_compp prod.rel_compp[of "op =", unfolded eq_OO]) -next - fix R - show "rel R = (\x y. \z. z \ {x. set x \ {(x, y). R x y}} \ map fst z = x \ map snd z = y)" - unfolding fun_eq_iff by transfer (fastforce simp: list.in_rel o_def intro: - exI[of _ "List.map (\p. ((fst p, fst (snd p)), (fst p, snd (snd p)))) z" for z] - exI[of _ "List.map (\p. (fst (fst p), snd (fst p), snd (snd p))) z" for z]) -next - fix z assume "z \ set empty" - then show False by transfer simp -qed (simp_all add: natLeq_cinfinite natLeq_card_order) +lift_bnf (dead 'k, set: 'v) alist [wits: "[] :: ('k \ 'v) list"] for map: map rel: rel + by auto hide_const valterm_empty valterm_update random_aux_alist diff -r 0ccb5fb83c24 -r 6584c0f3f0e0 src/HOL/ROOT --- a/src/HOL/ROOT Thu Aug 13 13:55:48 2015 +0200 +++ b/src/HOL/ROOT Thu Aug 13 13:55:52 2015 +0200 @@ -781,6 +781,7 @@ "Derivation_Trees/Gram_Lang" "Derivation_Trees/Parallel" Koenig + Lift_BNF Stream_Processor Misc_Codatatype Misc_Datatype diff -r 0ccb5fb83c24 -r 6584c0f3f0e0 src/HOL/Tools/BNF/bnf_def.ML --- a/src/HOL/Tools/BNF/bnf_def.ML Thu Aug 13 13:55:48 2015 +0200 +++ b/src/HOL/Tools/BNF/bnf_def.ML Thu Aug 13 13:55:52 2015 +0200 @@ -104,6 +104,7 @@ 'a list val mk_witness: int list * term -> thm list -> nonemptiness_witness + val mk_wit_goals: term list -> term list -> term list -> int list * term -> term list val minimize_wits: (''a list * 'b) list -> (''a list * 'b) list val wits_of_bnf: bnf -> nonemptiness_witness list @@ -759,6 +760,22 @@ |>> (fn [] => bnf | noted => morph_bnf (substitute_noted_thm noted) bnf) end; +fun mk_wit_goals zs bs sets (I, wit) = + let + val xs = map (nth bs) I; + fun wit_goal i = + let + val z = nth zs i; + val set_wit = nth sets i $ Term.list_comb (wit, xs); + val concl = HOLogic.mk_Trueprop + (if member (op =) I i then HOLogic.mk_eq (z, nth bs i) else @{term False}); + in + fold_rev Logic.all (z :: xs) (Logic.mk_implies (mk_Trueprop_mem (z, set_wit), concl)) + end; + in + map wit_goal (0 upto length sets - 1) + end; + (* Define new BNFs *) @@ -1117,22 +1134,7 @@ val goals = zip_axioms map_id0_goal map_comp0_goal map_cong0_goal set_map0s_goal card_order_bd_goal cinfinite_bd_goal set_bds_goal le_rel_OO_goal rel_OO_Grp_goal; - fun mk_wit_goals (I, wit) = - let - val xs = map (nth bs) I; - fun wit_goal i = - let - val z = nth zs i; - val set_wit = nth bnf_sets_As i $ Term.list_comb (wit, xs); - val concl = HOLogic.mk_Trueprop - (if member (op =) I i then HOLogic.mk_eq (z, nth bs i) else @{term False}); - in - fold_rev Logic.all (z :: xs) (Logic.mk_implies (mk_Trueprop_mem (z, set_wit), concl)) - end; - in - map wit_goal (0 upto live - 1) - end; - + val mk_wit_goals = mk_wit_goals bs zs bnf_sets_As; fun triv_wit_tac ctxt = mk_trivial_wit_tac ctxt bnf_wit_defs; val wit_goalss = diff -r 0ccb5fb83c24 -r 6584c0f3f0e0 src/HOL/Tools/BNF/bnf_lift.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/BNF/bnf_lift.ML Thu Aug 13 13:55:52 2015 +0200 @@ -0,0 +1,391 @@ +(* Title: HOL/Tools/BNF/bnf_lift.ML + Author: Julian Biendarra, TU Muenchen + Author: Dmitriy Traytel, ETH Zurich + Copyright 2015 + +Lifting of BNFs through typedefs. +*) + +signature BNF_LIFT = sig + datatype lift_bnf_option = Plugins_Option of Proof.context -> Plugin_Name.filter | No_Warn_Wits + + val copy_bnf: + (((lift_bnf_option list * (binding option * (string * sort option)) list) * + string) * thm option) * (binding * binding) -> + local_theory -> local_theory + val copy_bnf_cmd: + (((lift_bnf_option list * (binding option * (string * string option)) list) * + string) * (Facts.ref * Token.src list) option) * (binding * binding) -> + local_theory -> local_theory + val lift_bnf: + (((lift_bnf_option list * (binding option * (string * sort option)) list) * + string) * thm option) * (binding * binding) -> + ({context: Proof.context, prems: thm list} -> tactic) list -> + local_theory -> local_theory + val lift_bnf_cmd: + ((((lift_bnf_option list * (binding option * (string * string option)) list) * + string) * string list) * (Facts.ref * Token.src list) option) * (binding * binding) -> + local_theory -> Proof.state + end + +structure BNF_Lift : BNF_LIFT = struct + +open Ctr_Sugar_Tactics +open BNF_Util +open BNF_Comp +open BNF_Def + +datatype lift_bnf_option = Plugins_Option of Proof.context -> Plugin_Name.filter | No_Warn_Wits + +fun typedef_bnf thm wits specs map_b rel_b opts lthy = + let + val plugins = get_first (fn Plugins_Option f => SOME (f lthy) | _ => NONE) (rev opts) + |> the_default Plugin_Name.default_filter; + val no_warn_wits = exists (can (fn Sequential_Option => ())) opts; + + (* extract Rep Abs F RepT AbsT *) + val (_, [Rep_G, Abs_G, F]) = Thm.prop_of thm + |> HOLogic.dest_Trueprop + |> Term.strip_comb; + val typ_Abs_G = fastype_of Abs_G |> dest_funT; + val RepT = fst typ_Abs_G; (* F *) + val AbsT = snd typ_Abs_G; (* G *) + val AbsT_name = fst (dest_Type AbsT); + val tvs = AbsT |> dest_Type |> snd |> map (fst o dest_TVar); + val alpha0s = map (TFree o snd) specs; + + (* instantiate the new type variables newtvs to oldtvs *) + val subst = subst_TVars (tvs ~~ alpha0s); + val typ_subst = typ_subst_TVars (tvs ~~ alpha0s); + + val Rep_G = subst Rep_G; + val Abs_G = subst Abs_G; + val F = subst F; + val RepT = typ_subst RepT; + val AbsT = typ_subst AbsT; + + fun flatten_tyargs Ass = map dest_TFree alpha0s |> + filter (fn T => exists (fn Ts => member (op =) Ts T) Ass); + + val Ds0 = filter (is_none o fst) specs |> map snd; + + (* get the bnf for RepT *) + val ((bnf, (deads, alphas)),((_, unfolds), lthy)) = + bnf_of_typ Dont_Inline (Binding.qualify true AbsT_name) flatten_tyargs [] + Ds0 RepT ((empty_comp_cache, empty_unfolds), lthy); + + val set_bs = map (fn T => find_index (fn U => T = U) alpha0s) alphas + |> map (the_default Binding.empty o fst o nth specs); + + val _ = case alphas of [] => error "No live variables." | alphas => alphas; + + val defs = #map_unfolds unfolds @ flat (#set_unfoldss unfolds) @ #rel_unfolds unfolds; + + (* number of live variables *) + val lives = length alphas; + + (* state the three required properties *) + val sorts = map Type.sort_of_atyp alphas; + val names_lthy = fold Variable.declare_typ (alphas @ deads) lthy; + val (alphas', names_lthy) = mk_TFrees' sorts names_lthy; + val (betas, names_lthy) = mk_TFrees' sorts names_lthy; + + val map_F = mk_map_of_bnf deads alphas betas bnf; + + val (typ_fs, typ_aF) = fastype_of map_F |> strip_typeN lives ||> domain_type; + val typ_pairs = map HOLogic.mk_prodT (alphas ~~ alphas'); + val typ_subst_pair = typ_subst_atomic (alphas ~~ typ_pairs); + val typ_pair = typ_subst_pair RepT; + + val subst_b = subst_atomic_types (alphas ~~ betas); + val subst_a' = subst_atomic_types (alphas ~~ alphas'); + val subst_pair = subst_atomic_types (alphas ~~ typ_pairs); + val aF_set = F; + val bF_set = subst_b F; + val aF_set' = subst_a' F; + val pairF_set = subst_pair F; + val map_F_fst = mk_map_of_bnf deads typ_pairs alphas bnf; + val map_F_snd = mk_map_of_bnf deads typ_pairs alphas' bnf; + val wits_F = mk_wits_of_bnf + (replicate (nwits_of_bnf bnf) deads) + (replicate (nwits_of_bnf bnf) alphas) bnf; + + (* val map_closed_F = @{term "\f x. x \ F \ map_F f x \ F"}; *) + val (var_fs, names_lthy) = mk_Frees "f" typ_fs names_lthy; + val (var_x, names_lthy) = mk_Frees "x" [typ_aF] names_lthy |>> the_single; + val mem_x = HOLogic.mk_mem (var_x, aF_set) |> HOLogic.mk_Trueprop; + val map_f = list_comb (map_F, var_fs); + val mem_map = HOLogic.mk_mem (map_f $ var_x, bF_set) |> HOLogic.mk_Trueprop; + val imp_map = Logic.mk_implies (mem_x, mem_map); + val map_closed_F = Library.foldr (Library.uncurry Logic.all) (var_fs, Logic.all var_x imp_map); + + (* val zip_closed_F = @{term "\z. map_F fst z \ F \ map_F snd z \ F \ z \ F"}; *) + val (var_zs, names_lthy) = mk_Frees "z" [typ_pair] names_lthy; + val (pairs, names_lthy) = mk_Frees "tmp" typ_pairs names_lthy; + val var_z = hd var_zs; + val fsts = map (fst o Term.strip_comb o HOLogic.mk_fst) pairs; + val snds = map (fst o Term.strip_comb o HOLogic.mk_snd) pairs; + val map_fst = list_comb (list_comb (map_F_fst, fsts), var_zs); + val mem_map_fst = HOLogic.mk_mem (map_fst, aF_set) |> HOLogic.mk_Trueprop; + val map_snd = list_comb (list_comb (map_F_snd, snds), var_zs); + val mem_map_snd = HOLogic.mk_mem (map_snd, aF_set') |> HOLogic.mk_Trueprop; + val mem_z = HOLogic.mk_mem (var_z, pairF_set) |> HOLogic.mk_Trueprop; + val imp_zip = Logic.mk_implies (mem_map_fst, Logic.mk_implies (mem_map_snd, mem_z)); + val zip_closed_F = Logic.all var_z imp_zip; + + (* val wit_closed_F = @{term "wit_F a \ F"}; *) + val (var_as, names_lthy) = mk_Frees "a" alphas names_lthy; + val (var_bs, _) = mk_Frees "a" alphas names_lthy; + val Iwits = the_default wits_F (Option.map (map (`(map (fn T => + find_index (fn U => T = U) alphas) o fst o strip_type o fastype_of))) wits); + val wit_closed_Fs = + map (fn (I, wit_F) => + let + val vars = map (nth var_as) I; + val wit_a = list_comb (wit_F, vars); + in + Library.foldr (Library.uncurry Logic.all) (vars, + HOLogic.mk_mem (wit_a, aF_set) |> HOLogic.mk_Trueprop) + end) + Iwits; + + val mk_wit_goals = mk_wit_goals var_as var_bs + (mk_sets_of_bnf (replicate lives deads) (replicate lives alphas) bnf); + + val goals = [map_closed_F, zip_closed_F] @ wit_closed_Fs @ + (case wits of NONE => [] | _ => maps mk_wit_goals Iwits); + + val lost_wits = filter_out (fn (J, _) => exists (fn (I, _) => I = J) Iwits) wits_F; + val _ = if null lost_wits orelse no_warn_wits then () else + lost_wits + |> map (Syntax.pretty_typ lthy o fastype_of o snd) + |> Pretty.big_list + "The following types of nonemptiness witnesses of the raw type's BNF were lost:" + |> (fn pt => Pretty.chunks [pt, + Pretty.para "You can specify a liftable witness (e.g., a term of one of the above types\ + \ that satisfies the typedef's invariant)\ + \ using the annotation [wits: ]."]) + |> Pretty.string_of + |> warning; + + fun after_qed ([map_closed_thm] :: [zip_closed_thm] :: wit_thmss) lthy = + let + val (wit_closed_thms, wit_thms) = + (case wits of + NONE => (map the_single wit_thmss, wit_thms_of_bnf bnf) + | _ => chop (length wit_closed_Fs) (map the_single wit_thmss)) + + (* construct map set bd rel wit *) + (* val map_G = @{term "\f. Abs_G o map_F f o Rep_G"}; *) + val Abs_Gb = subst_b Abs_G; + val map_G = Library.foldr (uncurry HOLogic.tupled_lambda) + (var_fs, HOLogic.mk_comp (HOLogic.mk_comp (Abs_Gb, map_f), + Rep_G)); + + (* val sets_G = [@{term "set_F o Rep_G"}]; *) + val sets_F = mk_sets_of_bnf (replicate lives deads) (replicate lives alphas) bnf; + val sets_G = map (fn set_F => HOLogic.mk_comp (set_F, Rep_G)) sets_F; + + (* val bd_G = @{term "bd_F"}; *) + val bd_F = mk_bd_of_bnf deads alphas bnf; + val bd_G = bd_F; + + (* val rel_G = @{term "\R. BNF_Def.vimage2p Rep_G Rep_G (rel_F R)"}; *) + val rel_F = mk_rel_of_bnf deads alphas betas bnf; + val (typ_Rs, _) = fastype_of rel_F |> strip_typeN lives; + + val (var_Rs, names_lthy) = mk_Frees "R" typ_Rs lthy; + val Rep_Gb = subst_b Rep_G; + val rel_G = fold_rev absfree (map dest_Free var_Rs) + (mk_vimage2p Rep_G Rep_Gb $ list_comb (rel_F, var_Rs)); + + (* val wits_G = [@{term "Abs_G o wit_F"}]; *) + val (var_as, _) = mk_Frees "a" alphas names_lthy; + val wits_G = + map (fn (I, wit_F) => + let + val vs = map (nth var_as) I; + in fold_rev absfree (map dest_Free vs) (Abs_G $ (list_comb (wit_F, vs))) end) + Iwits; + + (* tactics *) + val Rep_thm = thm RS @{thm type_definition.Rep}; + val Abs_inverse_thm = thm RS @{thm type_definition.Abs_inverse}; + val Abs_inject_thm = thm RS @{thm type_definition.Abs_inject}; + val Rep_cases_thm = thm RS @{thm type_definition.Rep_cases}; + val Rep_inverse_thm = thm RS @{thm type_definition.Rep_inverse}; + + fun map_id0_tac ctxt = + HEADGOAL (EVERY' [rtac ctxt ext, + SELECT_GOAL (unfold_thms_tac ctxt [map_id0_of_bnf bnf, id_apply, o_apply, + Rep_inverse_thm]), + rtac ctxt refl]); + + fun map_comp0_tac ctxt = + HEADGOAL (EVERY' [rtac ctxt ext, + SELECT_GOAL (unfold_thms_tac ctxt [map_comp0_of_bnf bnf, o_apply, + Rep_thm RS (map_closed_thm RS Abs_inverse_thm)]), + rtac ctxt refl]); + + fun map_cong0_tac ctxt = + HEADGOAL (EVERY' ([SELECT_GOAL (unfold_thms_tac ctxt [o_apply]), + rtac ctxt (([Rep_thm RS map_closed_thm, Rep_thm RS map_closed_thm] MRS + Abs_inject_thm) RS iffD2), + rtac ctxt (map_cong0_of_bnf bnf)] @ replicate lives (Goal.assume_rule_tac ctxt))); + + val set_map0s_tac = + map (fn set_map => fn ctxt => + HEADGOAL (EVERY' [rtac ctxt ext, + SELECT_GOAL (unfold_thms_tac ctxt [set_map, o_apply, + Rep_thm RS (map_closed_thm RS Abs_inverse_thm)]), + rtac ctxt refl])) + (set_map_of_bnf bnf); + + fun card_order_bd_tac ctxt = HEADGOAL (rtac ctxt (bd_card_order_of_bnf bnf)); + + fun cinfinite_bd_tac ctxt = HEADGOAL (rtac ctxt (bd_cinfinite_of_bnf bnf)); + + val set_bds_tac = + map (fn set_bd => fn ctxt => + HEADGOAL (EVERY' [SELECT_GOAL (unfold_thms_tac ctxt [o_apply]), rtac ctxt set_bd])) + (set_bd_of_bnf bnf); + + fun le_rel_OO_tac ctxt = + HEADGOAL (EVERY' [rtac ctxt @{thm vimage2p_relcompp_mono}, + rtac ctxt ((rel_OO_of_bnf bnf RS sym) RS @{thm ord_eq_le_trans}), + rtac ctxt @{thm order_refl}]); + + fun rel_OO_Grp_tac ctxt = + HEADGOAL (EVERY' ([SELECT_GOAL (REPEAT_DETERM (HEADGOAL (rtac ctxt ext))), + SELECT_GOAL (unfold_thms_tac ctxt [@{thm OO_Grp_alt}, mem_Collect_eq, + o_apply, @{thm vimage2p_def}, in_rel_of_bnf bnf, Bex_def, mem_Collect_eq]), + rtac ctxt iffI, + SELECT_GOAL (REPEAT_DETERM (HEADGOAL (eresolve0_tac [exE,conjE]))), + rtac ctxt (zip_closed_thm OF (replicate 2 (Rep_thm RSN (2, @{thm ssubst_mem}))) RS + Rep_cases_thm), + assume_tac ctxt, + assume_tac ctxt, + hyp_subst_tac ctxt, + SELECT_GOAL (REPEAT_DETERM (HEADGOAL (rtac ctxt exI))), + rtac ctxt conjI] @ + replicate (lives - 1) (rtac ctxt conjI THEN' assume_tac ctxt) @ + [assume_tac ctxt, + SELECT_GOAL (REPEAT_DETERM (HEADGOAL (rtac ctxt conjI))), + REPEAT_DETERM_N 2 o + etac ctxt (trans OF [iffD2 OF [Abs_inject_thm OF + [map_closed_thm OF [Rep_thm], Rep_thm]], Rep_inverse_thm]), + SELECT_GOAL (REPEAT_DETERM (HEADGOAL (eresolve0_tac [exE,conjE]))), + rtac ctxt exI, + rtac ctxt conjI] @ + replicate (lives - 1) (rtac ctxt conjI THEN' assume_tac ctxt) @ + [assume_tac ctxt, + rtac ctxt conjI, + REPEAT_DETERM_N 2 o EVERY' + [rtac ctxt (iffD1 OF [Abs_inject_thm OF [map_closed_thm OF [Rep_thm], Rep_thm]]), + etac ctxt (Rep_inverse_thm RS sym RSN (2, trans))]])); + + fun wit_tac ctxt = + HEADGOAL (EVERY' + (map (fn thm => (EVERY' + [SELECT_GOAL (unfold_thms_tac ctxt (o_apply :: + (wit_closed_thms RL [Abs_inverse_thm]))), + dtac ctxt thm, assume_tac ctxt])) + wit_thms)); + + val tactics = [map_id0_tac, map_comp0_tac, map_cong0_tac] @ set_map0s_tac @ + [card_order_bd_tac, cinfinite_bd_tac] @ set_bds_tac @ [le_rel_OO_tac, rel_OO_Grp_tac]; + + val (bnf, lthy) = bnf_def Dont_Inline (user_policy Note_Some) false I + tactics wit_tac NONE map_b rel_b set_bs + ((((((Binding.empty, AbsT), map_G), sets_G), bd_G), wits_G), SOME rel_G) + lthy; + in + lthy |> BNF_Def.register_bnf plugins AbsT_name bnf + end + | after_qed _ _ = error "should not happen"; + in + (goals, after_qed, defs, lthy) + end; + +fun prepare_common prepare_name prepare_sort prepare_term prepare_thm + (((((plugins, raw_specs), raw_Tname), raw_wits), xthm_opt), (map_b, rel_b)) lthy = + let + val Tname = prepare_name lthy raw_Tname; + val input_thm = + (case xthm_opt of + SOME xthm => prepare_thm lthy xthm + | NONE => Typedef.get_info lthy Tname |> hd |> snd |> #type_definition); + val wits = Option.map (map (prepare_term lthy)) raw_wits; + val specs = map (apsnd (apsnd + (the_default @{sort type} o Option.map (prepare_sort lthy)))) raw_specs; + + (* analyze theorem here*) + fun is_typedef (t as (Const ("Typedef.type_definition", _) $ _ $ _ $ _)) = t + | is_typedef t = raise TERM("not a typedef",[t]); + + val _ = (HOLogic.dest_Trueprop o Thm.prop_of) input_thm |> is_typedef + handle TERM _ => error "Unsupported type of a theorem. Only type_definition is supported."; + in + typedef_bnf input_thm wits specs map_b rel_b plugins lthy + end; + +fun prepare_lift_bnf prepare_name prepare_sort prepare_term prepare_thm = + (fn (goals, after_qed, definitions, lthy) => + lthy + |> Proof.theorem NONE after_qed (map (single o rpair []) goals) + |> Proof.refine (Method.Basic (fn ctxt => SIMPLE_METHOD (unfold_thms_tac ctxt definitions))) + |> Seq.hd + |> Proof.refine (Method.primitive_text (K I)) + |> Seq.hd) oo + prepare_common prepare_name prepare_sort prepare_term prepare_thm o apfst (apfst (apsnd SOME)); + +val lift_bnf_cmd = prepare_lift_bnf + (fst o dest_Type oo Proof_Context.read_type_name {proper = true, strict = false}) + Syntax.read_sort Syntax.read_term (singleton o Attrib.eval_thms); + +fun prepare_solve prepare_name prepare_typ prepare_sort prepare_thm tacs = + (fn (goals, after_qed, _, lthy) => + lthy + |> after_qed (map2 (single oo Goal.prove lthy [] []) goals (tacs (length goals)))) oo + prepare_common prepare_name prepare_typ prepare_sort prepare_thm o apfst (apfst (rpair NONE)); + +fun lift_bnf args tacs = prepare_solve (K I) (K I) (K I) (K I) (K tacs) args; +val copy_bnf = prepare_solve (K I) (K I) (K I) (K I) + (fn n => replicate n (fn {context = ctxt, prems = _} => rtac ctxt UNIV_I 1)); +val copy_bnf_cmd = prepare_solve + (fst o dest_Type oo Proof_Context.read_type_name {proper = true, strict = false}) + Syntax.read_sort Syntax.read_term (singleton o Attrib.eval_thms) + (fn n => replicate n (fn {context = ctxt, prems = _} => rtac ctxt UNIV_I 1)); + +val parse_wits = + @{keyword "["} |-- (Parse.name --| @{keyword ":"} -- Scan.repeat Parse.term >> + (fn ("wits", Ts) => Ts + | (s, _) => error ("Unknown label " ^ quote s ^ " (expected \"wits\")"))) --| + @{keyword "]"} || Scan.succeed []; + +val parse_options = + Scan.optional (@{keyword "("} |-- + Parse.list1 (Parse.group (K "option") + (Plugin_Name.parse_filter >> Plugins_Option + || Parse.reserved "no_warn_wits" >> K No_Warn_Wits)) + --| @{keyword ")"}) []; + +val parse_plugins = + Scan.optional (@{keyword "("} |-- Plugin_Name.parse_filter --| @{keyword ")"}) + (K Plugin_Name.default_filter) >> Plugins_Option >> single; + +val parse_typedef_thm = Scan.option (Parse.reserved "via" |-- Parse.xthm); + +val _ = + Outer_Syntax.local_theory_to_proof @{command_keyword lift_bnf} + "register a subtype of a bounded natural functor (BNF) as a BNF" + ((parse_options -- parse_type_args_named_constrained -- Parse.type_const -- parse_wits -- + parse_typedef_thm -- parse_map_rel_bindings) >> lift_bnf_cmd); + +val _ = + Outer_Syntax.local_theory @{command_keyword copy_bnf} + "register a type copy of a bounded natural functor (BNF) as a BNF" + ((parse_plugins -- parse_type_args_named_constrained -- Parse.type_const -- + parse_typedef_thm -- parse_map_rel_bindings) >> copy_bnf_cmd); + +end \ No newline at end of file