1.1 --- a/src/HOL/Tools/typedef_package.ML Sun Jul 15 14:47:28 2001 +0200
1.2 +++ b/src/HOL/Tools/typedef_package.ML Sun Jul 15 14:48:36 2001 +0200
1.3 @@ -69,40 +69,20 @@
1.4 fun message s = if ! quiet_mode then () else writeln s;
1.5
1.6
1.7 -(* non-emptiness of set *) (*exception ERROR*)
1.8 +(* prove_nonempty -- tactical version *) (*exception ERROR*)
1.9
1.10 -fun check_nonempty cset thm =
1.11 - let
1.12 - val {t, sign, maxidx, ...} = Thm.rep_cterm cset;
1.13 - val {prop, ...} = Thm.rep_thm (Thm.transfer_sg sign (Drule.standard thm));
1.14 - val matches = Pattern.matches (Sign.tsig_of sign);
1.15 - in
1.16 - (case try (HOLogic.dest_mem o HOLogic.dest_Trueprop) prop of
1.17 - None => raise ERROR
1.18 - | Some (_, A) => if matches (Logic.incr_indexes ([], maxidx) A, t) then () else raise ERROR)
1.19 - end handle ERROR => error ("Bad non-emptiness theorem " ^ Display.string_of_thm thm ^
1.20 - "\nfor set " ^ quote (Display.string_of_cterm cset));
1.21 -
1.22 -fun goal_nonempty ex cset =
1.23 - let
1.24 - val {T = setT, t = A, maxidx, sign} = Thm.rep_cterm cset;
1.25 - val T = HOLogic.dest_setT setT;
1.26 - val tm =
1.27 - if ex then HOLogic.mk_exists ("x", T, HOLogic.mk_mem (Free ("x", T), A))
1.28 - else HOLogic.mk_mem (Var (("x", maxidx + 1), T), A); (*old-style version*)
1.29 - in Thm.cterm_of sign (HOLogic.mk_Trueprop tm) end;
1.30 -
1.31 -fun prove_nonempty thy cset (witn_names, witn_thms, witn_tac) =
1.32 +fun prove_nonempty thy cset goal (witn_names, witn_thms, witn_tac) =
1.33 let
1.34 val is_def = Logic.is_equals o #prop o Thm.rep_thm;
1.35 val thms = PureThy.get_thmss thy witn_names @ witn_thms;
1.36 val tac =
1.37 + rtac exI 1 THEN
1.38 TRY (rewrite_goals_tac (filter is_def thms)) THEN
1.39 TRY (REPEAT_FIRST (resolve_tac (filter_out is_def thms))) THEN
1.40 if_none witn_tac (TRY (ALLGOALS (CLASET' blast_tac)));
1.41 in
1.42 message ("Proving non-emptiness of set " ^ quote (string_of_cterm cset) ^ " ...");
1.43 - prove_goalw_cterm [] (goal_nonempty false cset) (K [tac])
1.44 + prove_goalw_cterm [] (cterm_of (sign_of thy) goal) (K [tac])
1.45 end handle ERROR => error ("Failed to prove non-emptiness of " ^ quote (string_of_cterm cset));
1.46
1.47
1.48 @@ -129,7 +109,10 @@
1.49 val rhs_tfrees = term_tfrees set;
1.50 val oldT = HOLogic.dest_setT setT handle TYPE _ =>
1.51 error ("Not a set type: " ^ quote (Sign.string_of_typ sign setT));
1.52 - val cset_pat = Thm.cterm_of sign (Var ((name, 0), setT));
1.53 + fun mk_nonempty A =
1.54 + HOLogic.mk_Trueprop (HOLogic.mk_exists ("x", oldT, HOLogic.mk_mem (Free ("x", oldT), A)));
1.55 + val goal = mk_nonempty set;
1.56 + val goal_pat = mk_nonempty (Var ((name, 0), setT));
1.57
1.58 (*lhs*)
1.59 val lhs_tfrees = map (fn v => (v, if_none (assoc (rhs_tfrees, v)) HOLogic.termS)) vs;
1.60 @@ -151,12 +134,11 @@
1.61 val typedef_name = "type_definition_" ^ name;
1.62 val typedefC =
1.63 Const (type_definitionN, (newT --> oldT) --> (oldT --> newT) --> setT --> HOLogic.boolT);
1.64 - val typedef_prop = HOLogic.mk_Trueprop (typedefC $ RepC $ AbsC $ set');
1.65 + val typedef_prop =
1.66 + Logic.mk_implies (goal, HOLogic.mk_Trueprop (typedefC $ RepC $ AbsC $ set'));
1.67
1.68 - (*theory extender*)
1.69 - fun do_typedef super_theory theory =
1.70 + fun typedef_att (theory, nonempty) =
1.71 theory
1.72 - |> Theory.assert_super super_theory
1.73 |> add_typedecls [(t, vs, mx)]
1.74 |> Theory.add_consts_i
1.75 ((if no_def then [] else [(name, setT, NoSyn)]) @
1.76 @@ -164,10 +146,11 @@
1.77 (Abs_name, oldT --> newT, NoSyn)])
1.78 |> (if no_def then I else (#1 oo (PureThy.add_defs_i false o map Thm.no_attributes))
1.79 [Logic.mk_defpair (setC, set)])
1.80 - |> PureThy.add_axioms_i [((typedef_name, typedef_prop), [])]
1.81 - |> (fn (theory', typedef_ax) =>
1.82 - let fun make th = standard (th OF typedef_ax) in
1.83 - rpair (hd typedef_ax) (theory'
1.84 + |> PureThy.add_axioms_i [((typedef_name, typedef_prop),
1.85 + [apsnd (fn cond_axm => standard (nonempty RS cond_axm))])]
1.86 + |> (fn (theory', axm) =>
1.87 + let fun make th = standard (th OF axm) in
1.88 + rpair (hd axm) (theory'
1.89 |> (#1 oo PureThy.add_thms)
1.90 ([((Rep_name, make Rep), []),
1.91 ((Rep_name ^ "_inverse", make Rep_inverse), []),
1.92 @@ -182,8 +165,7 @@
1.93 [RuleCases.case_names [Rep_name], InductAttrib.induct_set_global full_name]),
1.94 ((Abs_name ^ "_induct", make Abs_induct),
1.95 [RuleCases.case_names [Abs_name], InductAttrib.induct_type_global full_tname])]))
1.96 - end)
1.97 - handle ERROR => err_in_typedef name;
1.98 + end);
1.99
1.100
1.101 (* errors *)
1.102 @@ -207,19 +189,24 @@
1.103 | xs => ["Illegal variables on rhs: " ^ show_names (map dest_Free xs)]);
1.104
1.105 val errs = illegal_vars @ dup_lhs_tfrees @ extra_rhs_tfrees @ illegal_frees;
1.106 - in
1.107 - if null errs then () else error (cat_lines errs);
1.108 - (cset, cset_pat, do_typedef)
1.109 - end handle ERROR => err_in_typedef name;
1.110 + val _ = if null errs then () else error (cat_lines errs);
1.111 +
1.112 + (*test theory errors now!*)
1.113 + val test_thy = Theory.copy thy;
1.114 + val test_sign = Theory.sign_of test_thy;
1.115 + val _ = (test_thy, Thm.assume (Thm.cterm_of test_sign goal)) |> typedef_att;
1.116 +
1.117 + in (cset, goal, goal_pat, typedef_att) end
1.118 + handle ERROR => err_in_typedef name;
1.119
1.120
1.121 (* add_typedef interfaces *)
1.122
1.123 fun gen_add_typedef prep_term no_def name typ set names thms tac thy =
1.124 let
1.125 - val (cset, _, do_typedef) = prepare_typedef prep_term no_def name typ set thy;
1.126 - val result = prove_nonempty thy cset (names, thms, tac);
1.127 - in check_nonempty cset result; thy |> do_typedef thy |> #1 end;
1.128 + val (cset, goal, _, typedef_att) = prepare_typedef prep_term no_def name typ set thy;
1.129 + val result = prove_nonempty thy cset goal (names, thms, tac);
1.130 + in (thy, result) |> typedef_att |> #1 end;
1.131
1.132 val add_typedef = gen_add_typedef read_term false;
1.133 val add_typedef_i = gen_add_typedef cert_term false;
1.134 @@ -228,20 +215,9 @@
1.135
1.136 (* typedef_proof interface *)
1.137
1.138 -fun typedef_attribute cset do_typedef (thy, thm) =
1.139 - (check_nonempty cset (thm RS (some_eq_ex RS iffD2)); (thy |> do_typedef));
1.140 -
1.141 fun gen_typedef_proof prep_term ((name, typ, set), comment) int thy =
1.142 - let
1.143 - val (cset, cset_pat, do_typedef) = prepare_typedef prep_term false name typ set thy;
1.144 - val goal = Thm.term_of (goal_nonempty true cset);
1.145 - val goal_pat = Thm.term_of (goal_nonempty true cset_pat);
1.146 - val test_thy = Theory.copy thy;
1.147 - in
1.148 - test_thy |> do_typedef test_thy; (*preview errors!*)
1.149 - thy |> IsarThy.theorem_i ((("", [typedef_attribute cset (do_typedef thy)]),
1.150 - (goal, ([goal_pat], []))), comment) int
1.151 - end;
1.152 + let val (_, goal, goal_pat, att) = prepare_typedef prep_term false name typ set thy;
1.153 + in thy |> IsarThy.theorem_i ((("", [att]), (goal, ([goal_pat], []))), comment) int end;
1.154
1.155 val typedef_proof = gen_typedef_proof read_term;
1.156 val typedef_proof_i = gen_typedef_proof cert_term;