| author | Thomas Lindae <thomas.lindae@in.tum.de> | 
| Fri, 14 Jun 2024 10:21:03 +0200 | |
| changeset 81054 | 4bfcb14547c6 | 
| parent 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 66451 | 1 | (* Title: HOL/Library/Pattern_Aliases.thy | 
| 66483 | 2 | Author: Lars Hupel, Ondrej Kuncar, Tobias Nipkow | 
| 66451 | 3 | *) | 
| 4 | ||
| 5 | section \<open>Input syntax for pattern aliases (or ``as-patterns'' in Haskell)\<close> | |
| 6 | ||
| 7 | theory Pattern_Aliases | |
| 8 | imports Main | |
| 9 | begin | |
| 10 | ||
| 11 | text \<open> | |
| 12 | Most functional languages (Haskell, ML, Scala) support aliases in patterns. This allows to refer | |
| 13 | to a subpattern with a variable name. This theory implements this using a check phase. It works | |
| 66481 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 14 |   well for function definitions (see usage below). All features are packed into a @{command bundle}.
 | 
| 66451 | 15 | |
| 16 | The following caveats should be kept in mind: | |
| 69593 | 17 | \<^item> The translation expects a term of the form \<^prop>\<open>f x y = rhs\<close>, where \<open>x\<close> and \<open>y\<close> are patterns | 
| 18 | that may contain aliases. The result of the translation is a nested \<^const>\<open>Let\<close>-expression on | |
| 66451 | 19 | the right hand side. The code generator \<^emph>\<open>does not\<close> print Isabelle pattern aliases to target | 
| 20 | language pattern aliases. | |
| 21 | \<^item> The translation does not process nested equalities; only the top-level equality is translated. | |
| 22 | \<^item> Terms that do not adhere to the above shape may either stay untranslated or produce an error | |
| 23 |     message. The @{command fun} command will complain if pattern aliases are left untranslated.
 | |
| 24 | In particular, there are no checks whether the patterns are wellformed or linear. | |
| 66483 | 25 | \<^item> The corresponding uncheck phase attempts to reverse the translation (no guarantee). The | 
| 26 | additionally introduced variables are bound using a ``fake quantifier'' that does not | |
| 27 | appear in the output. | |
| 66481 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 28 | \<^item> To obtain reasonable induction principles in function definitions, the bundle also declares | 
| 69593 | 29 |     a custom congruence rule for \<^const>\<open>Let\<close> that only affects @{command fun}. This congruence
 | 
| 66485 | 30 | rule might lead to an explosion in term size (although that is rare)! In some circumstances | 
| 31 | (using \<open>let\<close> to destructure tuples), the internal construction of functions stumbles over this | |
| 32 | rule and prints an error. To mitigate this, either | |
| 69593 | 33 | \<^item> activate the bundle locally (\<^theory_text>\<open>context includes ... begin\<close>) or | 
| 34 | \<^item> rewrite the \<open>let\<close>-expression to use \<open>case\<close>: \<^term>\<open>let (a, b) = x in (b, a)\<close> becomes | |
| 35 | \<^term>\<open>case x of (a, b) \<Rightarrow> (b, a)\<close>. | |
| 66485 | 36 |   \<^item> The bundle also adds the @{thm Let_def} rule to the simpset.
 | 
| 66451 | 37 | \<close> | 
| 38 | ||
| 39 | ||
| 40 | subsection \<open>Definition\<close> | |
| 41 | ||
| 66483 | 42 | consts | 
| 43 | as :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" | |
| 44 |   fake_quant :: "('a \<Rightarrow> prop) \<Rightarrow> prop"
 | |
| 66451 | 45 | |
| 66481 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 46 | lemma let_cong_unfolding: "M = N \<Longrightarrow> f N = g N \<Longrightarrow> Let M f = Let N g" | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 47 | by simp | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 48 | |
| 66483 | 49 | translations "P" <= "CONST fake_quant (\<lambda>x. P)" | 
| 50 | ||
| 66451 | 51 | ML\<open> | 
| 52 | local | |
| 53 | ||
| 54 | fun let_typ a b = a --> (a --> b) --> b | |
| 66481 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 55 | fun as_typ a = a --> a --> a | 
| 66451 | 56 | |
| 57 | fun strip_all t = | |
| 74525 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74383diff
changeset | 58 | case try Logic.dest_all_global t of | 
| 66451 | 59 | NONE => ([], t) | 
| 60 | | SOME (var, t) => apfst (cons var) (strip_all t) | |
| 61 | ||
| 62 | fun all_Frees t = | |
| 67405 
e9ab4ad7bd15
uniform use of Standard ML op-infix -- eliminated warnings;
 wenzelm parents: 
67399diff
changeset | 63 | fold_aterms (fn Free (x, t) => insert (op =) (x, t) | _ => I) t [] | 
| 66451 | 64 | |
| 66481 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 65 | fun subst_once (old, new) t = | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 66 | let | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 67 | fun go t = | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 68 | if t = old then | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 69 | (new, true) | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 70 | else | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 71 | case t of | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 72 | u $ v => | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 73 | let | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 74 | val (u', substituted) = go u | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 75 | in | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 76 | if substituted then | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 77 | (u' $ v, true) | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 78 | else | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 79 | case go v of (v', substituted) => (u $ v', substituted) | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 80 | end | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 81 | | Abs (name, typ, t) => | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 82 | (case go t of (t', substituted) => (Abs (name, typ, t'), substituted)) | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 83 | | _ => (t, false) | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 84 | in fst (go t) end | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 85 | |
| 66483 | 86 | (* adapted from logic.ML *) | 
| 69593 | 87 | fun fake_const T = Const (\<^const_name>\<open>fake_quant\<close>, (T --> propT) --> propT); | 
| 66483 | 88 | |
| 89 | fun dependent_fake_name v t = | |
| 90 | let | |
| 91 | val x = Term.term_name v | |
| 92 | val T = Term.fastype_of v | |
| 93 | val t' = Term.abstract_over (v, t) | |
| 94 | in if Term.is_dependent t' then fake_const T $ Abs (x, T, t') else t end | |
| 95 | ||
| 66451 | 96 | in | 
| 97 | ||
| 98 | fun check_pattern_syntax t = | |
| 99 | case strip_all t of | |
| 74383 | 100 | (vars, \<^Const>\<open>Trueprop\<close> $ (Const (\<^const_name>\<open>HOL.eq\<close>, _) $ lhs $ rhs)) => | 
| 66451 | 101 | let | 
| 69593 | 102 | fun go (Const (\<^const_name>\<open>as\<close>, _) $ pat $ var, rhs) = | 
| 66451 | 103 | let | 
| 104 | val (pat', rhs') = go (pat, rhs) | |
| 66481 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 105 | val _ = if is_Free var then () else error "Right-hand side of =: must be a free variable" | 
| 66451 | 106 | val rhs'' = | 
| 69593 | 107 | Const (\<^const_name>\<open>Let\<close>, let_typ (fastype_of var) (fastype_of rhs)) $ | 
| 66451 | 108 | pat' $ lambda var rhs' | 
| 109 | in | |
| 110 | (pat', rhs'') | |
| 111 | end | |
| 112 | | go (t $ u, rhs) = | |
| 113 | let | |
| 114 | val (t', rhs') = go (t, rhs) | |
| 115 | val (u', rhs'') = go (u, rhs') | |
| 116 | in (t' $ u', rhs'') end | |
| 117 | | go (t, rhs) = (t, rhs) | |
| 118 | ||
| 119 | val (lhs', rhs') = go (lhs, rhs) | |
| 120 | ||
| 121 | val res = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs', rhs')) | |
| 122 | ||
| 67405 
e9ab4ad7bd15
uniform use of Standard ML op-infix -- eliminated warnings;
 wenzelm parents: 
67399diff
changeset | 123 | val frees = filter (member (op =) vars) (all_Frees res) | 
| 66451 | 124 |       in fold (fn v => Logic.dependent_all_name ("", v)) (map Free frees) res end
 | 
| 125 | | _ => t | |
| 126 | ||
| 66481 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 127 | fun uncheck_pattern_syntax ctxt t = | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 128 | case strip_all t of | 
| 74383 | 129 | (vars, \<^Const>\<open>Trueprop\<close> $ (Const (\<^const_name>\<open>HOL.eq\<close>, _) $ lhs $ rhs)) => | 
| 66481 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 130 | let | 
| 66483 | 131 | (* restricted to going down abstractions; ignores eta-contracted rhs *) | 
| 69593 | 132 | fun go lhs (rhs as Const (\<^const_name>\<open>Let\<close>, _) $ pat $ Abs (name, typ, t)) ctxt frees = | 
| 66481 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 133 | if exists_subterm (fn t' => t' = pat) lhs then | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 134 | let | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 135 | val ([name'], ctxt') = Variable.variant_fixes [name] ctxt | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 136 | val free = Free (name', typ) | 
| 69593 | 137 | val subst = (pat, Const (\<^const_name>\<open>as\<close>, as_typ typ) $ pat $ free) | 
| 66481 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 138 | val lhs' = subst_once subst lhs | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 139 | val rhs' = subst_bound (free, t) | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 140 | in | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 141 | go lhs' rhs' ctxt' (Free (name', typ) :: frees) | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 142 | end | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 143 | else | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 144 | (lhs, rhs, ctxt, frees) | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 145 | | go lhs rhs ctxt frees = (lhs, rhs, ctxt, frees) | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 146 | |
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 147 | val (lhs', rhs', _, frees) = go lhs rhs ctxt [] | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 148 | |
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 149 | val res = | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 150 | HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs', rhs')) | 
| 66483 | 151 |           |> fold (fn v => Logic.dependent_all_name ("", v)) (map Free vars)
 | 
| 152 | |> fold dependent_fake_name frees | |
| 66481 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 153 | in | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 154 | if null frees then t else res | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 155 | end | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 156 | | _ => t | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 157 | |
| 66451 | 158 | end | 
| 159 | \<close> | |
| 160 | ||
| 161 | bundle pattern_aliases begin | |
| 162 | ||
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
80874diff
changeset | 163 | notation as (infixr \<open>=:\<close> 1) | 
| 66451 | 164 | |
| 165 | declaration \<open>K (Syntax_Phases.term_check 98 "pattern_syntax" (K (map check_pattern_syntax)))\<close> | |
| 66481 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 166 | declaration \<open>K (Syntax_Phases.term_uncheck 98 "pattern_syntax" (map o uncheck_pattern_syntax))\<close> | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 167 | |
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 168 | declare let_cong_unfolding [fundef_cong] | 
| 66485 | 169 | declare Let_def [simp] | 
| 66451 | 170 | |
| 171 | end | |
| 172 | ||
| 66483 | 173 | hide_const as | 
| 174 | hide_const fake_quant | |
| 66451 | 175 | |
| 176 | ||
| 177 | subsection \<open>Usage\<close> | |
| 178 | ||
| 179 | context includes pattern_aliases begin | |
| 180 | ||
| 181 | text \<open>Not very useful for plain definitions, but works anyway.\<close> | |
| 182 | ||
| 183 | private definition "test_1 x (y =: z) = y + z" | |
| 184 | ||
| 185 | lemma "test_1 x y = y + y" | |
| 186 | by (rule test_1_def[unfolded Let_def]) | |
| 187 | ||
| 188 | text \<open>Very useful for function definitions.\<close> | |
| 189 | ||
| 190 | private fun test_2 where | |
| 66481 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 191 | "test_2 (y # (y' # ys =: x') =: x) = x @ x' @ x'" | | 
| 66451 | 192 | "test_2 _ = []" | 
| 193 | ||
| 66481 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 194 | lemma "test_2 (y # y' # ys) = (y # y' # ys) @ (y' # ys) @ (y' # ys)" | 
| 66451 | 195 | by (rule test_2.simps[unfolded Let_def]) | 
| 196 | ||
| 66481 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 197 | ML\<open> | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 198 | let | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 199 | val actual = | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 200 |     @{thm test_2.simps(1)}
 | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 201 | |> Thm.prop_of | 
| 80874 
9af593e9e454
prefer Pretty.pure_string_of to produce output without markup, instead of cleaning output afterwards;
 wenzelm parents: 
80820diff
changeset | 202 | |> Syntax.pretty_term \<^context> | 
| 
9af593e9e454
prefer Pretty.pure_string_of to produce output without markup, instead of cleaning output afterwards;
 wenzelm parents: 
80820diff
changeset | 203 | |> Pretty.pure_string_of | 
| 66483 | 204 | val expected = "test_2 (?y # (?y' # ?ys =: x') =: x) = x @ x' @ x'" | 
| 69593 | 205 | in \<^assert> (actual = expected) end | 
| 66481 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 206 | \<close> | 
| 
d35f7a9f92e2
output syntax for pattern aliases
 Lars Hupel <lars.hupel@mytum.de> parents: 
66451diff
changeset | 207 | |
| 66451 | 208 | end | 
| 209 | ||
| 67399 | 210 | end |