src/HOL/Basic_BNF_LFPs.thy
author immler
Sat, 23 Feb 2019 19:50:21 +0000
changeset 69830 54d19f1f0ba6
parent 69605 a96320074298
child 69850 5f993636ac07
permissions -rw-r--r--
no more shadowing of Min and Max by Approximation
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
59141
9a5c2e9b001e renamed theory file
blanchet
parents: 58916
diff changeset
     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
9a5c2e9b001e renamed theory file
blanchet
parents: 58916
diff changeset
     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: 62335
diff 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: 62335
diff changeset
    19
  unfolding o_def xtor_def .
e0b894bba6ff single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
traytel
parents: 62335
diff 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
c9f374b64d99 tuned fact visibility
blanchet
parents: 58352
diff changeset
    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
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67332
diff changeset
    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
c9f374b64d99 tuned fact visibility
blanchet
parents: 58352
diff changeset
    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
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61076
diff changeset
    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: 61169
diff 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: 61169
diff changeset
    39
  unfolding xtor_def ..
1abe81758619 keep 'ctor_iff_dtor' theorem around in BNF FP database
blanchet
parents: 61169
diff changeset
    40
58353
c9f374b64d99 tuned fact visibility
blanchet
parents: 58352
diff changeset
    41
lemma Inl_def_alt: "Inl \<equiv> (\<lambda>a. xtor (id_bnf (Inl a)))"
c9f374b64d99 tuned fact visibility
blanchet
parents: 58352
diff changeset
    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
c9f374b64d99 tuned fact visibility
blanchet
parents: 58352
diff changeset
    44
lemma Inr_def_alt: "Inr \<equiv> (\<lambda>a. xtor (id_bnf (Inr a)))"
c9f374b64d99 tuned fact visibility
blanchet
parents: 58352
diff changeset
    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
c9f374b64d99 tuned fact visibility
blanchet
parents: 58352
diff changeset
    47
lemma Pair_def_alt: "Pair \<equiv> (\<lambda>a b. xtor (id_bnf (a, b)))"
c9f374b64d99 tuned fact visibility
blanchet
parents: 58352
diff changeset
    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: 58353
diff changeset
    50
definition ctor_rec :: "'a \<Rightarrow> 'a" where
c6f93b8d2d8e moved old 'size' generator together with 'old_datatype'
blanchet
parents: 58353
diff changeset
    51
  "ctor_rec x = x"
c6f93b8d2d8e moved old 'size' generator together with 'old_datatype'
blanchet
parents: 58353
diff changeset
    52
c6f93b8d2d8e moved old 'size' generator together with 'old_datatype'
blanchet
parents: 58353
diff 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: 58353
diff 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: 58353
diff changeset
    55
62863
e0b894bba6ff single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
traytel
parents: 62335
diff 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: 62335
diff 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: 62335
diff changeset
    58
58377
c6f93b8d2d8e moved old 'size' generator together with 'old_datatype'
blanchet
parents: 58353
diff 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: 58353
diff 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: 58353
diff changeset
    61
c6f93b8d2d8e moved old 'size' generator together with 'old_datatype'
blanchet
parents: 58353
diff 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: 58353
diff 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: 58353
diff changeset
    64
63045
c50c764aab10 n2m operates on (un)folds
traytel
parents: 62863
diff changeset
    65
lemma ctor_rec_transfer: "rel_fun (rel_fun (vimage2p id_bnf id_bnf R) S) (rel_fun R S) ctor_rec ctor_rec"
c50c764aab10 n2m operates on (un)folds
traytel
parents: 62863
diff changeset
    66
  unfolding rel_fun_def vimage2p_def id_bnf_def ctor_rec_def by simp
c50c764aab10 n2m operates on (un)folds
traytel
parents: 62863
diff changeset
    67
58916
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
    68
lemma eq_fst_iff: "a = fst p \<longleftrightarrow> (\<exists>b. p = (a, b))"
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
    69
  by (cases p) auto
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
    70
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
    71
lemma eq_snd_iff: "b = snd p \<longleftrightarrow> (\<exists>a. p = (a, b))"
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
    72
  by (cases p) auto
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
    73
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
    74
lemma ex_neg_all_pos: "((\<exists>x. P x) \<Longrightarrow> Q) \<equiv> (\<And>x. P x \<Longrightarrow> Q)"
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61076
diff changeset
    75
  by standard blast+
58916
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
    76
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
    77
lemma hypsubst_in_prems: "(\<And>x. y = x \<Longrightarrow> z = f x \<Longrightarrow> P) \<equiv> (z = f y \<Longrightarrow> P)"
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61076
diff changeset
    78
  by standard blast+
58916
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
    79
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
    80
lemma isl_map_sum:
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
    81
  "isl (map_sum f g s) = isl s"
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
    82
  by (cases s) simp_all
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
    83
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
    84
lemma map_sum_sel:
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
    85
  "isl s \<Longrightarrow> projl (map_sum f g s) = f (projl s)"
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
    86
  "\<not> isl s \<Longrightarrow> projr (map_sum f g s) = g (projr s)"
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
    87
  by (case_tac [!] s) simp_all
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
    88
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
    89
lemma set_sum_sel:
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
    90
  "isl s \<Longrightarrow> projl s \<in> setl s"
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
    91
  "\<not> isl s \<Longrightarrow> projr s \<in> setr s"
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
    92
  by (case_tac [!] s) (auto intro: setl.intros setr.intros)
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
    93
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
    94
lemma rel_sum_sel: "rel_sum R1 R2 a b = (isl a = isl b \<and>
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
    95
  (isl a \<longrightarrow> isl b \<longrightarrow> R1 (projl a) (projl b)) \<and>
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
    96
  (\<not> isl a \<longrightarrow> \<not> isl b \<longrightarrow> R2 (projr a) (projr b)))"
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
    97
  by (cases a b rule: sum.exhaust[case_product sum.exhaust]) simp_all
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
    98
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67332
diff changeset
    99
lemma isl_transfer: "rel_fun (rel_sum A B) (=) isl isl"
58916
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
   100
  unfolding rel_fun_def rel_sum_sel by simp
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
   101
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
   102
lemma rel_prod_sel: "rel_prod R1 R2 p q = (R1 (fst p) (fst q) \<and> R2 (snd p) (snd q))"
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
   103
  by (force simp: rel_prod.simps elim: rel_prod.cases)
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58391
diff changeset
   104
69605
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 67399
diff changeset
   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
cb96edae56ef kill old size infrastructure
blanchet
parents: 63045
diff changeset
   107
declare prod.size [no_atp]
58391
fe0fc8aee49a reintroduced old setup for size of basic types
blanchet
parents: 58390
diff changeset
   108
58353
c9f374b64d99 tuned fact visibility
blanchet
parents: 58352
diff changeset
   109
hide_const (open) xtor ctor_rec
c9f374b64d99 tuned fact visibility
blanchet
parents: 58352
diff changeset
   110
c9f374b64d99 tuned fact visibility
blanchet
parents: 58352
diff changeset
   111
hide_fact (open)
c9f374b64d99 tuned fact visibility
blanchet
parents: 58352
diff changeset
   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: 58353
diff changeset
   113
  ctor_rec_def_alt ctor_rec_o_map xtor_rel_induct Inl_def_alt Inr_def_alt Pair_def_alt
58353
c9f374b64d99 tuned fact visibility
blanchet
parents: 58352
diff changeset
   114
58352
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   115
end