| author | wenzelm | 
| Mon, 20 Sep 2021 23:15:02 +0200 | |
| changeset 74333 | a9b20bc32fa6 | 
| parent 69850 | 5f993636ac07 | 
| permissions | -rw-r--r-- | 
| 59141 | 1 | (* Title: HOL/Basic_BNF_LFPs.thy | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 2 | Author: Jasmin Blanchette, TU Muenchen | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 3 | Copyright 2014 | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 4 | |
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 5 | Registration of basic types as BNF least fixpoints (datatypes). | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 6 | *) | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 7 | |
| 59141 | 8 | theory Basic_BNF_LFPs | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 9 | imports BNF_Least_Fixpoint | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 10 | begin | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 11 | |
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 12 | definition xtor :: "'a \<Rightarrow> 'a" where | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 13 | "xtor x = x" | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 14 | |
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 15 | lemma xtor_map: "f (xtor x) = xtor (f x)" | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 16 | unfolding xtor_def by (rule refl) | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 17 | |
| 62863 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 traytel parents: 
62335diff
changeset | 18 | lemma xtor_map_unique: "u \<circ> xtor = xtor \<circ> f \<Longrightarrow> u = f" | 
| 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 traytel parents: 
62335diff
changeset | 19 | unfolding o_def xtor_def . | 
| 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 traytel parents: 
62335diff
changeset | 20 | |
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 21 | lemma xtor_set: "f (xtor x) = f x" | 
| 58353 | 22 | unfolding xtor_def by (rule refl) | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 23 | |
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 24 | lemma xtor_rel: "R (xtor x) (xtor y) = R x y" | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 25 | unfolding xtor_def by (rule refl) | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 26 | |
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 27 | lemma xtor_induct: "(\<And>x. P (xtor x)) \<Longrightarrow> P z" | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 28 | unfolding xtor_def by assumption | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 29 | |
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 30 | lemma xtor_xtor: "xtor (xtor x) = x" | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 31 | unfolding xtor_def by (rule refl) | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 32 | |
| 67399 | 33 | lemmas xtor_inject = xtor_rel[of "(=)"] | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 34 | |
| 58353 | 35 | lemma xtor_rel_induct: "(\<And>x y. vimage2p id_bnf id_bnf R x y \<Longrightarrow> IR (xtor x) (xtor y)) \<Longrightarrow> R \<le> IR" | 
| 61169 | 36 | unfolding xtor_def vimage2p_def id_bnf_def .. | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 37 | |
| 62321 
1abe81758619
keep 'ctor_iff_dtor' theorem around in BNF FP database
 blanchet parents: 
61169diff
changeset | 38 | lemma xtor_iff_xtor: "u = xtor w \<longleftrightarrow> xtor u = w" | 
| 
1abe81758619
keep 'ctor_iff_dtor' theorem around in BNF FP database
 blanchet parents: 
61169diff
changeset | 39 | unfolding xtor_def .. | 
| 
1abe81758619
keep 'ctor_iff_dtor' theorem around in BNF FP database
 blanchet parents: 
61169diff
changeset | 40 | |
| 58353 | 41 | lemma Inl_def_alt: "Inl \<equiv> (\<lambda>a. xtor (id_bnf (Inl a)))" | 
| 42 | unfolding xtor_def id_bnf_def by (rule reflexive) | |
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 43 | |
| 58353 | 44 | lemma Inr_def_alt: "Inr \<equiv> (\<lambda>a. xtor (id_bnf (Inr a)))" | 
| 45 | unfolding xtor_def id_bnf_def by (rule reflexive) | |
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 46 | |
| 58353 | 47 | lemma Pair_def_alt: "Pair \<equiv> (\<lambda>a b. xtor (id_bnf (a, b)))" | 
| 48 | unfolding xtor_def id_bnf_def by (rule reflexive) | |
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 49 | |
| 58377 
c6f93b8d2d8e
moved old 'size' generator together with 'old_datatype'
 blanchet parents: 
58353diff
changeset | 50 | definition ctor_rec :: "'a \<Rightarrow> 'a" where | 
| 
c6f93b8d2d8e
moved old 'size' generator together with 'old_datatype'
 blanchet parents: 
58353diff
changeset | 51 | "ctor_rec x = x" | 
| 
c6f93b8d2d8e
moved old 'size' generator together with 'old_datatype'
 blanchet parents: 
58353diff
changeset | 52 | |
| 
c6f93b8d2d8e
moved old 'size' generator together with 'old_datatype'
 blanchet parents: 
58353diff
changeset | 53 | lemma ctor_rec: "g = id \<Longrightarrow> ctor_rec f (xtor x) = f ((id_bnf \<circ> g \<circ> id_bnf) x)" | 
| 
c6f93b8d2d8e
moved old 'size' generator together with 'old_datatype'
 blanchet parents: 
58353diff
changeset | 54 | unfolding ctor_rec_def id_bnf_def xtor_def comp_def id_def by hypsubst (rule refl) | 
| 
c6f93b8d2d8e
moved old 'size' generator together with 'old_datatype'
 blanchet parents: 
58353diff
changeset | 55 | |
| 62863 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 traytel parents: 
62335diff
changeset | 56 | lemma ctor_rec_unique: "g = id \<Longrightarrow> f \<circ> xtor = s \<circ> (id_bnf \<circ> g \<circ> id_bnf) \<Longrightarrow> f = ctor_rec s" | 
| 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 traytel parents: 
62335diff
changeset | 57 | unfolding ctor_rec_def id_bnf_def xtor_def comp_def id_def by hypsubst (rule refl) | 
| 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 traytel parents: 
62335diff
changeset | 58 | |
| 58377 
c6f93b8d2d8e
moved old 'size' generator together with 'old_datatype'
 blanchet parents: 
58353diff
changeset | 59 | lemma ctor_rec_def_alt: "f = ctor_rec (f \<circ> id_bnf)" | 
| 
c6f93b8d2d8e
moved old 'size' generator together with 'old_datatype'
 blanchet parents: 
58353diff
changeset | 60 | unfolding ctor_rec_def id_bnf_def comp_def by (rule refl) | 
| 
c6f93b8d2d8e
moved old 'size' generator together with 'old_datatype'
 blanchet parents: 
58353diff
changeset | 61 | |
| 
c6f93b8d2d8e
moved old 'size' generator together with 'old_datatype'
 blanchet parents: 
58353diff
changeset | 62 | lemma ctor_rec_o_map: "ctor_rec f \<circ> g = ctor_rec (f \<circ> (id_bnf \<circ> g \<circ> id_bnf))" | 
| 
c6f93b8d2d8e
moved old 'size' generator together with 'old_datatype'
 blanchet parents: 
58353diff
changeset | 63 | unfolding ctor_rec_def id_bnf_def comp_def by (rule refl) | 
| 
c6f93b8d2d8e
moved old 'size' generator together with 'old_datatype'
 blanchet parents: 
58353diff
changeset | 64 | |
| 63045 | 65 | lemma ctor_rec_transfer: "rel_fun (rel_fun (vimage2p id_bnf id_bnf R) S) (rel_fun R S) ctor_rec ctor_rec" | 
| 66 | unfolding rel_fun_def vimage2p_def id_bnf_def ctor_rec_def by simp | |
| 67 | ||
| 58916 | 68 | lemma eq_fst_iff: "a = fst p \<longleftrightarrow> (\<exists>b. p = (a, b))" | 
| 69 | by (cases p) auto | |
| 70 | ||
| 71 | lemma eq_snd_iff: "b = snd p \<longleftrightarrow> (\<exists>a. p = (a, b))" | |
| 72 | by (cases p) auto | |
| 73 | ||
| 74 | lemma ex_neg_all_pos: "((\<exists>x. P x) \<Longrightarrow> Q) \<equiv> (\<And>x. P x \<Longrightarrow> Q)" | |
| 61169 | 75 | by standard blast+ | 
| 58916 | 76 | |
| 77 | lemma hypsubst_in_prems: "(\<And>x. y = x \<Longrightarrow> z = f x \<Longrightarrow> P) \<equiv> (z = f y \<Longrightarrow> P)" | |
| 61169 | 78 | by standard blast+ | 
| 58916 | 79 | |
| 80 | lemma isl_map_sum: | |
| 81 | "isl (map_sum f g s) = isl s" | |
| 82 | by (cases s) simp_all | |
| 83 | ||
| 84 | lemma map_sum_sel: | |
| 85 | "isl s \<Longrightarrow> projl (map_sum f g s) = f (projl s)" | |
| 86 | "\<not> isl s \<Longrightarrow> projr (map_sum f g s) = g (projr s)" | |
| 69850 | 87 | by (cases s; simp)+ | 
| 58916 | 88 | |
| 89 | lemma set_sum_sel: | |
| 90 | "isl s \<Longrightarrow> projl s \<in> setl s" | |
| 91 | "\<not> isl s \<Longrightarrow> projr s \<in> setr s" | |
| 69850 | 92 | by (cases s; auto intro: setl.intros setr.intros)+ | 
| 58916 | 93 | |
| 94 | lemma rel_sum_sel: "rel_sum R1 R2 a b = (isl a = isl b \<and> | |
| 95 | (isl a \<longrightarrow> isl b \<longrightarrow> R1 (projl a) (projl b)) \<and> | |
| 96 | (\<not> isl a \<longrightarrow> \<not> isl b \<longrightarrow> R2 (projr a) (projr b)))" | |
| 97 | by (cases a b rule: sum.exhaust[case_product sum.exhaust]) simp_all | |
| 98 | ||
| 67399 | 99 | lemma isl_transfer: "rel_fun (rel_sum A B) (=) isl isl" | 
| 58916 | 100 | unfolding rel_fun_def rel_sum_sel by simp | 
| 101 | ||
| 102 | lemma rel_prod_sel: "rel_prod R1 R2 p q = (R1 (fst p) (fst q) \<and> R2 (snd p) (snd q))" | |
| 103 | by (force simp: rel_prod.simps elim: rel_prod.cases) | |
| 104 | ||
| 69605 | 105 | ML_file \<open>Tools/BNF/bnf_lfp_basic_sugar.ML\<close> | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 106 | |
| 67332 | 107 | declare prod.size [no_atp] | 
| 58391 | 108 | |
| 58353 | 109 | hide_const (open) xtor ctor_rec | 
| 110 | ||
| 111 | hide_fact (open) | |
| 112 | xtor_def xtor_map xtor_set xtor_rel xtor_induct xtor_xtor xtor_inject ctor_rec_def ctor_rec | |
| 58377 
c6f93b8d2d8e
moved old 'size' generator together with 'old_datatype'
 blanchet parents: 
58353diff
changeset | 113 | ctor_rec_def_alt ctor_rec_o_map xtor_rel_induct Inl_def_alt Inr_def_alt Pair_def_alt | 
| 58353 | 114 | |
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 115 | end |