| author | wenzelm | 
| Wed, 02 Sep 2009 14:11:45 +0200 | |
| changeset 32490 | 6a48db3e627c | 
| parent 30364 | 577edc39b501 | 
| child 33385 | fb2358edcfb6 | 
| permissions | -rw-r--r-- | 
| 24259 | 1 | (* Title: Pure/primitive_defs.ML | 
| 2 | Author: Makarius | |
| 3 | ||
| 4 | Primitive definition forms. | |
| 5 | *) | |
| 6 | ||
| 7 | signature PRIMITIVE_DEFS = | |
| 8 | sig | |
| 24981 
4ec3f95190bf
dest/cert_def: replaced Pretty.pp by explicit Proof.context;
 wenzelm parents: 
24259diff
changeset | 9 | val dest_def: Proof.context -> (term -> bool) -> (string -> bool) -> (string -> bool) -> | 
| 24259 | 10 | term -> (term * term) * term | 
| 11 | val abs_def: term -> term * term | |
| 12 | val mk_defpair: term * term -> string * term | |
| 13 | end; | |
| 14 | ||
| 15 | structure PrimitiveDefs: PRIMITIVE_DEFS = | |
| 16 | struct | |
| 17 | ||
| 18 | fun term_kind (Const _) = "existing constant " | |
| 19 | | term_kind (Free _) = "free variable " | |
| 20 | | term_kind (Bound _) = "bound variable " | |
| 21 | | term_kind _ = ""; | |
| 22 | ||
| 23 | (*c x == t[x] to !!x. c x == t[x]*) | |
| 24981 
4ec3f95190bf
dest/cert_def: replaced Pretty.pp by explicit Proof.context;
 wenzelm parents: 
24259diff
changeset | 24 | fun dest_def ctxt check_head is_fixed is_fixedT eq = | 
| 24259 | 25 | let | 
| 26 | fun err msg = raise TERM (msg, [eq]); | |
| 27 | val eq_vars = Term.strip_all_vars eq; | |
| 28 | val eq_body = Term.strip_all_body eq; | |
| 29 | ||
| 30 | val display_terms = | |
| 24981 
4ec3f95190bf
dest/cert_def: replaced Pretty.pp by explicit Proof.context;
 wenzelm parents: 
24259diff
changeset | 31 | commas_quote o map (Syntax.string_of_term ctxt o Syntax.bound_vars eq_vars); | 
| 
4ec3f95190bf
dest/cert_def: replaced Pretty.pp by explicit Proof.context;
 wenzelm parents: 
24259diff
changeset | 32 | val display_types = commas_quote o map (Syntax.string_of_typ ctxt); | 
| 24259 | 33 | |
| 34 | val (raw_lhs, rhs) = Logic.dest_equals eq_body handle TERM _ => err "Not a meta-equality (==)"; | |
| 35 | val lhs = Envir.beta_eta_contract raw_lhs; | |
| 36 | val (head, args) = Term.strip_comb lhs; | |
| 37 | val head_tfrees = Term.add_tfrees head []; | |
| 38 | ||
| 39 | fun check_arg (Bound _) = true | |
| 40 | | check_arg (Free (x, _)) = not (is_fixed x) | |
| 41 |       | check_arg (Const ("TYPE", Type ("itself", [TFree _]))) = true
 | |
| 42 | | check_arg _ = false; | |
| 43 | fun close_arg (Bound _) t = t | |
| 44 | | close_arg x t = Term.all (Term.fastype_of x) $ lambda x t; | |
| 45 | ||
| 46 | val lhs_bads = filter_out check_arg args; | |
| 47 | val lhs_dups = duplicates (op aconv) args; | |
| 48 | val rhs_extras = Term.fold_aterms (fn v as Free (x, _) => | |
| 49 | if is_fixed x orelse member (op aconv) args v then I | |
| 50 | else insert (op aconv) v | _ => I) rhs []; | |
| 51 | val rhs_extrasT = Term.fold_aterms (Term.fold_types (fn v as TFree (a, S) => | |
| 52 | if is_fixedT a orelse member (op =) head_tfrees (a, S) then I | |
| 53 | else insert (op =) v | _ => I)) rhs []; | |
| 54 | in | |
| 55 | if not (check_head head) then | |
| 56 |       err ("Bad head of lhs: " ^ term_kind head ^ display_terms [head])
 | |
| 57 | else if not (null lhs_bads) then | |
| 58 |       err ("Bad arguments on lhs: " ^ display_terms lhs_bads)
 | |
| 59 | else if not (null lhs_dups) then | |
| 60 |       err ("Duplicate arguments on lhs: " ^ display_terms lhs_dups)
 | |
| 61 | else if not (null rhs_extras) then | |
| 62 |       err ("Extra variables on rhs: " ^ display_terms rhs_extras)
 | |
| 63 | else if not (null rhs_extrasT) then | |
| 64 |       err ("Extra type variables on rhs: " ^ display_types rhs_extrasT)
 | |
| 65 | else if exists_subterm (fn t => t aconv head) rhs then | |
| 66 | err "Entity to be defined occurs on rhs" | |
| 67 | else | |
| 68 | ((lhs, rhs), fold_rev close_arg args (Term.list_all (eq_vars, (Logic.mk_equals (lhs, rhs))))) | |
| 69 | end; | |
| 70 | ||
| 71 | (*!!x. c x == t[x] to c == %x. t[x]*) | |
| 72 | fun abs_def eq = | |
| 73 | let | |
| 74 | val body = Term.strip_all_body eq; | |
| 75 | val vars = map Free (Term.rename_wrt_term body (Term.strip_all_vars eq)); | |
| 76 | val (lhs, rhs) = Logic.dest_equals (Term.subst_bounds (vars, body)); | |
| 77 | val (lhs', args) = Term.strip_comb lhs; | |
| 78 | val rhs' = Term.list_abs_free (map Term.dest_Free args, rhs); | |
| 79 | in (lhs', rhs') end; | |
| 80 | ||
| 81 | fun mk_defpair (lhs, rhs) = | |
| 82 | (case Term.head_of lhs of | |
| 83 | Const (name, _) => | |
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 84 | (Long_Name.base_name name ^ "_def", Logic.mk_equals (lhs, rhs)) | 
| 24259 | 85 |   | _ => raise TERM ("Malformed definition: head of lhs not a constant", [lhs, rhs]));
 | 
| 86 | ||
| 87 | end; |