| author | nipkow | 
| Tue, 29 Oct 2024 10:26:06 +0100 | |
| changeset 81285 | 34f3ec8d4d24 | 
| parent 80914 | d97fdabd9e2b | 
| child 81574 | c4abe6582ee5 | 
| permissions | -rw-r--r-- | 
| 42151 | 1 | (* Title: HOL/HOLCF/Fixrec.thy | 
| 16221 | 2 | Author: Amber Telfer and Brian Huffman | 
| 3 | *) | |
| 4 | ||
| 58880 | 5 | section "Package for defining recursive functions in HOLCF" | 
| 16221 | 6 | |
| 7 | theory Fixrec | |
| 65380 
ae93953746fc
eliminated Plain_HOLCF.thy (see also 8e92772bc0e8): it was modeled after HOL/Plain.thy which was discontinued later;
 wenzelm parents: 
62175diff
changeset | 8 | imports Cprod Sprod Ssum Up One Tr Fix | 
| 69913 | 9 | keywords "fixrec" :: thy_defn | 
| 16221 | 10 | begin | 
| 11 | ||
| 62175 | 12 | subsection \<open>Pattern-match monad\<close> | 
| 16221 | 13 | |
| 36452 | 14 | default_sort cpo | 
| 16776 
a3899ac14a1c
generalized types of monadic operators to class cpo; added match function for UU
 huffman parents: 
16758diff
changeset | 15 | |
| 49759 
ecf1d5d87e5e
removed support for set constant definitions in HOLCF {cpo,pcpo,domain}def commands;
 huffman parents: 
48891diff
changeset | 16 | pcpodef 'a match = "UNIV::(one ++ 'a u) set" | 
| 29063 
7619f0561cd7
pcpodef package: state two goals, instead of encoded conjunction;
 wenzelm parents: 
28891diff
changeset | 17 | by simp_all | 
| 16221 | 18 | |
| 29141 | 19 | definition | 
| 37108 
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
 huffman parents: 
37080diff
changeset | 20 | fail :: "'a match" where | 
| 
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
 huffman parents: 
37080diff
changeset | 21 | "fail = Abs_match (sinl\<cdot>ONE)" | 
| 16221 | 22 | |
| 29141 | 23 | definition | 
| 37108 
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
 huffman parents: 
37080diff
changeset | 24 | succeed :: "'a \<rightarrow> 'a match" where | 
| 
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
 huffman parents: 
37080diff
changeset | 25 | "succeed = (\<Lambda> x. Abs_match (sinr\<cdot>(up\<cdot>x)))" | 
| 19092 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 26 | |
| 37108 
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
 huffman parents: 
37080diff
changeset | 27 | lemma matchE [case_names bottom fail succeed, cases type: match]: | 
| 
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
 huffman parents: 
37080diff
changeset | 28 | "\<lbrakk>p = \<bottom> \<Longrightarrow> Q; p = fail \<Longrightarrow> Q; \<And>x. p = succeed\<cdot>x \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q" | 
| 
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
 huffman parents: 
37080diff
changeset | 29 | unfolding fail_def succeed_def | 
| 19092 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 30 | apply (cases p, rename_tac r) | 
| 37108 
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
 huffman parents: 
37080diff
changeset | 31 | apply (rule_tac p=r in ssumE, simp add: Abs_match_strict) | 
| 16221 | 32 | apply (rule_tac p=x in oneE, simp, simp) | 
| 37108 
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
 huffman parents: 
37080diff
changeset | 33 | apply (rule_tac p=y in upE, simp, simp add: cont_Abs_match) | 
| 16221 | 34 | done | 
| 35 | ||
| 37108 
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
 huffman parents: 
37080diff
changeset | 36 | lemma succeed_defined [simp]: "succeed\<cdot>x \<noteq> \<bottom>" | 
| 41029 
f7d8cfa6e7fc
pcpodef no longer generates _defined lemmas, use _bottom_iff lemmas instead
 huffman parents: 
40834diff
changeset | 37 | by (simp add: succeed_def cont_Abs_match Abs_match_bottom_iff) | 
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 38 | |
| 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 39 | lemma fail_defined [simp]: "fail \<noteq> \<bottom>" | 
| 41029 
f7d8cfa6e7fc
pcpodef no longer generates _defined lemmas, use _bottom_iff lemmas instead
 huffman parents: 
40834diff
changeset | 40 | by (simp add: fail_def Abs_match_bottom_iff) | 
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 41 | |
| 37108 
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
 huffman parents: 
37080diff
changeset | 42 | lemma succeed_eq [simp]: "(succeed\<cdot>x = succeed\<cdot>y) = (x = y)" | 
| 
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
 huffman parents: 
37080diff
changeset | 43 | by (simp add: succeed_def cont_Abs_match Abs_match_inject) | 
| 18293 
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
 huffman parents: 
18112diff
changeset | 44 | |
| 37108 
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
 huffman parents: 
37080diff
changeset | 45 | lemma succeed_neq_fail [simp]: | 
| 
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
 huffman parents: 
37080diff
changeset | 46 | "succeed\<cdot>x \<noteq> fail" "fail \<noteq> succeed\<cdot>x" | 
| 
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
 huffman parents: 
37080diff
changeset | 47 | by (simp_all add: succeed_def fail_def cont_Abs_match Abs_match_inject) | 
| 19092 
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
 huffman parents: 
18293diff
changeset | 48 | |
| 62175 | 49 | subsubsection \<open>Run operator\<close> | 
| 16221 | 50 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 51 | definition | 
| 37108 
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
 huffman parents: 
37080diff
changeset | 52 | run :: "'a match \<rightarrow> 'a::pcpo" where | 
| 40735 | 53 | "run = (\<Lambda> m. sscase\<cdot>\<bottom>\<cdot>(fup\<cdot>ID)\<cdot>(Rep_match m))" | 
| 16221 | 54 | |
| 62175 | 55 | text \<open>rewrite rules for run\<close> | 
| 16221 | 56 | |
| 57 | lemma run_strict [simp]: "run\<cdot>\<bottom> = \<bottom>" | |
| 40735 | 58 | unfolding run_def | 
| 59 | by (simp add: cont_Rep_match Rep_match_strict) | |
| 16221 | 60 | |
| 61 | lemma run_fail [simp]: "run\<cdot>fail = \<bottom>" | |
| 40735 | 62 | unfolding run_def fail_def | 
| 63 | by (simp add: cont_Rep_match Abs_match_inverse) | |
| 16221 | 64 | |
| 37108 
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
 huffman parents: 
37080diff
changeset | 65 | lemma run_succeed [simp]: "run\<cdot>(succeed\<cdot>x) = x" | 
| 40735 | 66 | unfolding run_def succeed_def | 
| 67 | by (simp add: cont_Rep_match cont_Abs_match Abs_match_inverse) | |
| 16221 | 68 | |
| 62175 | 69 | subsubsection \<open>Monad plus operator\<close> | 
| 16221 | 70 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 71 | definition | 
| 37108 
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
 huffman parents: 
37080diff
changeset | 72 | mplus :: "'a match \<rightarrow> 'a match \<rightarrow> 'a match" where | 
| 40735 | 73 | "mplus = (\<Lambda> m1 m2. sscase\<cdot>(\<Lambda> _. m2)\<cdot>(\<Lambda> _. m1)\<cdot>(Rep_match m1))" | 
| 18097 | 74 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 75 | abbreviation | 
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
69913diff
changeset | 76 | mplus_syn :: "['a match, 'a match] \<Rightarrow> 'a match" (infixr \<open>+++\<close> 65) where | 
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 77 | "m1 +++ m2 == mplus\<cdot>m1\<cdot>m2" | 
| 16221 | 78 | |
| 62175 | 79 | text \<open>rewrite rules for mplus\<close> | 
| 16221 | 80 | |
| 81 | lemma mplus_strict [simp]: "\<bottom> +++ m = \<bottom>" | |
| 40735 | 82 | unfolding mplus_def | 
| 40834 
a1249aeff5b6
change cpodef-generated cont_Rep rules to cont2cont format
 huffman parents: 
40795diff
changeset | 83 | by (simp add: cont_Rep_match Rep_match_strict) | 
| 16221 | 84 | |
| 85 | lemma mplus_fail [simp]: "fail +++ m = m" | |
| 40735 | 86 | unfolding mplus_def fail_def | 
| 40834 
a1249aeff5b6
change cpodef-generated cont_Rep rules to cont2cont format
 huffman parents: 
40795diff
changeset | 87 | by (simp add: cont_Rep_match Abs_match_inverse) | 
| 16221 | 88 | |
| 37108 
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
 huffman parents: 
37080diff
changeset | 89 | lemma mplus_succeed [simp]: "succeed\<cdot>x +++ m = succeed\<cdot>x" | 
| 40735 | 90 | unfolding mplus_def succeed_def | 
| 40834 
a1249aeff5b6
change cpodef-generated cont_Rep rules to cont2cont format
 huffman parents: 
40795diff
changeset | 91 | by (simp add: cont_Rep_match cont_Abs_match Abs_match_inverse) | 
| 16221 | 92 | |
| 16460 
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
 huffman parents: 
16417diff
changeset | 93 | lemma mplus_fail2 [simp]: "m +++ fail = m" | 
| 37108 
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
 huffman parents: 
37080diff
changeset | 94 | by (cases m, simp_all) | 
| 16460 
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
 huffman parents: 
16417diff
changeset | 95 | |
| 16221 | 96 | lemma mplus_assoc: "(x +++ y) +++ z = x +++ (y +++ z)" | 
| 37108 
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
 huffman parents: 
37080diff
changeset | 97 | by (cases x, simp_all) | 
| 16221 | 98 | |
| 62175 | 99 | subsection \<open>Match functions for built-in types\<close> | 
| 16221 | 100 | |
| 36452 | 101 | default_sort pcpo | 
| 16776 
a3899ac14a1c
generalized types of monadic operators to class cpo; added match function for UU
 huffman parents: 
16758diff
changeset | 102 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 103 | definition | 
| 40768 | 104 | match_bottom :: "'a \<rightarrow> 'c match \<rightarrow> 'c match" | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 105 | where | 
| 40768 | 106 | "match_bottom = (\<Lambda> x k. seq\<cdot>x\<cdot>fail)" | 
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 107 | |
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 108 | definition | 
| 39807 | 109 |   match_Pair :: "'a::cpo \<times> 'b::cpo \<rightarrow> ('a \<rightarrow> 'b \<rightarrow> 'c match) \<rightarrow> 'c match"
 | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 110 | where | 
| 39807 | 111 | "match_Pair = (\<Lambda> x k. csplit\<cdot>k\<cdot>x)" | 
| 16776 
a3899ac14a1c
generalized types of monadic operators to class cpo; added match function for UU
 huffman parents: 
16758diff
changeset | 112 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 113 | definition | 
| 37108 
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
 huffman parents: 
37080diff
changeset | 114 |   match_spair :: "'a \<otimes> 'b \<rightarrow> ('a \<rightarrow> 'b \<rightarrow> 'c match) \<rightarrow> 'c match"
 | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 115 | where | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 116 | "match_spair = (\<Lambda> x k. ssplit\<cdot>k\<cdot>x)" | 
| 16221 | 117 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 118 | definition | 
| 37108 
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
 huffman parents: 
37080diff
changeset | 119 |   match_sinl :: "'a \<oplus> 'b \<rightarrow> ('a \<rightarrow> 'c match) \<rightarrow> 'c match"
 | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 120 | where | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 121 | "match_sinl = (\<Lambda> x k. sscase\<cdot>k\<cdot>(\<Lambda> b. fail)\<cdot>x)" | 
| 16551 | 122 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 123 | definition | 
| 37108 
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
 huffman parents: 
37080diff
changeset | 124 |   match_sinr :: "'a \<oplus> 'b \<rightarrow> ('b \<rightarrow> 'c match) \<rightarrow> 'c match"
 | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 125 | where | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 126 | "match_sinr = (\<Lambda> x k. sscase\<cdot>(\<Lambda> a. fail)\<cdot>k\<cdot>x)" | 
| 16551 | 127 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 128 | definition | 
| 37108 
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
 huffman parents: 
37080diff
changeset | 129 |   match_up :: "'a::cpo u \<rightarrow> ('a \<rightarrow> 'c match) \<rightarrow> 'c match"
 | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 130 | where | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 131 | "match_up = (\<Lambda> x k. fup\<cdot>k\<cdot>x)" | 
| 16221 | 132 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 133 | definition | 
| 37108 
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
 huffman parents: 
37080diff
changeset | 134 | match_ONE :: "one \<rightarrow> 'c match \<rightarrow> 'c match" | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 135 | where | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 136 | "match_ONE = (\<Lambda> ONE k. k)" | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 137 | |
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 138 | definition | 
| 37108 
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
 huffman parents: 
37080diff
changeset | 139 | match_TT :: "tr \<rightarrow> 'c match \<rightarrow> 'c match" | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 140 | where | 
| 40322 
707eb30e8a53
make syntax of continuous if-then-else consistent with HOL if-then-else
 huffman parents: 
40046diff
changeset | 141 | "match_TT = (\<Lambda> x k. If x then k else fail)" | 
| 65380 
ae93953746fc
eliminated Plain_HOLCF.thy (see also 8e92772bc0e8): it was modeled after HOL/Plain.thy which was discontinued later;
 wenzelm parents: 
62175diff
changeset | 142 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
23152diff
changeset | 143 | definition | 
| 37108 
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
 huffman parents: 
37080diff
changeset | 144 | match_FF :: "tr \<rightarrow> 'c match \<rightarrow> 'c match" | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 145 | where | 
| 40322 
707eb30e8a53
make syntax of continuous if-then-else consistent with HOL if-then-else
 huffman parents: 
40046diff
changeset | 146 | "match_FF = (\<Lambda> x k. If x then fail else k)" | 
| 16460 
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
 huffman parents: 
16417diff
changeset | 147 | |
| 40768 | 148 | lemma match_bottom_simps [simp]: | 
| 40795 
c52cd8bc426d
change match_bottom_simps to produce if-then-else, making more uses of bottom-patterns work with fixrec
 huffman parents: 
40774diff
changeset | 149 | "match_bottom\<cdot>x\<cdot>k = (if x = \<bottom> then \<bottom> else fail)" | 
| 
c52cd8bc426d
change match_bottom_simps to produce if-then-else, making more uses of bottom-patterns work with fixrec
 huffman parents: 
40774diff
changeset | 150 | by (simp add: match_bottom_def) | 
| 16776 
a3899ac14a1c
generalized types of monadic operators to class cpo; added match function for UU
 huffman parents: 
16758diff
changeset | 151 | |
| 39807 | 152 | lemma match_Pair_simps [simp]: | 
| 153 | "match_Pair\<cdot>(x, y)\<cdot>k = k\<cdot>x\<cdot>y" | |
| 154 | by (simp_all add: match_Pair_def) | |
| 16221 | 155 | |
| 16551 | 156 | lemma match_spair_simps [simp]: | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 157 | "\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> match_spair\<cdot>(:x, y:)\<cdot>k = k\<cdot>x\<cdot>y" | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 158 | "match_spair\<cdot>\<bottom>\<cdot>k = \<bottom>" | 
| 16551 | 159 | by (simp_all add: match_spair_def) | 
| 160 | ||
| 161 | lemma match_sinl_simps [simp]: | |
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 162 | "x \<noteq> \<bottom> \<Longrightarrow> match_sinl\<cdot>(sinl\<cdot>x)\<cdot>k = k\<cdot>x" | 
| 30914 
ceeb5f2eac75
fix too-specific types in lemmas match_{sinl,sinr}_simps
 huffman parents: 
30912diff
changeset | 163 | "y \<noteq> \<bottom> \<Longrightarrow> match_sinl\<cdot>(sinr\<cdot>y)\<cdot>k = fail" | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 164 | "match_sinl\<cdot>\<bottom>\<cdot>k = \<bottom>" | 
| 16551 | 165 | by (simp_all add: match_sinl_def) | 
| 166 | ||
| 167 | lemma match_sinr_simps [simp]: | |
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 168 | "x \<noteq> \<bottom> \<Longrightarrow> match_sinr\<cdot>(sinl\<cdot>x)\<cdot>k = fail" | 
| 30914 
ceeb5f2eac75
fix too-specific types in lemmas match_{sinl,sinr}_simps
 huffman parents: 
30912diff
changeset | 169 | "y \<noteq> \<bottom> \<Longrightarrow> match_sinr\<cdot>(sinr\<cdot>y)\<cdot>k = k\<cdot>y" | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 170 | "match_sinr\<cdot>\<bottom>\<cdot>k = \<bottom>" | 
| 16551 | 171 | by (simp_all add: match_sinr_def) | 
| 172 | ||
| 16221 | 173 | lemma match_up_simps [simp]: | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 174 | "match_up\<cdot>(up\<cdot>x)\<cdot>k = k\<cdot>x" | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 175 | "match_up\<cdot>\<bottom>\<cdot>k = \<bottom>" | 
| 16221 | 176 | by (simp_all add: match_up_def) | 
| 177 | ||
| 16460 
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
 huffman parents: 
16417diff
changeset | 178 | lemma match_ONE_simps [simp]: | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 179 | "match_ONE\<cdot>ONE\<cdot>k = k" | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 180 | "match_ONE\<cdot>\<bottom>\<cdot>k = \<bottom>" | 
| 18094 | 181 | by (simp_all add: match_ONE_def) | 
| 16460 
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
 huffman parents: 
16417diff
changeset | 182 | |
| 
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
 huffman parents: 
16417diff
changeset | 183 | lemma match_TT_simps [simp]: | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 184 | "match_TT\<cdot>TT\<cdot>k = k" | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 185 | "match_TT\<cdot>FF\<cdot>k = fail" | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 186 | "match_TT\<cdot>\<bottom>\<cdot>k = \<bottom>" | 
| 18094 | 187 | by (simp_all add: match_TT_def) | 
| 16460 
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
 huffman parents: 
16417diff
changeset | 188 | |
| 
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
 huffman parents: 
16417diff
changeset | 189 | lemma match_FF_simps [simp]: | 
| 30912 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 190 | "match_FF\<cdot>FF\<cdot>k = k" | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 191 | "match_FF\<cdot>TT\<cdot>k = fail" | 
| 
4022298c1a86
change definition of match combinators for fixrec package
 huffman parents: 
30131diff
changeset | 192 | "match_FF\<cdot>\<bottom>\<cdot>k = \<bottom>" | 
| 18094 | 193 | by (simp_all add: match_FF_def) | 
| 16460 
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
 huffman parents: 
16417diff
changeset | 194 | |
| 62175 | 195 | subsection \<open>Mutual recursion\<close> | 
| 16401 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 196 | |
| 62175 | 197 | text \<open> | 
| 16401 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 198 | The following rules are used to prove unfolding theorems from | 
| 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 199 | fixed-point definitions of mutually recursive functions. | 
| 62175 | 200 | \<close> | 
| 16401 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 201 | |
| 31095 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31008diff
changeset | 202 | lemma Pair_equalI: "\<lbrakk>x \<equiv> fst p; y \<equiv> snd p\<rbrakk> \<Longrightarrow> (x, y) \<equiv> p" | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31008diff
changeset | 203 | by simp | 
| 16401 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 204 | |
| 31095 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31008diff
changeset | 205 | lemma Pair_eqD1: "(x, y) = (x', y') \<Longrightarrow> x = x'" | 
| 16401 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 206 | by simp | 
| 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 207 | |
| 31095 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31008diff
changeset | 208 | lemma Pair_eqD2: "(x, y) = (x', y') \<Longrightarrow> y = y'" | 
| 16401 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 209 | by simp | 
| 
57c35ede00b9
fixrec package now handles mutually-recursive definitions
 huffman parents: 
16229diff
changeset | 210 | |
| 31095 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31008diff
changeset | 211 | lemma def_cont_fix_eq: | 
| 40327 | 212 | "\<lbrakk>f \<equiv> fix\<cdot>(Abs_cfun F); cont F\<rbrakk> \<Longrightarrow> f = F f" | 
| 31095 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31008diff
changeset | 213 | by (simp, subst fix_eq, simp) | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31008diff
changeset | 214 | |
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31008diff
changeset | 215 | lemma def_cont_fix_ind: | 
| 40327 | 216 | "\<lbrakk>f \<equiv> fix\<cdot>(Abs_cfun F); cont F; adm P; P \<bottom>; \<And>x. P x \<Longrightarrow> P (F x)\<rbrakk> \<Longrightarrow> P f" | 
| 31095 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31008diff
changeset | 217 | by (simp add: fix_ind) | 
| 
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
 huffman parents: 
31008diff
changeset | 218 | |
| 62175 | 219 | text \<open>lemma for proving rewrite rules\<close> | 
| 16463 
342d74ca8815
fixrec shows unsolved subgoals when proofs of rewrites fail
 huffman parents: 
16460diff
changeset | 220 | |
| 
342d74ca8815
fixrec shows unsolved subgoals when proofs of rewrites fail
 huffman parents: 
16460diff
changeset | 221 | lemma ssubst_lhs: "\<lbrakk>t = s; P s = Q\<rbrakk> \<Longrightarrow> P t = Q" | 
| 
342d74ca8815
fixrec shows unsolved subgoals when proofs of rewrites fail
 huffman parents: 
16460diff
changeset | 222 | by simp | 
| 
342d74ca8815
fixrec shows unsolved subgoals when proofs of rewrites fail
 huffman parents: 
16460diff
changeset | 223 | |
| 16221 | 224 | |
| 62175 | 225 | subsection \<open>Initializing the fixrec package\<close> | 
| 16221 | 226 | |
| 69605 | 227 | ML_file \<open>Tools/holcf_library.ML\<close> | 
| 228 | ML_file \<open>Tools/fixrec.ML\<close> | |
| 16221 | 229 | |
| 62175 | 230 | method_setup fixrec_simp = \<open> | 
| 47432 | 231 | Scan.succeed (SIMPLE_METHOD' o Fixrec.fixrec_simp_tac) | 
| 62175 | 232 | \<close> "pattern prover for fixrec constants" | 
| 30131 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29530diff
changeset | 233 | |
| 62175 | 234 | setup \<open> | 
| 31738 
7b9b9ba532ca
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31095diff
changeset | 235 | Fixrec.add_matchers | 
| 69597 | 236 | [ (\<^const_name>\<open>up\<close>, \<^const_name>\<open>match_up\<close>), | 
| 237 | (\<^const_name>\<open>sinl\<close>, \<^const_name>\<open>match_sinl\<close>), | |
| 238 | (\<^const_name>\<open>sinr\<close>, \<^const_name>\<open>match_sinr\<close>), | |
| 239 | (\<^const_name>\<open>spair\<close>, \<^const_name>\<open>match_spair\<close>), | |
| 240 | (\<^const_name>\<open>Pair\<close>, \<^const_name>\<open>match_Pair\<close>), | |
| 241 | (\<^const_name>\<open>ONE\<close>, \<^const_name>\<open>match_ONE\<close>), | |
| 242 | (\<^const_name>\<open>TT\<close>, \<^const_name>\<open>match_TT\<close>), | |
| 243 | (\<^const_name>\<open>FF\<close>, \<^const_name>\<open>match_FF\<close>), | |
| 244 | (\<^const_name>\<open>bottom\<close>, \<^const_name>\<open>match_bottom\<close>) ] | |
| 62175 | 245 | \<close> | 
| 30131 
6be1be402ef0
use TheoryData to keep track of pattern match combinators
 huffman parents: 
29530diff
changeset | 246 | |
| 37109 
e67760c1b851
move unused pattern match syntax stuff into HOLCF/ex
 huffman parents: 
37108diff
changeset | 247 | hide_const (open) succeed fail run | 
| 19104 | 248 | |
| 16221 | 249 | end |