blanchet@58352: (* Title: HOL/Basic_BNF_Least_Fixpoints.thy blanchet@58352: Author: Jasmin Blanchette, TU Muenchen blanchet@58352: Copyright 2014 blanchet@58352: blanchet@58352: Registration of basic types as BNF least fixpoints (datatypes). blanchet@58352: *) blanchet@58352: blanchet@58352: theory Basic_BNF_Least_Fixpoints blanchet@58352: imports BNF_Least_Fixpoint blanchet@58352: begin blanchet@58352: blanchet@58352: definition xtor :: "'a \ 'a" where blanchet@58352: "xtor x = x" blanchet@58352: blanchet@58352: lemma xtor_map: "f (xtor x) = xtor (f x)" blanchet@58352: unfolding xtor_def by (rule refl) blanchet@58352: blanchet@58352: lemma xtor_set: "f (xtor x) = f x" blanchet@58353: unfolding xtor_def by (rule refl) blanchet@58352: blanchet@58352: lemma xtor_rel: "R (xtor x) (xtor y) = R x y" blanchet@58352: unfolding xtor_def by (rule refl) blanchet@58352: blanchet@58352: lemma xtor_induct: "(\x. P (xtor x)) \ P z" blanchet@58352: unfolding xtor_def by assumption blanchet@58352: blanchet@58352: lemma xtor_xtor: "xtor (xtor x) = x" blanchet@58352: unfolding xtor_def by (rule refl) blanchet@58352: blanchet@58352: lemmas xtor_inject = xtor_rel[of "op ="] blanchet@58352: blanchet@58353: lemma xtor_rel_induct: "(\x y. vimage2p id_bnf id_bnf R x y \ IR (xtor x) (xtor y)) \ R \ IR" blanchet@58353: unfolding xtor_def vimage2p_def id_bnf_def by default blanchet@58352: blanchet@58353: lemma Inl_def_alt: "Inl \ (\a. xtor (id_bnf (Inl a)))" blanchet@58353: unfolding xtor_def id_bnf_def by (rule reflexive) blanchet@58352: blanchet@58353: lemma Inr_def_alt: "Inr \ (\a. xtor (id_bnf (Inr a)))" blanchet@58353: unfolding xtor_def id_bnf_def by (rule reflexive) blanchet@58352: blanchet@58353: lemma Pair_def_alt: "Pair \ (\a b. xtor (id_bnf (a, b)))" blanchet@58353: unfolding xtor_def id_bnf_def by (rule reflexive) blanchet@58352: blanchet@58377: definition ctor_rec :: "'a \ 'a" where blanchet@58377: "ctor_rec x = x" blanchet@58377: blanchet@58377: lemma ctor_rec: "g = id \ ctor_rec f (xtor x) = f ((id_bnf \ g \ id_bnf) x)" blanchet@58377: unfolding ctor_rec_def id_bnf_def xtor_def comp_def id_def by hypsubst (rule refl) blanchet@58377: blanchet@58377: lemma ctor_rec_def_alt: "f = ctor_rec (f \ id_bnf)" blanchet@58377: unfolding ctor_rec_def id_bnf_def comp_def by (rule refl) blanchet@58377: blanchet@58377: lemma ctor_rec_o_map: "ctor_rec f \ g = ctor_rec (f \ (id_bnf \ g \ id_bnf))" blanchet@58377: unfolding ctor_rec_def id_bnf_def comp_def by (rule refl) blanchet@58377: blanchet@58352: ML_file "Tools/BNF/bnf_lfp_basic_sugar.ML" blanchet@58352: blanchet@58377: thm sum.rec_o_map blanchet@58377: thm sum.size_o_map blanchet@58377: blanchet@58377: thm prod.rec_o_map blanchet@58377: thm prod.size_o_map blanchet@58377: blanchet@58353: hide_const (open) xtor ctor_rec blanchet@58353: blanchet@58353: hide_fact (open) blanchet@58353: xtor_def xtor_map xtor_set xtor_rel xtor_induct xtor_xtor xtor_inject ctor_rec_def ctor_rec blanchet@58377: ctor_rec_def_alt ctor_rec_o_map xtor_rel_induct Inl_def_alt Inr_def_alt Pair_def_alt blanchet@58353: blanchet@58352: end