| author | haftmann | 
| Tue, 09 Nov 2010 14:02:13 +0100 | |
| changeset 40465 | 2989f9f3aa10 | 
| parent 40321 | d065b195ec89 | 
| child 40503 | 4094d788b904 | 
| permissions | -rw-r--r-- | 
| 15741 | 1 | (* Title: HOLCF/Domain.thy | 
| 2 | Author: Brian Huffman | |
| 3 | *) | |
| 4 | ||
| 5 | header {* Domain package *}
 | |
| 6 | ||
| 7 | theory Domain | |
| 33800 | 8 | imports Ssum Sprod Up One Tr Fixrec Representable | 
| 30910 | 9 | uses | 
| 10 |   ("Tools/cont_consts.ML")
 | |
| 11 |   ("Tools/cont_proc.ML")
 | |
| 35444 | 12 |   ("Tools/Domain/domain_constructors.ML")
 | 
| 32126 | 13 |   ("Tools/Domain/domain_axioms.ML")
 | 
| 40040 
3adb92ee2f22
rename domain_theorems.ML to domain_induction.ML; rename domain_extender.ML to domain.ML
 huffman parents: 
40039diff
changeset | 14 |   ("Tools/Domain/domain_induction.ML")
 | 
| 
3adb92ee2f22
rename domain_theorems.ML to domain_induction.ML; rename domain_extender.ML to domain.ML
 huffman parents: 
40039diff
changeset | 15 |   ("Tools/Domain/domain.ML")
 | 
| 15741 | 16 | begin | 
| 17 | ||
| 36452 | 18 | default_sort pcpo | 
| 15741 | 19 | |
| 23376 | 20 | |
| 15741 | 21 | subsection {* Casedist *}
 | 
| 22 | ||
| 40039 
391746914dba
simplify some proofs; remove some unused lists of lemmas
 huffman parents: 
40026diff
changeset | 23 | text {* Lemmas for proving nchotomy rule: *}
 | 
| 
391746914dba
simplify some proofs; remove some unused lists of lemmas
 huffman parents: 
40026diff
changeset | 24 | |
| 40321 
d065b195ec89
rename lemmas *_defined_iff and *_strict_iff to *_bottom_iff
 huffman parents: 
40040diff
changeset | 25 | lemma ex_one_bottom_iff: | 
| 15741 | 26 | "(\<exists>x. P x \<and> x \<noteq> \<bottom>) = P ONE" | 
| 40039 
391746914dba
simplify some proofs; remove some unused lists of lemmas
 huffman parents: 
40026diff
changeset | 27 | by simp | 
| 15741 | 28 | |
| 40321 
d065b195ec89
rename lemmas *_defined_iff and *_strict_iff to *_bottom_iff
 huffman parents: 
40040diff
changeset | 29 | lemma ex_up_bottom_iff: | 
| 15741 | 30 | "(\<exists>x. P x \<and> x \<noteq> \<bottom>) = (\<exists>x. P (up\<cdot>x))" | 
| 40039 
391746914dba
simplify some proofs; remove some unused lists of lemmas
 huffman parents: 
40026diff
changeset | 31 | by (safe, case_tac x, auto) | 
| 15741 | 32 | |
| 40321 
d065b195ec89
rename lemmas *_defined_iff and *_strict_iff to *_bottom_iff
 huffman parents: 
40040diff
changeset | 33 | lemma ex_sprod_bottom_iff: | 
| 15741 | 34 | "(\<exists>y. P y \<and> y \<noteq> \<bottom>) = | 
| 35 | (\<exists>x y. (P (:x, y:) \<and> x \<noteq> \<bottom>) \<and> y \<noteq> \<bottom>)" | |
| 40039 
391746914dba
simplify some proofs; remove some unused lists of lemmas
 huffman parents: 
40026diff
changeset | 36 | by (safe, case_tac y, auto) | 
| 15741 | 37 | |
| 40321 
d065b195ec89
rename lemmas *_defined_iff and *_strict_iff to *_bottom_iff
 huffman parents: 
40040diff
changeset | 38 | lemma ex_sprod_up_bottom_iff: | 
| 15741 | 39 | "(\<exists>y. P y \<and> y \<noteq> \<bottom>) = | 
| 40 | (\<exists>x y. P (:up\<cdot>x, y:) \<and> y \<noteq> \<bottom>)" | |
| 40039 
391746914dba
simplify some proofs; remove some unused lists of lemmas
 huffman parents: 
40026diff
changeset | 41 | by (safe, case_tac y, simp, case_tac x, auto) | 
| 15741 | 42 | |
| 40321 
d065b195ec89
rename lemmas *_defined_iff and *_strict_iff to *_bottom_iff
 huffman parents: 
40040diff
changeset | 43 | lemma ex_ssum_bottom_iff: | 
| 15741 | 44 | "(\<exists>x. P x \<and> x \<noteq> \<bottom>) = | 
| 45 | ((\<exists>x. P (sinl\<cdot>x) \<and> x \<noteq> \<bottom>) \<or> | |
| 46 | (\<exists>x. P (sinr\<cdot>x) \<and> x \<noteq> \<bottom>))" | |
| 40039 
391746914dba
simplify some proofs; remove some unused lists of lemmas
 huffman parents: 
40026diff
changeset | 47 | by (safe, case_tac x, auto) | 
| 15741 | 48 | |
| 49 | lemma exh_start: "p = \<bottom> \<or> (\<exists>x. p = x \<and> x \<noteq> \<bottom>)" | |
| 23376 | 50 | by auto | 
| 15741 | 51 | |
| 40321 
d065b195ec89
rename lemmas *_defined_iff and *_strict_iff to *_bottom_iff
 huffman parents: 
40040diff
changeset | 52 | lemmas ex_bottom_iffs = | 
| 
d065b195ec89
rename lemmas *_defined_iff and *_strict_iff to *_bottom_iff
 huffman parents: 
40040diff
changeset | 53 | ex_ssum_bottom_iff | 
| 
d065b195ec89
rename lemmas *_defined_iff and *_strict_iff to *_bottom_iff
 huffman parents: 
40040diff
changeset | 54 | ex_sprod_up_bottom_iff | 
| 
d065b195ec89
rename lemmas *_defined_iff and *_strict_iff to *_bottom_iff
 huffman parents: 
40040diff
changeset | 55 | ex_sprod_bottom_iff | 
| 
d065b195ec89
rename lemmas *_defined_iff and *_strict_iff to *_bottom_iff
 huffman parents: 
40040diff
changeset | 56 | ex_up_bottom_iff | 
| 
d065b195ec89
rename lemmas *_defined_iff and *_strict_iff to *_bottom_iff
 huffman parents: 
40040diff
changeset | 57 | ex_one_bottom_iff | 
| 15741 | 58 | |
| 40039 
391746914dba
simplify some proofs; remove some unused lists of lemmas
 huffman parents: 
40026diff
changeset | 59 | text {* Rules for turning nchotomy into exhaust: *}
 | 
| 15741 | 60 | |
| 61 | lemma exh_casedist0: "\<lbrakk>R; R \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P" (* like make_elim *) | |
| 23376 | 62 | by auto | 
| 15741 | 63 | |
| 64 | lemma exh_casedist1: "((P \<or> Q \<Longrightarrow> R) \<Longrightarrow> S) \<equiv> (\<lbrakk>P \<Longrightarrow> R; Q \<Longrightarrow> R\<rbrakk> \<Longrightarrow> S)" | |
| 23376 | 65 | by rule auto | 
| 15741 | 66 | |
| 67 | lemma exh_casedist2: "(\<exists>x. P x \<Longrightarrow> Q) \<equiv> (\<And>x. P x \<Longrightarrow> Q)" | |
| 23376 | 68 | by rule auto | 
| 15741 | 69 | |
| 70 | lemma exh_casedist3: "(P \<and> Q \<Longrightarrow> R) \<equiv> (P \<Longrightarrow> Q \<Longrightarrow> R)" | |
| 23376 | 71 | by rule auto | 
| 15741 | 72 | |
| 73 | lemmas exh_casedists = exh_casedist1 exh_casedist2 exh_casedist3 | |
| 74 | ||
| 30910 | 75 | |
| 76 | subsection {* Installing the domain package *}
 | |
| 77 | ||
| 30911 
7809cbaa1b61
domain package: simplify internal proofs of con_rews
 huffman parents: 
30910diff
changeset | 78 | lemmas con_strict_rules = | 
| 
7809cbaa1b61
domain package: simplify internal proofs of con_rews
 huffman parents: 
30910diff
changeset | 79 | sinl_strict sinr_strict spair_strict1 spair_strict2 | 
| 
7809cbaa1b61
domain package: simplify internal proofs of con_rews
 huffman parents: 
30910diff
changeset | 80 | |
| 40321 
d065b195ec89
rename lemmas *_defined_iff and *_strict_iff to *_bottom_iff
 huffman parents: 
40040diff
changeset | 81 | lemmas con_bottom_iff_rules = | 
| 
d065b195ec89
rename lemmas *_defined_iff and *_strict_iff to *_bottom_iff
 huffman parents: 
40040diff
changeset | 82 | sinl_bottom_iff sinr_bottom_iff spair_bottom_iff up_defined ONE_defined | 
| 30911 
7809cbaa1b61
domain package: simplify internal proofs of con_rews
 huffman parents: 
30910diff
changeset | 83 | |
| 35117 
eeec2a320a77
change generated lemmas dist_eqs and dist_les to iff-style
 huffman parents: 
33800diff
changeset | 84 | lemmas con_below_iff_rules = | 
| 40321 
d065b195ec89
rename lemmas *_defined_iff and *_strict_iff to *_bottom_iff
 huffman parents: 
40040diff
changeset | 85 | sinl_below sinr_below sinl_below_sinr sinr_below_sinl con_bottom_iff_rules | 
| 35117 
eeec2a320a77
change generated lemmas dist_eqs and dist_les to iff-style
 huffman parents: 
33800diff
changeset | 86 | |
| 
eeec2a320a77
change generated lemmas dist_eqs and dist_les to iff-style
 huffman parents: 
33800diff
changeset | 87 | lemmas con_eq_iff_rules = | 
| 40321 
d065b195ec89
rename lemmas *_defined_iff and *_strict_iff to *_bottom_iff
 huffman parents: 
40040diff
changeset | 88 | sinl_eq sinr_eq sinl_eq_sinr sinr_eq_sinl con_bottom_iff_rules | 
| 35117 
eeec2a320a77
change generated lemmas dist_eqs and dist_les to iff-style
 huffman parents: 
33800diff
changeset | 89 | |
| 35446 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35444diff
changeset | 90 | lemmas sel_strict_rules = | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35444diff
changeset | 91 | cfcomp2 sscase1 sfst_strict ssnd_strict fup1 | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35444diff
changeset | 92 | |
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35444diff
changeset | 93 | lemma sel_app_extra_rules: | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35444diff
changeset | 94 | "sscase\<cdot>ID\<cdot>\<bottom>\<cdot>(sinr\<cdot>x) = \<bottom>" | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35444diff
changeset | 95 | "sscase\<cdot>ID\<cdot>\<bottom>\<cdot>(sinl\<cdot>x) = x" | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35444diff
changeset | 96 | "sscase\<cdot>\<bottom>\<cdot>ID\<cdot>(sinl\<cdot>x) = \<bottom>" | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35444diff
changeset | 97 | "sscase\<cdot>\<bottom>\<cdot>ID\<cdot>(sinr\<cdot>x) = x" | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35444diff
changeset | 98 | "fup\<cdot>ID\<cdot>(up\<cdot>x) = x" | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35444diff
changeset | 99 | by (cases "x = \<bottom>", simp, simp)+ | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35444diff
changeset | 100 | |
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35444diff
changeset | 101 | lemmas sel_app_rules = | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35444diff
changeset | 102 | sel_strict_rules sel_app_extra_rules | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35444diff
changeset | 103 | ssnd_spair sfst_spair up_defined spair_defined | 
| 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35444diff
changeset | 104 | |
| 40321 
d065b195ec89
rename lemmas *_defined_iff and *_strict_iff to *_bottom_iff
 huffman parents: 
40040diff
changeset | 105 | lemmas sel_bottom_iff_rules = | 
| 
d065b195ec89
rename lemmas *_defined_iff and *_strict_iff to *_bottom_iff
 huffman parents: 
40040diff
changeset | 106 | cfcomp2 sfst_bottom_iff ssnd_bottom_iff | 
| 35446 
b719dad322fa
rewrite domain package code for selector functions
 huffman parents: 
35444diff
changeset | 107 | |
| 35494 
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
 huffman parents: 
35457diff
changeset | 108 | lemmas take_con_rules = | 
| 36160 
f84fa49a0b69
add rule deflation_ID to proof script for take + constructor rules
 huffman parents: 
35653diff
changeset | 109 | ssum_map_sinl' ssum_map_sinr' sprod_map_spair' u_map_up | 
| 
f84fa49a0b69
add rule deflation_ID to proof script for take + constructor rules
 huffman parents: 
35653diff
changeset | 110 | deflation_strict deflation_ID ID1 cfcomp2 | 
| 35494 
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
 huffman parents: 
35457diff
changeset | 111 | |
| 30910 | 112 | use "Tools/cont_consts.ML" | 
| 113 | use "Tools/cont_proc.ML" | |
| 32126 | 114 | use "Tools/Domain/domain_axioms.ML" | 
| 35457 
d63655b88369
move proofs of casedist into domain_constructors.ML
 huffman parents: 
35446diff
changeset | 115 | use "Tools/Domain/domain_constructors.ML" | 
| 40040 
3adb92ee2f22
rename domain_theorems.ML to domain_induction.ML; rename domain_extender.ML to domain.ML
 huffman parents: 
40039diff
changeset | 116 | use "Tools/Domain/domain_induction.ML" | 
| 
3adb92ee2f22
rename domain_theorems.ML to domain_induction.ML; rename domain_extender.ML to domain.ML
 huffman parents: 
40039diff
changeset | 117 | use "Tools/Domain/domain.ML" | 
| 30910 | 118 | |
| 15741 | 119 | end |