| author | wenzelm | 
| Mon, 25 Sep 2023 20:56:44 +0200 | |
| changeset 78709 | ebafb2daabb7 | 
| parent 74383 | 107941e8fa01 | 
| permissions | -rw-r--r-- | 
| 66646 | 1 | (* Title: HOL/Tools/Nunchaku/nunchaku_reconstruct.ML | 
| 66614 
1f1c5d85d232
moved Nunchaku to Main; the goal is to move Nitpick out in the next 1-2 years
 blanchet parents: 
64389diff
changeset | 2 | Author: Jasmin Blanchette, VU Amsterdam | 
| 
1f1c5d85d232
moved Nunchaku to Main; the goal is to move Nitpick out in the next 1-2 years
 blanchet parents: 
64389diff
changeset | 3 | Copyright 2015, 2016, 2017 | 
| 64389 | 4 | |
| 5 | Reconstruction of Nunchaku models in Isabelle/HOL. | |
| 6 | *) | |
| 7 | ||
| 8 | signature NUNCHAKU_RECONSTRUCT = | |
| 9 | sig | |
| 10 | type nun_model = Nunchaku_Model.nun_model | |
| 11 | ||
| 12 | type typ_entry = typ * term list | |
| 13 | type term_entry = term * term | |
| 14 | ||
| 15 | type isa_model = | |
| 16 |     {type_model: typ_entry list,
 | |
| 17 | free_model: term_entry list, | |
| 18 | pat_complete_model: term_entry list, | |
| 19 | pat_incomplete_model: term_entry list, | |
| 66623 | 20 | skolem_model: term_entry list, | 
| 21 | auxiliary_model: term_entry list} | |
| 64389 | 22 | |
| 23 | val str_of_isa_model: Proof.context -> isa_model -> string | |
| 24 | ||
| 25 | val isa_model_of_nun: Proof.context -> term list -> (typ option * string list) list -> | |
| 26 | nun_model -> isa_model | |
| 66623 | 27 | val clean_up_isa_model: Proof.context -> isa_model -> isa_model | 
| 64389 | 28 | end; | 
| 29 | ||
| 30 | structure Nunchaku_Reconstruct : NUNCHAKU_RECONSTRUCT = | |
| 31 | struct | |
| 32 | ||
| 33 | open Nunchaku_Util; | |
| 34 | open Nunchaku_Problem; | |
| 35 | open Nunchaku_Translate; | |
| 36 | open Nunchaku_Model; | |
| 37 | ||
| 38 | type typ_entry = typ * term list; | |
| 39 | type term_entry = term * term; | |
| 40 | ||
| 41 | type isa_model = | |
| 42 |   {type_model: typ_entry list,
 | |
| 43 | free_model: term_entry list, | |
| 44 | pat_complete_model: term_entry list, | |
| 45 | pat_incomplete_model: term_entry list, | |
| 66623 | 46 | skolem_model: term_entry list, | 
| 47 | auxiliary_model: term_entry list}; | |
| 64389 | 48 | |
| 49 | val anonymousN = "anonymous"; | |
| 50 | val irrelevantN = "irrelevant"; | |
| 51 | val unparsableN = "unparsable"; | |
| 52 | ||
| 53 | val nun_arrow_exploded = String.explode nun_arrow; | |
| 54 | ||
| 67405 
e9ab4ad7bd15
uniform use of Standard ML op-infix -- eliminated warnings;
 wenzelm parents: 
67399diff
changeset | 55 | val is_ty_meta = member (op =) (String.explode "()->,"); | 
| 64389 | 56 | |
| 57 | fun next_token_lowlevel [] = (End_of_Stream, []) | |
| 58 | | next_token_lowlevel (c :: cs) = | |
| 59 | if Char.isSpace c then | |
| 60 | next_token_lowlevel cs | |
| 61 | else if not (is_ty_meta c) then | |
| 62 | let val n = find_index (Char.isSpace orf is_ty_meta) cs in | |
| 63 | (if n = ~1 then (cs, []) else chop n cs) | |
| 64 | |>> (cons c #> String.implode #> ident_of_str #> Ident) | |
| 65 | end | |
| 66 | else if is_prefix (op =) nun_arrow_exploded (c :: cs) then | |
| 67 | (Ident nun_arrow, tl cs) | |
| 68 | else | |
| 69 | (Symbol (String.str c), cs); | |
| 70 | ||
| 71 | val tokenize_lowlevel = | |
| 72 | let | |
| 73 | fun toks cs = | |
| 74 | (case next_token_lowlevel cs of | |
| 75 | (End_of_Stream, []) => [] | |
| 76 | | (tok, cs') => tok :: toks cs'); | |
| 77 | in | |
| 78 | toks o String.explode | |
| 79 | end; | |
| 80 | ||
| 81 | fun parse_lowlevel_ty tok = | |
| 82 | (Scan.optional | |
| 83 |      (parse_sym "(" |-- Scan.repeat (parse_lowlevel_ty --| Scan.option (parse_sym ",")) --|
 | |
| 84 | parse_sym ")") | |
| 85 | [] | |
| 86 | -- parse_ident >> (swap #> NType)) tok; | |
| 87 | ||
| 88 | val ty_of_lowlevel_str = fst o parse_lowlevel_ty o tokenize_lowlevel; | |
| 89 | ||
| 90 | fun ident_of_const (NConst (id, _, _)) = id | |
| 91 | | ident_of_const _ = nun_dummy; | |
| 92 | ||
| 93 | fun str_of_typ_entry ctxt (T, ts) = | |
| 94 | "type " ^ Syntax.string_of_typ ctxt T ^ | |
| 95 |   " := {" ^ commas (map (Syntax.string_of_term ctxt) ts) ^ "}.";
 | |
| 96 | ||
| 97 | fun str_of_term_entry ctxt (tm, value) = | |
| 98 | "val " ^ Syntax.string_of_term ctxt tm ^ " := " ^ Syntax.string_of_term ctxt value ^ "."; | |
| 99 | ||
| 100 | fun str_of_isa_model ctxt | |
| 66623 | 101 |     {type_model, free_model, pat_complete_model, pat_incomplete_model, skolem_model,
 | 
| 102 | auxiliary_model} = | |
| 64389 | 103 | map (str_of_typ_entry ctxt) type_model @ "" :: | 
| 104 | map (str_of_term_entry ctxt) free_model @ "" :: | |
| 105 | map (str_of_term_entry ctxt) pat_complete_model @ "" :: | |
| 106 | map (str_of_term_entry ctxt) pat_incomplete_model @ "" :: | |
| 66623 | 107 | map (str_of_term_entry ctxt) skolem_model @ "" :: | 
| 108 | map (str_of_term_entry ctxt) auxiliary_model | |
| 64389 | 109 | |> cat_lines; | 
| 110 | ||
| 111 | fun typ_of_nun ctxt = | |
| 112 | let | |
| 113 | fun typ_of (NType (id, tys)) = | |
| 114 | let val Ts = map typ_of tys in | |
| 115 | if id = nun_dummy then | |
| 116 | dummyT | |
| 117 | else if id = nun_prop then | |
| 69593 | 118 | \<^typ>\<open>bool\<close> | 
| 64389 | 119 | else if id = nun_arrow then | 
| 69593 | 120 | Type (\<^type_name>\<open>fun\<close>, Ts) | 
| 64389 | 121 | else | 
| 122 | (case try str_of_nun_tconst id of | |
| 123 | SOME (args, s) => | |
| 124 | let val tys' = map ty_of_lowlevel_str args in | |
| 125 | Type (s, map typ_of (tys' @ tys)) | |
| 126 | end | |
| 127 | | NONE => | |
| 128 | (case try str_of_nun_tfree id of | |
| 129 | SOME s => TFree (Proof_Context.check_tfree ctxt (flip_quote s, dummyS)) | |
| 130 |             | NONE => raise Fail ("unknown type constructor: " ^ quote (str_of_ident id))))
 | |
| 131 | end; | |
| 132 | in | |
| 133 | typ_of | |
| 134 | end; | |
| 135 | ||
| 136 | fun one_letter_of s = | |
| 137 | let val c = String.sub (Long_Name.base_name s, 0) in | |
| 138 | String.str (if Char.isAlpha c then c else #"x") | |
| 139 | end; | |
| 140 | ||
| 141 | fun base_of_typ (Type (s, _)) = s | |
| 142 | | base_of_typ (TFree (s, _)) = flip_quote s | |
| 143 | | base_of_typ (TVar ((s, _), _)) = flip_quote s; | |
| 144 | ||
| 145 | fun term_of_nun ctxt atomss = | |
| 146 | let | |
| 147 | val thy = Proof_Context.theory_of ctxt; | |
| 148 | ||
| 149 | val typ_of = typ_of_nun ctxt; | |
| 150 | ||
| 151 | fun nth_atom T j = | |
| 152 | let val ss = these (triple_lookup (typ_match thy) atomss T) in | |
| 153 | if j >= 0 andalso j < length ss then nth ss j | |
| 154 | else one_letter_of (base_of_typ T) ^ nat_subscript (j + 1) | |
| 155 | end; | |
| 156 | ||
| 157 | fun term_of _ (NAtom (j, ty)) = | |
| 158 | let val T = typ_of ty in Var ((nth_atom T j, 0), T) end | |
| 159 | | term_of bounds (NConst (id, tys0, ty)) = | |
| 160 | if id = nun_conj then | |
| 161 | HOLogic.conj | |
| 66623 | 162 | else if id = nun_choice then | 
| 69593 | 163 | Const (\<^const_name>\<open>Eps\<close>, typ_of ty) | 
| 64389 | 164 | else if id = nun_disj then | 
| 165 | HOLogic.disj | |
| 166 | else if id = nun_equals then | |
| 69593 | 167 | Const (\<^const_name>\<open>HOL.eq\<close>, typ_of ty) | 
| 64389 | 168 | else if id = nun_false then | 
| 74383 | 169 | \<^Const>\<open>False\<close> | 
| 64389 | 170 | else if id = nun_if then | 
| 69593 | 171 | Const (\<^const_name>\<open>If\<close>, typ_of ty) | 
| 64389 | 172 | else if id = nun_implies then | 
| 74383 | 173 | \<^Const>\<open>implies\<close> | 
| 66623 | 174 | else if id = nun_not then | 
| 175 | HOLogic.Not | |
| 64389 | 176 | else if id = nun_unique then | 
| 69593 | 177 | Const (\<^const_name>\<open>The\<close>, typ_of ty) | 
| 64389 | 178 | else if id = nun_unique_unsafe then | 
| 69593 | 179 | Const (\<^const_name>\<open>The_unsafe\<close>, typ_of ty) | 
| 64389 | 180 | else if id = nun_true then | 
| 74383 | 181 | \<^Const>\<open>True\<close> | 
| 66623 | 182 | else if String.isPrefix nun_dollar_anon_fun_prefix id then | 
| 183 | let val j = Int.fromString (unprefix nun_dollar_anon_fun_prefix id) |> the_default ~1 in | |
| 64389 | 184 | Var ((anonymousN ^ nat_subscript (j + 1), 0), typ_of ty) | 
| 185 | end | |
| 66616 | 186 | else if String.isPrefix nun_irrelevant id then | 
| 187 | Var ((irrelevantN ^ nat_subscript (Value.parse_int (unprefix nun_irrelevant id)), 0), | |
| 188 | dummyT) | |
| 64389 | 189 | else if id = nun_unparsable then | 
| 66616 | 190 | Var ((unparsableN, 0), typ_of ty) | 
| 64389 | 191 | else | 
| 192 | (case try str_of_nun_const id of | |
| 193 | SOME (args, s) => | |
| 194 | let val tys = map ty_of_lowlevel_str args in | |
| 195 | Sign.mk_const thy (s, map typ_of (tys @ tys0)) | |
| 196 | end | |
| 197 | | NONE => | |
| 198 | (case try str_of_nun_free id of | |
| 199 | SOME s => Free (s, typ_of ty) | |
| 200 | | NONE => | |
| 201 | (case try str_of_nun_var id of | |
| 202 | SOME s => Var ((s, 0), typ_of ty) | |
| 203 | | NONE => | |
| 204 | (case find_index (fn bound => ident_of_const bound = id) bounds of | |
| 205 | ~1 => Var ((str_of_ident id, 0), typ_of ty) (* shouldn't happen? *) | |
| 206 | | j => Bound j)))) | |
| 207 | | term_of bounds (NAbs (var, body)) = | |
| 208 | let val T = typ_of (safe_ty_of var) in | |
| 209 | Abs (one_letter_of (base_of_typ T), T, term_of (var :: bounds) body) | |
| 210 | end | |
| 211 | | term_of bounds (NApp (func, arg)) = | |
| 212 | let | |
| 213 | fun same () = term_of bounds func $ term_of bounds arg; | |
| 214 | in | |
| 215 | (case (func, arg) of | |
| 216 | (NConst (id, _, _), NAbs _) => | |
| 217 | if id = nun_mu then | |
| 218 | let val Abs (s, T, body) = term_of bounds arg in | |
| 69593 | 219 | Const (\<^const_name>\<open>The\<close>, (T --> HOLogic.boolT) --> T) | 
| 64389 | 220 | $ Abs (s, T, HOLogic.eq_const T $ Bound 0 $ body) | 
| 221 | end | |
| 222 | else | |
| 223 | same () | |
| 224 | | _ => same ()) | |
| 225 | end | |
| 226 | | term_of _ (NMatch _) = raise Fail "unexpected match"; | |
| 66633 
ec8fceca7fb6
nicer numeral output for nats and ints in Nunchaku
 blanchet parents: 
66623diff
changeset | 227 | |
| 74383 | 228 | fun rewrite_numbers (t as \<^Const>\<open>Suc\<close> $ _) = | 
| 66633 
ec8fceca7fb6
nicer numeral output for nats and ints in Nunchaku
 blanchet parents: 
66623diff
changeset | 229 | (case try HOLogic.dest_nat t of | 
| 74383 | 230 | SOME n => HOLogic.mk_number \<^Type>\<open>nat\<close> n | 
| 66633 
ec8fceca7fb6
nicer numeral output for nats and ints in Nunchaku
 blanchet parents: 
66623diff
changeset | 231 | | NONE => t) | 
| 74383 | 232 | | rewrite_numbers \<^Const_>\<open>Abs_Integ for \<^Const>\<open>Pair \<^Type>\<open>nat\<close> \<^Type>\<open>nat\<close> for t u\<close>\<close> = | 
| 233 | HOLogic.mk_number \<^Type>\<open>int\<close> (HOLogic.dest_nat t - HOLogic.dest_nat u) | |
| 66633 
ec8fceca7fb6
nicer numeral output for nats and ints in Nunchaku
 blanchet parents: 
66623diff
changeset | 234 | | rewrite_numbers (t $ u) = rewrite_numbers t $ rewrite_numbers u | 
| 
ec8fceca7fb6
nicer numeral output for nats and ints in Nunchaku
 blanchet parents: 
66623diff
changeset | 235 | | rewrite_numbers (Abs (s, T, t)) = Abs (s, T, rewrite_numbers t) | 
| 
ec8fceca7fb6
nicer numeral output for nats and ints in Nunchaku
 blanchet parents: 
66623diff
changeset | 236 | | rewrite_numbers t = t; | 
| 64389 | 237 | in | 
| 238 | term_of [] | |
| 66633 
ec8fceca7fb6
nicer numeral output for nats and ints in Nunchaku
 blanchet parents: 
66623diff
changeset | 239 | #> rewrite_numbers | 
| 64389 | 240 | end; | 
| 241 | ||
| 242 | fun isa_typ_entry_of_nun ctxt atomss (ty, atoms) = | |
| 243 | (typ_of_nun ctxt ty, map (term_of_nun ctxt atomss) atoms); | |
| 244 | ||
| 245 | fun isa_term_entry_of_nun ctxt atomss (tm, value) = | |
| 246 | (term_of_nun ctxt atomss tm, term_of_nun ctxt atomss value); | |
| 247 | ||
| 66623 | 248 | fun isa_model_of_nun ctxt pat_completes atomss | 
| 249 |     {type_model, const_model, skolem_model, auxiliary_model} =
 | |
| 64389 | 250 | let | 
| 251 | val free_and_const_model = map (isa_term_entry_of_nun ctxt atomss) const_model; | |
| 252 | val (free_model, (pat_complete_model, pat_incomplete_model)) = | |
| 253 | List.partition (is_Free o fst) free_and_const_model | |
| 254 | ||> List.partition (member (op aconv) pat_completes o fst); | |
| 255 | in | |
| 256 |     {type_model = map (isa_typ_entry_of_nun ctxt atomss) type_model, free_model = free_model,
 | |
| 257 | pat_complete_model = pat_complete_model, pat_incomplete_model = pat_incomplete_model, | |
| 66623 | 258 | skolem_model = map (isa_term_entry_of_nun ctxt atomss) skolem_model, | 
| 259 | auxiliary_model = map (isa_term_entry_of_nun ctxt atomss) auxiliary_model} | |
| 260 | end; | |
| 261 | ||
| 262 | fun clean_up_isa_model ctxt {type_model, free_model, pat_complete_model, pat_incomplete_model,
 | |
| 263 | skolem_model, auxiliary_model} = | |
| 264 | let | |
| 265 | val pat_complete_model' = if Config.get ctxt show_consts then pat_complete_model else []; | |
| 266 | val pat_incomplete_model' = pat_incomplete_model | |
| 69593 | 267 | |> filter_out (can (fn Const (\<^const_name>\<open>unreachable\<close>, _) => ()) o fst); | 
| 66623 | 268 | |
| 269 | val term_model = free_model @ pat_complete_model' @ pat_incomplete_model' @ | |
| 270 | skolem_model @ auxiliary_model; | |
| 271 | ||
| 272 | (* Incomplete test: Can be led astray by references between the auxiliaries. *) | |
| 273 | fun is_auxiliary_referenced (aux, _) = | |
| 274 | exists (fn (lhs, rhs) => | |
| 275 | not (lhs aconv aux) andalso exists_subterm (curry (op aconv) aux) rhs) | |
| 276 | term_model; | |
| 277 | ||
| 278 | val auxiliary_model' = filter is_auxiliary_referenced auxiliary_model; | |
| 279 | in | |
| 280 |     {type_model = type_model, free_model = free_model, pat_complete_model = pat_complete_model',
 | |
| 281 | pat_incomplete_model = pat_incomplete_model', skolem_model = skolem_model, | |
| 282 | auxiliary_model = auxiliary_model'} | |
| 64389 | 283 | end; | 
| 284 | ||
| 285 | end; |