diff --git a/src/HOL/Proofs/ex/XML_Data.thy b/src/HOL/Proofs/ex/XML_Data.thy --- a/src/HOL/Proofs/ex/XML_Data.thy +++ b/src/HOL/Proofs/ex/XML_Data.thy @@ -1,58 +1,58 @@ (* Title: HOL/Proofs/ex/XML_Data.thy Author: Makarius Author: Stefan Berghofer XML data representation of proof terms. *) theory XML_Data imports "HOL-Isar_Examples.Drinker" begin subsection \Export and re-import of global proof terms\ ML \ fun export_proof thy thm = thm - |> Thm.standard_proof_of {full = true, expand_name = Thm.expand_name thm} + |> Proof_Syntax.standard_proof_of {full = true, expand_name = Thm.expand_name thm} |> Proofterm.encode (Sign.consts_of thy); fun import_proof thy xml = let val prf = Proofterm.decode (Sign.consts_of thy) xml; val (prf', _) = Proofterm.freeze_thaw_prf prf; in Drule.export_without_context (Proof_Checker.thm_of_proof thy prf') end; \ subsection \Examples\ ML \val thy1 = \<^theory>\ lemma ex: "A \ A" .. ML_val \ val xml = export_proof thy1 @{thm ex}; val thm = import_proof thy1 xml; \ ML_val \ val xml = export_proof thy1 @{thm de_Morgan}; val thm = import_proof thy1 xml; \ ML_val \ val xml = export_proof thy1 @{thm Drinker's_Principle}; val thm = import_proof thy1 xml; \ text \Some fairly large proof:\ ML_val \ val xml = export_proof thy1 @{thm abs_less_iff}; val thm = import_proof thy1 xml; val xml_size = size (YXML.string_of_body xml); \<^assert> (xml_size > 400000); \ end diff --git a/src/Pure/Proof/proof_rewrite_rules.ML b/src/Pure/Proof/proof_rewrite_rules.ML --- a/src/Pure/Proof/proof_rewrite_rules.ML +++ b/src/Pure/Proof/proof_rewrite_rules.ML @@ -1,384 +1,386 @@ (* Title: Pure/Proof/proof_rewrite_rules.ML Author: Stefan Berghofer, TU Muenchen Simplification functions for proof terms involving meta level rules. *) signature PROOF_REWRITE_RULES = sig val rew : bool -> typ list -> term option list -> Proofterm.proof -> (Proofterm.proof * Proofterm.proof) option val rprocs : bool -> (typ list -> term option list -> Proofterm.proof -> (Proofterm.proof * Proofterm.proof) option) list val rewrite_terms : (term -> term) -> Proofterm.proof -> Proofterm.proof val elim_defs : theory -> bool -> thm list -> Proofterm.proof -> Proofterm.proof val elim_vars : (typ -> term) -> Proofterm.proof -> Proofterm.proof val hhf_proof : term -> term -> Proofterm.proof -> Proofterm.proof val un_hhf_proof : term -> term -> Proofterm.proof -> Proofterm.proof val expand_of_sort_proof : theory -> term option list -> typ * sort -> Proofterm.proof list val expand_of_class_proof : theory -> term option list -> typ * class -> proof val expand_of_class : theory -> typ list -> term option list -> Proofterm.proof -> (Proofterm.proof * Proofterm.proof) option val standard_preproc: (proof * proof) list -> theory -> proof -> proof end; structure Proof_Rewrite_Rules : PROOF_REWRITE_RULES = struct fun rew b _ _ = let fun ?? x = if b then SOME x else NONE; fun ax (prf as PAxm (s, prop, _)) Ts = if b then PAxm (s, prop, SOME Ts) else prf; - fun ty T = if b then - let val Type (_, [Type (_, [U, _]), _]) = T - in SOME U end + fun ty T = + if b then + (case T of + Type (_, [Type (_, [U, _]), _]) => SOME U + | _ => NONE) else NONE; val equal_intr_axm = ax Proofterm.equal_intr_axm []; val equal_elim_axm = ax Proofterm.equal_elim_axm []; val symmetric_axm = ax Proofterm.symmetric_axm [propT]; fun rew' (PThm ({name = "Pure.protectD", ...}, _) % _ %% (PThm ({name = "Pure.protectI", ...}, _) % _ %% prf)) = SOME prf | rew' (PThm ({name = "Pure.conjunctionD1", ...}, _) % _ % _ %% (PThm ({name = "Pure.conjunctionI", ...}, _) % _ % _ %% prf %% _)) = SOME prf | rew' (PThm ({name = "Pure.conjunctionD2", ...}, _) % _ % _ %% (PThm ({name = "Pure.conjunctionI", ...}, _) % _ % _ %% _ %% prf)) = SOME prf | rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %% (PAxm ("Pure.equal_intr", _, _) % _ % _ %% prf %% _)) = SOME prf | rew' (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.equal_intr", _, _) % A % B %% prf1 %% prf2)) = SOME (equal_intr_axm % B % A %% prf2 %% prf1) | rew' (PAxm ("Pure.equal_elim", _, _) % SOME (_ $ A) % SOME (_ $ B) %% (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.prop", _)) % _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% ((tg as PThm ({name = "Pure.protectI", ...}, _)) % _ %% prf2)) = SOME (tg %> B %% (equal_elim_axm %> A %> B %% prf1 %% prf2)) | rew' (PAxm ("Pure.equal_elim", _, _) % SOME (_ $ A) % SOME (_ $ B) %% (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.prop", _)) % _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% prf1)) %% ((tg as PThm ({name = "Pure.protectI", ...}, _)) % _ %% prf2)) = SOME (tg %> B %% (equal_elim_axm %> A %> B %% (symmetric_axm % ?? B % ?? A %% prf1) %% prf2)) | rew' (PAxm ("Pure.equal_elim", _, _) % SOME X % SOME Y %% (PAxm ("Pure.combination", _, _) % _ % _ % _ % _ %% (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.imp", _)) % _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% prf2)) = let val _ $ A $ C = Envir.beta_norm X; val _ $ B $ D = Envir.beta_norm Y in SOME (AbsP ("H1", ?? X, AbsP ("H2", ?? B, Proofterm.equal_elim_axm %> C %> D %% Proofterm.incr_pboundvars 2 0 prf2 %% (PBound 1 %% (equal_elim_axm %> B %> A %% (Proofterm.symmetric_axm % ?? A % ?? B %% Proofterm.incr_pboundvars 2 0 prf1) %% PBound 0))))) end | rew' (PAxm ("Pure.equal_elim", _, _) % SOME X % SOME Y %% (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.combination", _, _) % _ % _ % _ % _ %% (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.imp", _)) % _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% prf2))) = let val _ $ A $ C = Envir.beta_norm Y; val _ $ B $ D = Envir.beta_norm X in SOME (AbsP ("H1", ?? X, AbsP ("H2", ?? A, equal_elim_axm %> D %> C %% (symmetric_axm % ?? C % ?? D %% Proofterm.incr_pboundvars 2 0 prf2) %% (PBound 1 %% (equal_elim_axm %> A %> B %% Proofterm.incr_pboundvars 2 0 prf1 %% PBound 0))))) end | rew' (PAxm ("Pure.equal_elim", _, _) % SOME X % SOME Y %% (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.all", _)) % _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% (PAxm ("Pure.abstract_rule", _, _) % _ % _ %% prf))) = let val Const (_, T) $ P = Envir.beta_norm X; val _ $ Q = Envir.beta_norm Y; in SOME (AbsP ("H", ?? X, Abst ("x", ty T, equal_elim_axm %> incr_boundvars 1 P $ Bound 0 %> incr_boundvars 1 Q $ Bound 0 %% (Proofterm.incr_pboundvars 1 1 prf %> Bound 0) %% (PBound 0 %> Bound 0)))) end | rew' (PAxm ("Pure.equal_elim", _, _) % SOME X % SOME Y %% (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.all", _)) % _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% (PAxm ("Pure.abstract_rule", _, _) % _ % _ %% prf)))) = let val Const (_, T) $ P = Envir.beta_norm X; val _ $ Q = Envir.beta_norm Y; val t = incr_boundvars 1 P $ Bound 0; val u = incr_boundvars 1 Q $ Bound 0 in SOME (AbsP ("H", ?? X, Abst ("x", ty T, equal_elim_axm %> t %> u %% (symmetric_axm % ?? u % ?? t %% (Proofterm.incr_pboundvars 1 1 prf %> Bound 0)) %% (PBound 0 %> Bound 0)))) end | rew' (PAxm ("Pure.equal_elim", _, _) % SOME A % SOME C %% (PAxm ("Pure.transitive", _, _) % _ % SOME B % _ %% prf1 %% prf2) %% prf3) = SOME (equal_elim_axm %> B %> C %% prf2 %% (equal_elim_axm %> A %> B %% prf1 %% prf3)) | rew' (PAxm ("Pure.equal_elim", _, _) % SOME A % SOME C %% (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.transitive", _, _) % _ % SOME B % _ %% prf1 %% prf2)) %% prf3) = SOME (equal_elim_axm %> B %> C %% (symmetric_axm % ?? C % ?? B %% prf1) %% (equal_elim_axm %> A %> B %% (symmetric_axm % ?? B % ?? A %% prf2) %% prf3)) | rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% prf) = SOME prf | rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %% (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _)) %% prf) = SOME prf | rew' (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.symmetric", _, _) % _ % _ %% prf)) = SOME prf | rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %% (PAxm ("Pure.equal_elim", _, _) % SOME (_ $ A $ C) % SOME (_ $ B $ D) %% (PAxm ("Pure.combination", _, _) % _ % _ % _ % _ %% (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.eq", _)) % _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% prf2) %% prf3) %% prf4) = SOME (equal_elim_axm %> C %> D %% prf2 %% (equal_elim_axm %> A %> C %% prf3 %% (equal_elim_axm %> B %> A %% (symmetric_axm % ?? A % ?? B %% prf1) %% prf4))) | rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %% (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.equal_elim", _, _) % SOME (_ $ A $ C) % SOME (_ $ B $ D) %% (PAxm ("Pure.combination", _, _) % _ % _ % _ % _ %% (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.eq", _)) % _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% prf2) %% prf3)) %% prf4) = SOME (equal_elim_axm %> A %> B %% prf1 %% (equal_elim_axm %> C %> A %% (symmetric_axm % ?? A % ?? C %% prf3) %% (equal_elim_axm %> D %> C %% (symmetric_axm % ?? C % ?? D %% prf2) %% prf4))) | rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %% (PAxm ("Pure.equal_elim", _, _) % SOME (_ $ B $ D) % SOME (_ $ A $ C) %% (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.combination", _, _) % _ % _ % _ % _ %% (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.eq", _)) % _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% prf2)) %% prf3) %% prf4) = SOME (equal_elim_axm %> D %> C %% (symmetric_axm % ?? C % ?? D %% prf2) %% (equal_elim_axm %> B %> D %% prf3 %% (equal_elim_axm %> A %> B %% prf1 %% prf4))) | rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %% (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.equal_elim", _, _) % SOME (_ $ B $ D) % SOME (_ $ A $ C) %% (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.combination", _, _) % _ % _ % _ % _ %% (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.eq", _)) % _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% prf2)) %% prf3)) %% prf4) = SOME (equal_elim_axm %> B %> A %% (symmetric_axm % ?? A % ?? B %% prf1) %% (equal_elim_axm %> D %> B %% (symmetric_axm % ?? B % ?? D %% prf3) %% (equal_elim_axm %> C %> D %% prf2 %% prf4))) | rew' ((prf as PAxm ("Pure.combination", _, _) % SOME ((eq as Const ("Pure.eq", T)) $ t) % _ % _ % _) %% (PAxm ("Pure.reflexive", _, _) % _)) = let val (U, V) = (case T of Type (_, [U, V]) => (U, V) | _ => (dummyT, dummyT)) in SOME (prf %% (ax Proofterm.combination_axm [U, V] %> eq % ?? eq % ?? t % ?? t %% (ax Proofterm.reflexive_axm [T] % ?? eq) %% (ax Proofterm.reflexive_axm [U] % ?? t))) end | rew' _ = NONE; in rew' #> Option.map (rpair Proofterm.no_skel) end; fun rprocs b = [rew b]; val _ = Theory.setup (fold Proofterm.add_prf_rproc (rprocs false)); (**** apply rewriting function to all terms in proof ****) fun rewrite_terms r = let fun rew_term Ts t = let val frees = map Free (Name.invent (Term.declare_term_frees t Name.context) "xa" (length Ts) ~~ Ts); val t' = r (subst_bounds (frees, t)); fun strip [] t = t | strip (_ :: xs) (Abs (_, _, t)) = strip xs t; in strip Ts (fold lambda frees t') end; fun rew Ts (prf1 %% prf2) = rew Ts prf1 %% rew Ts prf2 | rew Ts (prf % SOME t) = rew Ts prf % SOME (rew_term Ts t) | rew Ts (Abst (s, SOME T, prf)) = Abst (s, SOME T, rew (T :: Ts) prf) | rew Ts (AbsP (s, SOME t, prf)) = AbsP (s, SOME (rew_term Ts t), rew Ts prf) | rew _ prf = prf in rew [] end; (**** eliminate definitions in proof ****) fun vars_of t = rev (fold_aterms (fn v as Var _ => insert (op =) v | _ => I) t []); fun insert_refl defs Ts (prf1 %% prf2) = let val (prf1', b) = insert_refl defs Ts prf1 in if b then (prf1', true) else (prf1' %% fst (insert_refl defs Ts prf2), false) end | insert_refl defs Ts (Abst (s, SOME T, prf)) = (Abst (s, SOME T, fst (insert_refl defs (T :: Ts) prf)), false) | insert_refl defs Ts (AbsP (s, t, prf)) = (AbsP (s, t, fst (insert_refl defs Ts prf)), false) | insert_refl defs Ts prf = (case Proofterm.strip_combt prf of (PThm ({name = s, prop, types = SOME Ts, ...}, _), ts) => if member (op =) defs s then let val vs = vars_of prop; val tvars = Term.add_tvars prop [] |> rev; val (_, rhs) = Logic.dest_equals (Logic.strip_imp_concl prop); val rhs' = Term.betapplys (subst_TVars (map fst tvars ~~ Ts) (fold_rev (fn x => fn b => Abs ("", dummyT, abstract_over (x, b))) vs rhs), map the ts); in (Proofterm.change_types (SOME [fastype_of1 (Ts, rhs')]) Proofterm.reflexive_axm %> rhs', true) end else (prf, false) | (_, []) => (prf, false) | (prf', ts) => (Proofterm.proof_combt' (fst (insert_refl defs Ts prf'), ts), false)); fun elim_defs thy r defs prf = let val defs' = map (Logic.dest_equals o map_types Type.strip_sorts o Thm.prop_of o Drule.abs_def) defs; val defnames = map Thm.derivation_name defs; val prf' = if r then let val cnames = map (fst o dest_Const o fst) defs'; val expand = Proofterm.fold_proof_atoms true (fn PThm ({serial, name, prop, ...}, _) => if member (op =) defnames name orelse not (exists_Const (member (op =) cnames o #1) prop) then I else Inttab.update (serial, "") | _ => I) [prf] Inttab.empty; in Proofterm.expand_proof thy (Inttab.lookup expand o #serial) prf end else prf; in rewrite_terms (Pattern.rewrite_term thy defs' []) (fst (insert_refl defnames [] prf')) end; (**** eliminate all variables that don't occur in the proposition ****) fun elim_vars mk_default prf = let val prop = Proofterm.prop_of prf; val tv = Term.add_vars prop []; val tf = Term.add_frees prop []; fun hidden_variable (Var v) = not (member (op =) tv v) | hidden_variable (Free f) = not (member (op =) tf f) | hidden_variable _ = false; fun mk_default' T = fold_rev (Term.abs o pair "x") (binder_types T) (mk_default (body_type T)); fun elim_varst (t $ u) = elim_varst t $ elim_varst u | elim_varst (Abs (s, T, t)) = Abs (s, T, elim_varst t) | elim_varst (t as Free (x, T)) = if member (op =) tf (x, T) then t else mk_default' T | elim_varst (t as Var (xi, T)) = if member (op =) tv (xi, T) then t else mk_default' T | elim_varst t = t; in Proofterm.map_proof_terms (fn t => if Term.exists_subterm hidden_variable t then Envir.beta_norm (elim_varst t) else t) I prf end; (**** convert between hhf and non-hhf form ****) fun hhf_proof P Q prf = let val params = Logic.strip_params Q; val Hs = Logic.strip_assums_hyp P; val Hs' = Logic.strip_assums_hyp Q; val k = length Hs; val l = length params; fun mk_prf i j Hs Hs' (Const ("Pure.all", _) $ Abs (_, _, P)) prf = mk_prf i (j - 1) Hs Hs' P (prf %> Bound j) | mk_prf i j (H :: Hs) (H' :: Hs') (Const ("Pure.imp", _) $ _ $ P) prf = mk_prf (i - 1) j Hs Hs' P (prf %% un_hhf_proof H' H (PBound i)) | mk_prf _ _ _ _ _ prf = prf in prf |> Proofterm.incr_pboundvars k l |> mk_prf (k - 1) (l - 1) Hs Hs' P |> fold_rev (fn P => fn prf => AbsP ("H", SOME P, prf)) Hs' |> fold_rev (fn (s, T) => fn prf => Abst (s, SOME T, prf)) params end and un_hhf_proof P Q prf = let val params = Logic.strip_params Q; val Hs = Logic.strip_assums_hyp P; val Hs' = Logic.strip_assums_hyp Q; val k = length Hs; val l = length params; fun mk_prf (Const ("Pure.all", _) $ Abs (s, T, P)) prf = Abst (s, SOME T, mk_prf P prf) | mk_prf (Const ("Pure.imp", _) $ P $ Q) prf = AbsP ("H", SOME P, mk_prf Q prf) | mk_prf _ prf = prf in prf |> Proofterm.incr_pboundvars k l |> fold (fn i => fn prf => prf %> Bound i) (l - 1 downto 0) |> fold (fn ((H, H'), i) => fn prf => prf %% hhf_proof H' H (PBound i)) (Hs ~~ Hs' ~~ (k - 1 downto 0)) |> mk_prf Q end; (**** expand OfClass proofs ****) fun expand_of_sort_proof thy hyps (T, S) = let val of_class_hyps = map (fn SOME t => try Logic.dest_of_class t | NONE => NONE) hyps; fun of_class_index p = (case find_index (fn SOME h => h = p | NONE => false) of_class_hyps of ~1 => raise Fail "expand_of_class_proof: missing class hypothesis" | i => PBound i); val sorts = AList.coalesce (op =) (rev (map_filter I of_class_hyps)); fun get_sort T = the_default [] (AList.lookup (op =) sorts T); val subst = map_atyps (fn T as TVar (ai, _) => TVar (ai, get_sort T) | T as TFree (a, _) => TFree (a, get_sort T)); fun reconstruct prop = Proofterm.reconstruct_proof thy prop #> Proofterm.expand_proof thy Proofterm.expand_name_empty #> Same.commit (Proofterm.map_proof_same Same.same Same.same of_class_index); in map2 reconstruct (Logic.mk_of_sort (T, S)) (Proofterm.of_sort_proof (Sign.classes_of thy) (Thm.classrel_proof thy) (Thm.arity_proof thy) (OfClass o apfst Type.strip_sorts) (subst T, S)) end; fun expand_of_class_proof thy hyps (T, c) = hd (expand_of_sort_proof thy hyps (T, [c])); fun expand_of_class thy (_: typ list) hyps (OfClass (T, c)) = SOME (expand_of_class_proof thy hyps (T, c), Proofterm.no_skel) | expand_of_class _ _ _ _ = NONE; (* standard preprocessor *) fun standard_preproc rews thy = Proofterm.rewrite_proof thy (rews, rprocs true @ [expand_of_class thy]); end; diff --git a/src/Pure/Proof/proof_syntax.ML b/src/Pure/Proof/proof_syntax.ML --- a/src/Pure/Proof/proof_syntax.ML +++ b/src/Pure/Proof/proof_syntax.ML @@ -1,283 +1,297 @@ (* Title: Pure/Proof/proof_syntax.ML Author: Stefan Berghofer, TU Muenchen Function for parsing and printing proof terms. *) signature PROOF_SYNTAX = sig val add_proof_syntax: theory -> theory val term_of_proof: proof -> term val proof_of_term: theory -> bool -> term -> Proofterm.proof val read_term: theory -> bool -> typ -> string -> term val read_proof: theory -> bool -> bool -> string -> Proofterm.proof val proof_syntax: Proofterm.proof -> theory -> theory val proof_of: bool -> thm -> Proofterm.proof val pretty_proof: Proof.context -> Proofterm.proof -> Pretty.T - val pretty_standard_proof_of: Proof.context -> bool -> thm -> Pretty.T val pretty_proof_boxes_of: Proof.context -> {full: bool, preproc: theory -> proof -> proof} -> thm -> Pretty.T + val standard_proof_of: {full: bool, expand_name: Proofterm.thm_header -> string option} -> + thm -> Proofterm.proof + val pretty_standard_proof_of: Proof.context -> bool -> thm -> Pretty.T end; structure Proof_Syntax : PROOF_SYNTAX = struct (**** add special syntax for embedding proof terms ****) val proofT = Type ("Pure.proof", []); local val paramT = Type ("param", []); val paramsT = Type ("params", []); val idtT = Type ("idt", []); val aT = Term.aT []; fun mixfix (sy, ps, p) = Mixfix (Input.string sy, ps, p, Position.no_range); in fun add_proof_syntax thy = thy |> Sign.root_path |> Sign.set_defsort [] |> Sign.add_nonterminals_global [Binding.make ("param", \<^here>), Binding.make ("params", \<^here>)] |> Sign.add_syntax Syntax.mode_default [("_Lam", [paramsT, proofT] ---> proofT, mixfix ("(1\<^bold>\_./ _)", [0, 3], 3)), ("_Lam0", [paramT, paramsT] ---> paramsT, mixfix ("_/ _", [1, 0], 0)), ("_Lam0", [idtT, paramsT] ---> paramsT, mixfix ("_/ _", [1, 0], 0)), ("_Lam1", [idtT, propT] ---> paramT, mixfix ("_: _", [0, 0], 0)), ("", paramT --> paramT, Mixfix.mixfix "'(_')"), ("", idtT --> paramsT, Mixfix.mixfix "_"), ("", paramT --> paramsT, Mixfix.mixfix "_"), (Lexicon.mark_const "Pure.Appt", [proofT, aT] ---> proofT, mixfix ("(1_ \/ _)", [4, 5], 4)), (Lexicon.mark_const "Pure.AppP", [proofT, proofT] ---> proofT, mixfix ("(1_ \/ _)", [4, 5], 4)), (Lexicon.mark_const "Pure.MinProof", proofT, Mixfix.mixfix "\<^bold>?")] |> Sign.add_trrules (map Syntax.Parse_Print_Rule [(Ast.mk_appl (Ast.Constant "_Lam") [Ast.mk_appl (Ast.Constant "_Lam0") [Ast.Variable "l", Ast.Variable "m"], Ast.Variable "A"], Ast.mk_appl (Ast.Constant "_Lam") [Ast.Variable "l", Ast.mk_appl (Ast.Constant "_Lam") [Ast.Variable "m", Ast.Variable "A"]]), (Ast.mk_appl (Ast.Constant "_Lam") [Ast.mk_appl (Ast.Constant "_Lam1") [Ast.Variable "x", Ast.Variable "A"], Ast.Variable "B"], Ast.mk_appl (Ast.Constant (Lexicon.mark_const "Pure.AbsP")) [Ast.Variable "A", (Ast.mk_appl (Ast.Constant "_abs") [Ast.Variable "x", Ast.Variable "B"])]), (Ast.mk_appl (Ast.Constant "_Lam") [Ast.Variable "x", Ast.Variable "A"], Ast.mk_appl (Ast.Constant (Lexicon.mark_const "Pure.Abst")) [(Ast.mk_appl (Ast.Constant "_abs") [Ast.Variable "x", Ast.Variable "A"])])]); end; (** constants for theorems and axioms **) fun add_proof_atom_consts names thy = thy |> Sign.root_path |> Sign.add_consts (map (fn name => (Binding.qualified_name name, proofT, NoSyn)) names); (** proof terms as pure terms **) (* term_of_proof *) local val AbsPt = Const ("Pure.AbsP", propT --> (proofT --> proofT) --> proofT); val AppPt = Const ("Pure.AppP", proofT --> proofT --> proofT); val Hypt = Const ("Pure.Hyp", propT --> proofT); val Oraclet = Const ("Pure.Oracle", propT --> proofT); val MinProoft = Const ("Pure.MinProof", proofT); fun AppT T prf = Const ("Pure.Appt", proofT --> Term.itselfT T --> proofT) $ prf $ Logic.mk_type T; fun OfClasst (T, c) = let val U = Term.itselfT T --> propT in Const ("Pure.OfClass", U --> proofT) $ Const (Logic.const_of_class c, U) end; fun term_of _ (PThm ({serial = i, name, types = Ts, ...}, _)) = fold AppT (these Ts) (Const (Long_Name.append "thm" (if name = "" then string_of_int i else name), proofT)) | term_of _ (PAxm (name, _, Ts)) = fold AppT (these Ts) (Const (Long_Name.append "axm" name, proofT)) | term_of _ (OfClass (T, c)) = AppT T (OfClasst (T, c)) | term_of _ (PBound i) = Bound i | term_of Ts (Abst (s, opT, prf)) = let val T = the_default dummyT opT in Const ("Pure.Abst", (T --> proofT) --> proofT) $ Abs (s, T, term_of (T::Ts) (Proofterm.incr_pboundvars 1 0 prf)) end | term_of Ts (AbsP (s, t, prf)) = AbsPt $ the_default Term.dummy_prop t $ Abs (s, proofT, term_of (proofT::Ts) (Proofterm.incr_pboundvars 0 1 prf)) | term_of Ts (prf1 %% prf2) = AppPt $ term_of Ts prf1 $ term_of Ts prf2 | term_of Ts (prf % opt) = let val t = the_default Term.dummy opt; val T = fastype_of1 (Ts, t) handle TERM _ => dummyT; in Const ("Pure.Appt", proofT --> T --> proofT) $ term_of Ts prf $ t end | term_of _ (Hyp t) = Hypt $ t | term_of _ (Oracle (_, t, _)) = Oraclet $ t | term_of _ MinProof = MinProoft; in val term_of_proof = term_of []; end; (* proof_of_term *) fun proof_of_term thy ty = let val thms = Global_Theory.all_thms_of thy true; val axms = Theory.all_axioms_of thy; fun mk_term t = (if ty then I else map_types (K dummyT)) (Term.no_dummy_patterns t); fun prf_of [] (Bound i) = PBound i | prf_of Ts (Const (s, Type ("Pure.proof", _))) = Proofterm.change_types (if ty then SOME Ts else NONE) (case Long_Name.explode s of "axm" :: xs => let val name = Long_Name.implode xs; val prop = (case AList.lookup (op =) axms name of SOME prop => prop | NONE => error ("Unknown axiom " ^ quote name)) in PAxm (name, prop, NONE) end | "thm" :: xs => let val name = Long_Name.implode xs; in (case AList.lookup (op =) thms name of SOME thm => fst (Proofterm.strip_combt (fst (Proofterm.strip_combP (Thm.proof_of thm)))) | NONE => error ("Unknown theorem " ^ quote name)) end | _ => error ("Illegal proof constant name: " ^ quote s)) | prf_of Ts (Const ("Pure.OfClass", _) $ Const (c_class, _)) = (case try Logic.class_of_const c_class of SOME c => Proofterm.change_types (if ty then SOME Ts else NONE) (OfClass (TVar ((Name.aT, 0), []), c)) | NONE => error ("Bad class constant: " ^ quote c_class)) | prf_of Ts (Const ("Pure.Hyp", _) $ prop) = Hyp prop | prf_of Ts (v as Var ((_, Type ("Pure.proof", _)))) = Hyp v | prf_of [] (Const ("Pure.Abst", _) $ Abs (s, T, prf)) = if T = proofT then error ("Term variable abstraction may not bind proof variable " ^ quote s) else Abst (s, if ty then SOME T else NONE, Proofterm.incr_pboundvars (~1) 0 (prf_of [] prf)) | prf_of [] (Const ("Pure.AbsP", _) $ t $ Abs (s, _, prf)) = AbsP (s, case t of Const ("Pure.dummy_pattern", _) => NONE | _ $ Const ("Pure.dummy_pattern", _) => NONE | _ => SOME (mk_term t), Proofterm.incr_pboundvars 0 (~1) (prf_of [] prf)) | prf_of [] (Const ("Pure.AppP", _) $ prf1 $ prf2) = prf_of [] prf1 %% prf_of [] prf2 | prf_of Ts (Const ("Pure.Appt", _) $ prf $ Const ("Pure.type", Type ("itself", [T]))) = prf_of (T::Ts) prf | prf_of [] (Const ("Pure.Appt", _) $ prf $ t) = prf_of [] prf % (case t of Const ("Pure.dummy_pattern", _) => NONE | _ => SOME (mk_term t)) | prf_of _ t = error ("Not a proof term:\n" ^ Syntax.string_of_term_global thy t) in prf_of [] end; fun read_term thy topsort = let val thm_names = filter_out (fn s => s = "") (map fst (Global_Theory.all_thms_of thy true)); val axm_names = map fst (Theory.all_axioms_of thy); val ctxt = thy |> add_proof_syntax |> add_proof_atom_consts (map (Long_Name.append "axm") axm_names @ map (Long_Name.append "thm") thm_names) |> Proof_Context.init_global |> Proof_Context.allow_dummies |> Proof_Context.set_mode Proof_Context.mode_schematic |> topsort ? (Proof_Context.set_defsort [] #> Config.put Type_Infer.object_logic false #> Config.put Type_Infer_Context.const_sorts false); in fn ty => fn s => (if ty = propT then Syntax.parse_prop else Syntax.parse_term) ctxt s |> Type.constraint ty |> Syntax.check_term ctxt end; fun read_proof thy topsort = let val rd = read_term thy topsort proofT in fn ty => fn s => proof_of_term thy ty (Logic.varify_global (rd s)) end; fun proof_syntax prf = let val thm_names = Symtab.keys (Proofterm.fold_proof_atoms true (fn PThm ({name, ...}, _) => if name <> "" then Symtab.update (name, ()) else I | _ => I) [prf] Symtab.empty); val axm_names = Symtab.keys (Proofterm.fold_proof_atoms true (fn PAxm (name, _, _) => Symtab.update (name, ()) | _ => I) [prf] Symtab.empty); in add_proof_syntax #> add_proof_atom_consts (map (Long_Name.append "thm") thm_names @ map (Long_Name.append "axm") axm_names) end; fun proof_of full thm = let val thy = Thm.theory_of_thm thm; val prop = Thm.full_prop_of thm; val prf = Thm.proof_of thm; in (case fst (Proofterm.strip_combt (fst (Proofterm.strip_combP prf))) of PThm ({prop = prop', ...}, thm_body) => if prop = prop' then Proofterm.thm_body_proof_raw thm_body else prf | _ => prf) |> full ? Proofterm.reconstruct_proof thy prop end; fun pretty_proof ctxt prf = Proof_Context.pretty_term_abbrev (Proof_Context.transfer (proof_syntax prf (Proof_Context.theory_of ctxt)) ctxt) (term_of_proof prf); -fun pretty_standard_proof_of ctxt full thm = - pretty_proof ctxt (Thm.standard_proof_of {full = full, expand_name = Thm.expand_name thm} thm); - fun pretty_proof_boxes_of ctxt {full, preproc} thm = let val thy = Proof_Context.theory_of ctxt; val selection = {included = Proofterm.this_id (Thm.derivation_id thm), excluded = is_some o Global_Theory.lookup_thm_id thy} in Proofterm.proof_boxes selection [Thm.proof_of thm] |> map (fn ({serial = i, pos, prop, ...}, proof) => let val proof' = proof |> Proofterm.reconstruct_proof thy prop |> preproc thy |> not full ? Proofterm.shrink_proof |> Proofterm.forall_intr_variables prop; val prop' = prop |> Proofterm.forall_intr_variables_term; val name = Long_Name.append "thm" (string_of_int i); in Pretty.item [Pretty.str (name ^ Position.here_list pos ^ ":"), Pretty.brk 1, Syntax.pretty_term ctxt prop', Pretty.fbrk, pretty_proof ctxt proof'] end) |> Pretty.chunks end; + +(* standardized proofs *) + +fun standard_proof_of {full, expand_name} thm = + let val thy = Thm.theory_of_thm thm in + Thm.reconstruct_proof_of thm + |> Proofterm.expand_proof thy expand_name + |> Proofterm.rewrite_proof thy ([], Proof_Rewrite_Rules.rprocs true) + |> Proofterm.no_thm_proofs + |> not full ? Proofterm.shrink_proof + end; + +fun pretty_standard_proof_of ctxt full thm = + pretty_proof ctxt (standard_proof_of {full = full, expand_name = Thm.expand_name thm} thm); + end; diff --git a/src/Pure/ROOT.ML b/src/Pure/ROOT.ML --- a/src/Pure/ROOT.ML +++ b/src/Pure/ROOT.ML @@ -1,355 +1,355 @@ (* Title: Pure/ROOT.ML Author: Makarius Main entry point for the Isabelle/Pure bootstrap process. Note: When this file is open in the Prover IDE, the ML files of Isabelle/Pure can be explored interactively. This is a separate copy of Pure within Pure: it does not affect the running logic session. *) chapter "Isabelle/Pure bootstrap"; ML_file "ML/ml_name_space.ML"; section "Bootstrap phase 0: Poly/ML setup"; ML_file "ML/ml_init.ML"; ML_file "ML/ml_system.ML"; ML_file "System/distribution.ML"; ML_file "General/basics.ML"; ML_file "General/symbol_explode.ML"; ML_file "Concurrent/multithreading.ML"; ML_file "Concurrent/unsynchronized.ML"; ML_file "Concurrent/synchronized.ML"; ML_file "Concurrent/counter.ML"; ML_file "ML/ml_heap.ML"; ML_file "ML/ml_profiling.ML"; ML_file "ML/ml_print_depth0.ML"; ML_file "ML/ml_pretty.ML"; ML_file "ML/ml_compiler0.ML"; section "Bootstrap phase 1: towards ML within position context"; subsection "Library of general tools"; ML_file "library.ML"; ML_file "General/print_mode.ML"; ML_file "General/alist.ML"; ML_file "General/table.ML"; ML_file "General/random.ML"; ML_file "General/value.ML"; ML_file "General/properties.ML"; ML_file "General/output.ML"; ML_file "PIDE/markup.ML"; ML_file "General/utf8.ML"; ML_file "General/scan.ML"; ML_file "General/source.ML"; ML_file "General/symbol.ML"; ML_file "General/position.ML"; ML_file "General/symbol_pos.ML"; ML_file "General/input.ML"; ML_file "General/comment.ML"; ML_file "General/antiquote.ML"; ML_file "ML/ml_lex.ML"; ML_file "ML/ml_compiler1.ML"; section "Bootstrap phase 2: towards ML within Isar context"; subsection "Library of general tools"; ML_file "General/integer.ML"; ML_file "General/stack.ML"; ML_file "General/queue.ML"; ML_file "General/heap.ML"; ML_file "General/same.ML"; ML_file "General/ord_list.ML"; ML_file "General/balanced_tree.ML"; ML_file "General/linear_set.ML"; ML_file "General/buffer.ML"; ML_file "General/pretty.ML"; ML_file "General/rat.ML"; ML_file "PIDE/xml.ML"; ML_file "General/path.ML"; ML_file "General/url.ML"; ML_file "System/bash_syntax.ML"; ML_file "General/file.ML"; ML_file "General/long_name.ML"; ML_file "General/binding.ML"; ML_file "General/socket_io.ML"; ML_file "General/seq.ML"; ML_file "General/timing.ML"; ML_file "General/sha1.ML"; ML_file "PIDE/byte_message.ML"; ML_file "PIDE/yxml.ML"; ML_file "PIDE/protocol_message.ML"; ML_file "PIDE/document_id.ML"; ML_file "General/change_table.ML"; ML_file "General/graph.ML"; subsection "Fundamental structures"; ML_file "name.ML"; ML_file "term.ML"; ML_file "context.ML"; ML_file "context_position.ML"; ML_file "System/options.ML"; ML_file "config.ML"; ML_file "soft_type_system.ML"; subsection "Concurrency within the ML runtime"; ML_file "ML/exn_properties.ML"; +ML_file_no_debug "ML/exn_debugger.ML"; ML_file "ML/ml_statistics.ML"; ML_file "Concurrent/thread_data_virtual.ML"; ML_file "Concurrent/standard_thread.ML"; ML_file "Concurrent/single_assignment.ML"; ML_file "Concurrent/par_exn.ML"; ML_file "Concurrent/task_queue.ML"; ML_file "Concurrent/future.ML"; ML_file "Concurrent/event_timer.ML"; ML_file "Concurrent/timeout.ML"; ML_file "Concurrent/lazy.ML"; ML_file "Concurrent/par_list.ML"; ML_file "Concurrent/mailbox.ML"; ML_file "Concurrent/cache.ML"; ML_file "PIDE/active.ML"; ML_file "Thy/export.ML"; subsection "Inner syntax"; ML_file "Syntax/type_annotation.ML"; ML_file "Syntax/term_position.ML"; ML_file "Syntax/lexicon.ML"; ML_file "Syntax/ast.ML"; ML_file "Syntax/syntax_ext.ML"; ML_file "Syntax/parser.ML"; ML_file "Syntax/syntax_trans.ML"; ML_file "Syntax/mixfix.ML"; ML_file "Syntax/printer.ML"; ML_file "Syntax/syntax.ML"; subsection "Core of tactical proof system"; ML_file "term_ord.ML"; ML_file "term_subst.ML"; ML_file "General/completion.ML"; ML_file "General/name_space.ML"; ML_file "sorts.ML"; ML_file "type.ML"; ML_file "logic.ML"; ML_file "Syntax/simple_syntax.ML"; ML_file "net.ML"; ML_file "item_net.ML"; ML_file "envir.ML"; ML_file "consts.ML"; ML_file "term_xml.ML"; ML_file "primitive_defs.ML"; ML_file "sign.ML"; ML_file "defs.ML"; ML_file "term_sharing.ML"; ML_file "pattern.ML"; ML_file "unify.ML"; ML_file "theory.ML"; ML_file "proofterm.ML"; ML_file "thm.ML"; ML_file "more_pattern.ML"; ML_file "more_unify.ML"; ML_file "more_thm.ML"; ML_file "facts.ML"; ML_file "thm_name.ML"; ML_file "global_theory.ML"; ML_file "pure_thy.ML"; ML_file "drule.ML"; ML_file "morphism.ML"; ML_file "variable.ML"; ML_file "conv.ML"; ML_file "goal_display.ML"; ML_file "tactical.ML"; ML_file "search.ML"; ML_file "tactic.ML"; ML_file "raw_simplifier.ML"; ML_file "conjunction.ML"; ML_file "assumption.ML"; subsection "Isar -- Intelligible Semi-Automated Reasoning"; (*ML support and global execution*) ML_file "ML/ml_syntax.ML"; ML_file "ML/ml_env.ML"; ML_file "ML/ml_options.ML"; ML_file "ML/ml_print_depth.ML"; -ML_file_no_debug "ML/exn_debugger.ML"; ML_file_no_debug "Isar/runtime.ML"; ML_file "PIDE/execution.ML"; ML_file "ML/ml_compiler.ML"; ML_file "skip_proof.ML"; ML_file "goal.ML"; (*outer syntax*) ML_file "Isar/keyword.ML"; ML_file "Isar/token.ML"; ML_file "Isar/parse.ML"; ML_file "Thy/document_source.ML"; ML_file "Thy/thy_header.ML"; ML_file "Thy/document_marker.ML"; (*proof context*) ML_file "Isar/object_logic.ML"; ML_file "Isar/rule_cases.ML"; ML_file "Isar/auto_bind.ML"; ML_file "type_infer.ML"; ML_file "Syntax/local_syntax.ML"; ML_file "Isar/proof_context.ML"; ML_file "type_infer_context.ML"; ML_file "Syntax/syntax_phases.ML"; ML_file "Isar/args.ML"; (*theory specifications*) ML_file "Isar/local_defs.ML"; ML_file "Isar/local_theory.ML"; ML_file "Isar/entity.ML"; ML_file "PIDE/command_span.ML"; ML_file "Thy/thy_element.ML"; ML_file "Thy/markdown.ML"; ML_file "Thy/html.ML"; ML_file "Thy/latex.ML"; (*ML with context and antiquotations*) ML_file "ML/ml_context.ML"; ML_file "ML/ml_antiquotation.ML"; ML_file "ML/ml_compiler2.ML"; section "Bootstrap phase 3: towards theory Pure and final ML toplevel setup"; (*basic proof engine*) ML_file "par_tactical.ML"; ML_file "context_tactic.ML"; ML_file "Isar/proof_display.ML"; ML_file "Isar/attrib.ML"; ML_file "Isar/context_rules.ML"; ML_file "Isar/method.ML"; ML_file "Isar/proof.ML"; ML_file "Isar/element.ML"; ML_file "Isar/obtain.ML"; ML_file "Isar/subgoal.ML"; ML_file "Isar/calculation.ML"; (*local theories and targets*) ML_file "Isar/locale.ML"; ML_file "Isar/generic_target.ML"; ML_file "Isar/overloading.ML"; ML_file "axclass.ML"; ML_file "Isar/class.ML"; ML_file "Isar/named_target.ML"; ML_file "Isar/expression.ML"; ML_file "Isar/interpretation.ML"; ML_file "Isar/class_declaration.ML"; ML_file "Isar/bundle.ML"; ML_file "Isar/experiment.ML"; ML_file "simplifier.ML"; ML_file "Tools/plugin.ML"; (*executable theory content*) ML_file "Isar/code.ML"; (*specifications*) ML_file "Isar/spec_rules.ML"; ML_file "Isar/specification.ML"; ML_file "Isar/parse_spec.ML"; ML_file "Isar/typedecl.ML"; (*toplevel transactions*) ML_file "Isar/proof_node.ML"; ML_file "Isar/toplevel.ML"; (*proof term operations*) +ML_file "Proof/proof_rewrite_rules.ML"; ML_file "Proof/proof_syntax.ML"; -ML_file "Proof/proof_rewrite_rules.ML"; ML_file "Proof/proof_checker.ML"; ML_file "Proof/extraction.ML"; (*Isabelle system*) ML_file "System/bash.ML"; ML_file "System/isabelle_system.ML"; (*theory documents*) ML_file "Thy/term_style.ML"; ML_file "Isar/outer_syntax.ML"; ML_file "ML/ml_antiquotations.ML"; ML_file "Thy/document_antiquotation.ML"; ML_file "Thy/thy_output.ML"; ML_file "Thy/document_antiquotations.ML"; ML_file "General/graph_display.ML"; ML_file "pure_syn.ML"; ML_file "PIDE/command.ML"; ML_file "PIDE/query_operation.ML"; ML_file "PIDE/resources.ML"; ML_file "Thy/present.ML"; ML_file "Thy/thy_info.ML"; ML_file "thm_deps.ML"; ML_file "Thy/export_theory.ML"; ML_file "Thy/sessions.ML"; ML_file "PIDE/session.ML"; ML_file "PIDE/document.ML"; (*theory and proof operations*) ML_file "Isar/isar_cmd.ML"; subsection "Isabelle/Isar system"; ML_file "System/command_line.ML"; ML_file "System/message_channel.ML"; ML_file "System/isabelle_process.ML"; ML_file "System/invoke_scala.ML"; ML_file "Thy/bibtex.ML"; ML_file "PIDE/protocol.ML"; ML_file "General/output_primitives_virtual.ML"; subsection "Miscellaneous tools and packages for Pure Isabelle"; ML_file "ML/ml_pp.ML"; ML_file "ML/ml_thms.ML"; ML_file "ML/ml_file.ML"; ML_file "Tools/build.ML"; ML_file "Tools/named_thms.ML"; ML_file "Tools/print_operation.ML"; ML_file "Tools/rail.ML"; ML_file "Tools/rule_insts.ML"; ML_file "Tools/thy_deps.ML"; ML_file "Tools/class_deps.ML"; ML_file "Tools/find_theorems.ML"; ML_file "Tools/find_consts.ML"; ML_file "Tools/simplifier_trace.ML"; ML_file_no_debug "Tools/debugger.ML"; ML_file "Tools/named_theorems.ML"; ML_file "Tools/jedit.ML"; ML_file "Tools/ghc.ML"; ML_file "Tools/generated_files.ML" diff --git a/src/Pure/Thy/export_theory.ML b/src/Pure/Thy/export_theory.ML --- a/src/Pure/Thy/export_theory.ML +++ b/src/Pure/Thy/export_theory.ML @@ -1,406 +1,407 @@ (* Title: Pure/Thy/export_theory.ML Author: Makarius Export foundational theory content and locale/class structure. *) signature EXPORT_THEORY = sig val setup_presentation: (Thy_Info.presentation_context -> theory -> unit) -> unit val export_body: theory -> string -> XML.body -> unit end; structure Export_Theory: EXPORT_THEORY = struct (* approximative syntax *) val get_syntax = Syntax.get_approx o Proof_Context.syn_of; fun get_syntax_type ctxt = get_syntax ctxt o Lexicon.mark_type; fun get_syntax_const ctxt = get_syntax ctxt o Lexicon.mark_const; fun get_syntax_fixed ctxt = get_syntax ctxt o Lexicon.mark_fixed; fun get_syntax_param ctxt loc x = let val thy = Proof_Context.theory_of ctxt in if Class.is_class thy loc then (case AList.lookup (op =) (Class.these_params thy [loc]) x of NONE => NONE | SOME (_, (c, _)) => get_syntax_const ctxt c) else get_syntax_fixed ctxt x end; val encode_syntax = XML.Encode.variant [fn NONE => ([], []), fn SOME (Syntax.Prefix delim) => ([delim], []), fn SOME (Syntax.Infix {assoc, delim, pri}) => let val ass = (case assoc of Printer.No_Assoc => 0 | Printer.Left_Assoc => 1 | Printer.Right_Assoc => 2); open XML.Encode Term_XML.Encode; in ([], triple int string int (ass, delim, pri)) end]; (* free variables: not declared in the context *) val is_free = not oo Name.is_declared; fun add_frees used = fold_aterms (fn Free (x, T) => is_free used x ? insert (op =) (x, T) | _ => I); fun add_tfrees used = (fold_types o fold_atyps) (fn TFree (a, S) => is_free used a ? insert (op =) (a, S) | _ => I); (* spec rules *) fun primrec_types ctxt const = Spec_Rules.retrieve ctxt (Const const) |> get_first (fn (Spec_Rules.Equational (Spec_Rules.Primrec types), _) => SOME (types, false) | (Spec_Rules.Equational (Spec_Rules.Primcorec types), _) => SOME (types, true) | _ => NONE) |> the_default ([], false); (* locales content *) fun locale_content thy loc = let val ctxt = Locale.init loc thy; val args = Locale.params_of thy loc |> map (fn ((x, T), _) => ((x, T), get_syntax_param ctxt loc x)); val axioms = let val (asm, defs) = Locale.specification_of thy loc; val cprops = map (Thm.cterm_of ctxt) (the_list asm @ defs); val (intro1, intro2) = Locale.intros_of thy loc; val intros_tac = Method.try_intros_tac ctxt (the_list intro1 @ the_list intro2) []; val res = Goal.init (Conjunction.mk_conjunction_balanced cprops) |> (ALLGOALS Goal.conjunction_tac THEN intros_tac) |> try Seq.hd; in (case res of SOME goal => Thm.prems_of goal | NONE => raise Fail ("Cannot unfold locale " ^ quote loc)) end; val typargs = rev (fold Term.add_tfrees (map (Free o #1) args @ axioms) []); in {typargs = typargs, args = args, axioms = axioms} end; fun get_locales thy = Locale.get_locales thy |> map_filter (fn loc => if Experiment.is_experiment thy loc then NONE else SOME (loc, ())); fun get_dependencies prev_thys thy = Locale.dest_dependencies prev_thys thy |> map_filter (fn dep => if Experiment.is_experiment thy (#source dep) orelse Experiment.is_experiment thy (#target dep) then NONE else let val (type_params, params) = Locale.parameters_of thy (#source dep); val typargs = fold (Term.add_tfreesT o #2 o #1) params type_params; val substT = typargs |> map_filter (fn v => let val T = TFree v; val T' = Morphism.typ (#morphism dep) T; in if T = T' then NONE else SOME (v, T') end); val subst = params |> map_filter (fn (v, _) => let val t = Free v; val t' = Morphism.term (#morphism dep) t; in if t aconv t' then NONE else SOME (v, t') end); in SOME (dep, (substT, subst)) end); (* general setup *) fun setup_presentation f = Theory.setup (Thy_Info.add_presentation (fn context => fn thy => if Options.bool (#options context) "export_theory" then f context thy else ())); fun export_body thy name body = if XML.is_empty_body body then () else Export.export thy (Path.binding0 (Path.make ["theory", name])) body; (* presentation *) val _ = setup_presentation (fn {adjust_pos, ...} => fn thy => let val parents = Theory.parents_of thy; val rep_tsig = Type.rep_tsig (Sign.tsig_of thy); val thy_ctxt = Proof_Context.init_global thy; (* entities *) fun make_entity_markup name xname pos serial = let val props = Position.offset_properties_of (adjust_pos pos) @ Position.id_properties_of pos @ Markup.serial_properties serial; in (Markup.entityN, (Markup.nameN, name) :: (Markup.xnameN, xname) :: props) end; fun entity_markup space name = let val xname = Name_Space.extern_shortest thy_ctxt space name; val {serial, pos, ...} = Name_Space.the_entry space name; in make_entity_markup name xname pos serial end; fun export_entities export_name export get_space decls = let val parent_spaces = map get_space parents; val space = get_space thy; in (decls, []) |-> fold (fn (name, decl) => if exists (fn space => Name_Space.declared space name) parent_spaces then I else (case export name decl of NONE => I | SOME body => cons (#serial (Name_Space.the_entry space name), XML.Elem (entity_markup space name, body)))) |> sort (int_ord o apply2 #1) |> map #2 |> export_body thy export_name end; (* types *) val encode_type = let open XML.Encode Term_XML.Encode in triple encode_syntax (list string) (option typ) end; fun export_type c (Type.LogicalType n) = SOME (encode_type (get_syntax_type thy_ctxt c, Name.invent Name.context Name.aT n, NONE)) | export_type c (Type.Abbreviation (args, U, false)) = SOME (encode_type (get_syntax_type thy_ctxt c, args, SOME U)) | export_type _ _ = NONE; val _ = export_entities "types" export_type Sign.type_space (Name_Space.dest_table (#types rep_tsig)); (* consts *) val consts = Sign.consts_of thy; val encode_term = Term_XML.Encode.term consts; val encode_const = let open XML.Encode Term_XML.Encode in pair encode_syntax (pair (list string) (pair typ (pair (option encode_term) (pair bool (pair (list string) bool))))) end; fun export_const c (T, abbrev) = let val syntax = get_syntax_const thy_ctxt c; val U = Logic.unvarifyT_global T; val U0 = Type.strip_sorts U; val recursion = primrec_types thy_ctxt (c, U); val abbrev' = abbrev |> Option.map (Proofterm.standard_vars_term Name.context #> map_types Type.strip_sorts); val args = map (#1 o dest_TFree) (Consts.typargs consts (c, U0)); val propositional = Object_Logic.is_propositional thy_ctxt (Term.body_type U0); in encode_const (syntax, (args, (U0, (abbrev', (propositional, recursion))))) end; val _ = export_entities "consts" (SOME oo export_const) Sign.const_space (#constants (Consts.dest consts)); (* axioms *) fun standard_prop used extra_shyps raw_prop raw_proof = let val (prop, proof) = Proofterm.standard_vars used (raw_prop, raw_proof); val args = rev (add_frees used prop []); val typargs = rev (add_tfrees used prop []); val used_typargs = fold (Name.declare o #1) typargs used; val sorts = Name.invent used_typargs Name.aT (length extra_shyps) ~~ extra_shyps; in ((sorts @ typargs, args, prop), proof) end; val encode_prop = let open XML.Encode Term_XML.Encode in triple (list (pair string sort)) (list (pair string typ)) encode_term end; fun encode_axiom used prop = encode_prop (#1 (standard_prop used [] prop NONE)); val _ = export_entities "axioms" (K (SOME o encode_axiom Name.context)) Theory.axiom_space (Theory.all_axioms_of thy); (* theorems and proof terms *) val clean_thm = Thm.check_hyps (Context.Theory thy) #> Thm.strip_shyps; val lookup_thm_id = Global_Theory.lookup_thm_id thy; fun expand_name thm_id (header: Proofterm.thm_header) = if #serial header = #serial thm_id then "" else (case lookup_thm_id (Proofterm.thm_header_id header) of NONE => "" | SOME thm_name => Thm_Name.print thm_name); fun entity_markup_thm (serial, (name, i)) = let val space = Facts.space_of (Global_Theory.facts_of thy); val xname = Name_Space.extern_shortest thy_ctxt space name; val {pos, ...} = Name_Space.the_entry space name; in make_entity_markup (Thm_Name.print (name, i)) (Thm_Name.print (xname, i)) pos serial end; fun encode_thm thm_id raw_thm = let val deps = map (Thm_Name.print o #2) (Thm_Deps.thm_deps thy [raw_thm]); val thm = clean_thm (Thm.unconstrainT raw_thm); val proof0 = if Proofterm.export_standard_enabled () then - Thm.standard_proof_of {full = true, expand_name = SOME o expand_name thm_id} thm + Proof_Syntax.standard_proof_of + {full = true, expand_name = SOME o expand_name thm_id} thm else if Proofterm.export_enabled () then Thm.reconstruct_proof_of thm else MinProof; val (prop, SOME proof) = standard_prop Name.context (Thm.extra_shyps thm) (Thm.full_prop_of thm) (SOME proof0); val _ = Thm.expose_proofs thy [thm]; in (prop, deps, proof) |> let open XML.Encode Term_XML.Encode; val encode_proof = Proofterm.encode_standard_proof consts; in triple encode_prop (list string) encode_proof end end; fun export_thm (thm_id, thm_name) = let val markup = entity_markup_thm (#serial thm_id, thm_name); val thm = Global_Theory.get_thm_name thy (thm_name, Position.none); in XML.Elem (markup, encode_thm thm_id thm) end; val _ = export_body thy "thms" (map export_thm (Global_Theory.dest_thm_names thy)); (* type classes *) val encode_class = let open XML.Encode Term_XML.Encode in pair (list (pair string typ)) (list (encode_axiom Name.context)) end; fun export_class name = (case try (Axclass.get_info thy) name of NONE => ([], []) | SOME {params, axioms, ...} => (params, map (Thm.plain_prop_of o clean_thm) axioms)) |> encode_class |> SOME; val _ = export_entities "classes" (fn name => fn () => export_class name) Sign.class_space (map (rpair ()) (Graph.keys (Sorts.classes_of (#2 (#classes rep_tsig))))); (* sort algebra *) local val prop = encode_axiom Name.context o Logic.varify_global; val encode_classrel = let open XML.Encode in list (pair prop (pair string string)) end; val encode_arities = let open XML.Encode Term_XML.Encode in list (pair prop (triple string (list sort) string)) end; in val export_classrel = maps (fn (c, cs) => map (pair c) cs) #> map (`Logic.mk_classrel) #> encode_classrel; val export_arities = map (`Logic.mk_arity) #> encode_arities; val {classrel, arities} = Sorts.dest_algebra (map (#2 o #classes o Type.rep_tsig o Sign.tsig_of) parents) (#2 (#classes rep_tsig)); end; val _ = if null classrel then () else export_body thy "classrel" (export_classrel classrel); val _ = if null arities then () else export_body thy "arities" (export_arities arities); (* locales *) fun encode_locale used = let open XML.Encode Term_XML.Encode in triple (list (pair string sort)) (list (pair (pair string typ) encode_syntax)) (list (encode_axiom used)) end; fun export_locale loc = let val {typargs, args, axioms} = locale_content thy loc; val used = fold Name.declare (map #1 typargs @ map (#1 o #1) args) Name.context; in encode_locale used (typargs, args, axioms) end handle ERROR msg => cat_error msg ("The error(s) above occurred in locale " ^ quote (Locale.markup_name thy_ctxt loc)); val _ = export_entities "locales" (fn loc => fn () => SOME (export_locale loc)) Locale.locale_space (get_locales thy); (* locale dependencies *) fun encode_locale_dependency (dep: Locale.locale_dependency, subst) = (#source dep, (#target dep, (#prefix dep, subst))) |> let open XML.Encode Term_XML.Encode; val encode_subst = pair (list (pair (pair string sort) typ)) (list (pair (pair string typ) (term consts))); in pair string (pair string (pair (list (pair string bool)) encode_subst)) end; val _ = get_dependencies parents thy |> map_index (fn (i, dep) => let val xname = string_of_int (i + 1); val name = Long_Name.implode [Context.theory_name thy, xname]; val markup = make_entity_markup name xname (#pos (#1 dep)) (#serial (#1 dep)); val body = encode_locale_dependency dep; in XML.Elem (markup, body) end) |> export_body thy "locale_dependencies"; (* constdefs *) val constdefs = Defs.dest_constdefs (map Theory.defs_of (Theory.parents_of thy)) (Theory.defs_of thy) |> sort_by #1; val encode_constdefs = let open XML.Encode in list (pair string string) end; val _ = if null constdefs then () else export_body thy "constdefs" (encode_constdefs constdefs); (* parents *) val _ = Export.export thy \<^path_binding>\theory/parents\ (XML.Encode.string (cat_lines (map Context.theory_long_name parents))); in () end); end; diff --git a/src/Pure/more_thm.ML b/src/Pure/more_thm.ML --- a/src/Pure/more_thm.ML +++ b/src/Pure/more_thm.ML @@ -1,755 +1,736 @@ (* Title: Pure/more_thm.ML Author: Makarius Further operations on type ctyp/cterm/thm, outside the inference kernel. *) infix aconvc; signature BASIC_THM = sig include BASIC_THM val show_consts: bool Config.T val show_hyps: bool Config.T val show_tags: bool Config.T structure Ctermtab: TABLE structure Thmtab: TABLE val aconvc: cterm * cterm -> bool type attribute = Context.generic * thm -> Context.generic option * thm option end; signature THM = sig include THM structure Ctermtab: TABLE structure Thmtab: TABLE val eq_ctyp: ctyp * ctyp -> bool val aconvc: cterm * cterm -> bool val add_tvars: thm -> ctyp list -> ctyp list val add_frees: thm -> cterm list -> cterm list val add_vars: thm -> cterm list -> cterm list val dest_funT: ctyp -> ctyp * ctyp val strip_type: ctyp -> ctyp list * ctyp val all_name: Proof.context -> string * cterm -> cterm -> cterm val all: Proof.context -> cterm -> cterm -> cterm val mk_binop: cterm -> cterm -> cterm -> cterm val dest_binop: cterm -> cterm * cterm val dest_implies: cterm -> cterm * cterm val dest_equals: cterm -> cterm * cterm val dest_equals_lhs: cterm -> cterm val dest_equals_rhs: cterm -> cterm val lhs_of: thm -> cterm val rhs_of: thm -> cterm val fast_term_ord: cterm ord val term_ord: cterm ord val thm_ord: thm ord val cterm_cache: (cterm -> 'a) -> cterm -> 'a val thm_cache: (thm -> 'a) -> thm -> 'a val is_reflexive: thm -> bool val eq_thm: thm * thm -> bool val eq_thm_prop: thm * thm -> bool val eq_thm_strict: thm * thm -> bool val equiv_thm: theory -> thm * thm -> bool val class_triv: theory -> class -> thm val of_sort: ctyp * sort -> thm list val is_dummy: thm -> bool val add_thm: thm -> thm list -> thm list val del_thm: thm -> thm list -> thm list val merge_thms: thm list * thm list -> thm list val full_rules: thm Item_Net.T val intro_rules: thm Item_Net.T val elim_rules: thm Item_Net.T val declare_hyps: cterm -> Proof.context -> Proof.context val assume_hyps: cterm -> Proof.context -> thm * Proof.context val unchecked_hyps: Proof.context -> Proof.context val restore_hyps: Proof.context -> Proof.context -> Proof.context val undeclared_hyps: Context.generic -> thm -> term list val check_hyps: Context.generic -> thm -> thm val declare_term_sorts: term -> Proof.context -> Proof.context val extra_shyps': Proof.context -> thm -> sort list val check_shyps: Proof.context -> thm -> thm val weaken_sorts': Proof.context -> cterm -> cterm val elim_implies: thm -> thm -> thm val forall_intr_name: string * cterm -> thm -> thm val forall_elim_var: int -> thm -> thm val forall_elim_vars: int -> thm -> thm val instantiate_frees: ((string * sort) * ctyp) list * ((string * typ) * cterm) list -> thm -> thm val instantiate': ctyp option list -> cterm option list -> thm -> thm val forall_intr_frees: thm -> thm val unvarify_global: theory -> thm -> thm val unvarify_axiom: theory -> string -> thm val rename_params_rule: string list * int -> thm -> thm val rename_boundvars: term -> term -> thm -> thm val add_axiom: Proof.context -> binding * term -> theory -> (string * thm) * theory val add_axiom_global: binding * term -> theory -> (string * thm) * theory val add_def: Defs.context -> bool -> bool -> binding * term -> theory -> (string * thm) * theory val add_def_global: bool -> bool -> binding * term -> theory -> (string * thm) * theory type attribute = Context.generic * thm -> Context.generic option * thm option type binding = binding * attribute list val tag_rule: string * string -> thm -> thm val untag_rule: string -> thm -> thm val is_free_dummy: thm -> bool val tag_free_dummy: thm -> thm val def_name: string -> string val def_name_optional: string -> string -> string val def_binding: Binding.binding -> Binding.binding val def_binding_optional: Binding.binding -> Binding.binding -> Binding.binding val make_def_binding: bool -> Binding.binding -> Binding.binding val has_name_hint: thm -> bool val get_name_hint: thm -> string val put_name_hint: string -> thm -> thm val theoremK: string val legacy_get_kind: thm -> string val kind_rule: string -> thm -> thm val rule_attribute: thm list -> (Context.generic -> thm -> thm) -> attribute val declaration_attribute: (thm -> Context.generic -> Context.generic) -> attribute val mixed_attribute: (Context.generic * thm -> Context.generic * thm) -> attribute val apply_attribute: attribute -> thm -> Context.generic -> thm * Context.generic val attribute_declaration: attribute -> thm -> Context.generic -> Context.generic val theory_attributes: attribute list -> thm -> theory -> thm * theory val proof_attributes: attribute list -> thm -> Proof.context -> thm * Proof.context val no_attributes: 'a -> 'a * 'b list val simple_fact: 'a -> ('a * 'b list) list val tag: string * string -> attribute val untag: string -> attribute val kind: string -> attribute - val reconstruct_proof_of: thm -> Proofterm.proof - val standard_proof_of: {full: bool, expand_name: Proofterm.thm_header -> string option} -> - thm -> Proofterm.proof val register_proofs: thm list lazy -> theory -> theory val consolidate_theory: theory -> unit val expose_theory: theory -> unit val show_consts: bool Config.T val show_hyps: bool Config.T val show_tags: bool Config.T val pretty_thm_raw: Proof.context -> {quote: bool, show_hyps: bool} -> thm -> Pretty.T val pretty_thm: Proof.context -> thm -> Pretty.T val pretty_thm_item: Proof.context -> thm -> Pretty.T val pretty_thm_global: theory -> thm -> Pretty.T val string_of_thm: Proof.context -> thm -> string val string_of_thm_global: theory -> thm -> string end; structure Thm: THM = struct (** basic operations **) (* collecting ctyps and cterms *) val eq_ctyp = op = o apply2 Thm.typ_of; val op aconvc = op aconv o apply2 Thm.term_of; val add_tvars = Thm.fold_atomic_ctyps (fn a => is_TVar (Thm.typ_of a) ? insert eq_ctyp a); val add_frees = Thm.fold_atomic_cterms (fn a => is_Free (Thm.term_of a) ? insert (op aconvc) a); val add_vars = Thm.fold_atomic_cterms (fn a => is_Var (Thm.term_of a) ? insert (op aconvc) a); (* ctyp operations *) fun dest_funT cT = (case Thm.typ_of cT of Type ("fun", _) => let val [A, B] = Thm.dest_ctyp cT in (A, B) end | T => raise TYPE ("dest_funT", [T], [])); (* ctyp version of strip_type: maps [T1,...,Tn]--->T to ([T1,T2,...,Tn], T) *) fun strip_type cT = (case Thm.typ_of cT of Type ("fun", _) => let val (cT1, cT2) = dest_funT cT; val (cTs, cT') = strip_type cT2 in (cT1 :: cTs, cT') end | _ => ([], cT)); (* cterm operations *) fun all_name ctxt (x, t) A = let val T = Thm.typ_of_cterm t; val all_const = Thm.cterm_of ctxt (Const ("Pure.all", (T --> propT) --> propT)); in Thm.apply all_const (Thm.lambda_name (x, t) A) end; fun all ctxt t A = all_name ctxt ("", t) A; fun mk_binop c a b = Thm.apply (Thm.apply c a) b; fun dest_binop ct = (Thm.dest_arg1 ct, Thm.dest_arg ct); fun dest_implies ct = (case Thm.term_of ct of Const ("Pure.imp", _) $ _ $ _ => dest_binop ct | _ => raise TERM ("dest_implies", [Thm.term_of ct])); fun dest_equals ct = (case Thm.term_of ct of Const ("Pure.eq", _) $ _ $ _ => dest_binop ct | _ => raise TERM ("dest_equals", [Thm.term_of ct])); fun dest_equals_lhs ct = (case Thm.term_of ct of Const ("Pure.eq", _) $ _ $ _ => Thm.dest_arg1 ct | _ => raise TERM ("dest_equals_lhs", [Thm.term_of ct])); fun dest_equals_rhs ct = (case Thm.term_of ct of Const ("Pure.eq", _) $ _ $ _ => Thm.dest_arg ct | _ => raise TERM ("dest_equals_rhs", [Thm.term_of ct])); val lhs_of = dest_equals_lhs o Thm.cprop_of; val rhs_of = dest_equals_rhs o Thm.cprop_of; (* certified term order *) val fast_term_ord = Term_Ord.fast_term_ord o apply2 Thm.term_of; val term_ord = Term_Ord.term_ord o apply2 Thm.term_of; (* thm order: ignores theory context! *) val thm_ord = Term_Ord.fast_term_ord o apply2 Thm.prop_of <<< list_ord (prod_ord Term_Ord.fast_term_ord Term_Ord.fast_term_ord) o apply2 Thm.tpairs_of <<< list_ord Term_Ord.fast_term_ord o apply2 Thm.hyps_of <<< list_ord Term_Ord.sort_ord o apply2 Thm.shyps_of; (* tables and caches *) structure Ctermtab = Table(type key = cterm val ord = fast_term_ord); structure Thmtab = Table(type key = thm val ord = thm_ord); fun cterm_cache f = Cache.create Ctermtab.empty Ctermtab.lookup Ctermtab.update f; fun thm_cache f = Cache.create Thmtab.empty Thmtab.lookup Thmtab.update f; (* equality *) fun is_reflexive th = op aconv (Logic.dest_equals (Thm.prop_of th)) handle TERM _ => false; val eq_thm = is_equal o thm_ord; val eq_thm_prop = op aconv o apply2 Thm.full_prop_of; fun eq_thm_strict ths = eq_thm ths andalso Context.eq_thy_id (apply2 Thm.theory_id ths) andalso op = (apply2 Thm.maxidx_of ths) andalso op = (apply2 Thm.get_tags ths); (* pattern equivalence *) fun equiv_thm thy ths = Pattern.equiv thy (apply2 (Thm.full_prop_of o Thm.transfer thy) ths); (* type classes and sorts *) fun class_triv thy c = Thm.of_class (Thm.global_ctyp_of thy (TVar ((Name.aT, 0), [c])), c); fun of_sort (T, S) = map (fn c => Thm.of_class (T, c)) S; (* misc operations *) fun is_dummy thm = (case try Logic.dest_term (Thm.concl_of thm) of NONE => false | SOME t => Term.is_dummy_pattern (Term.head_of t)); (* collections of theorems in canonical order *) val add_thm = update eq_thm_prop; val del_thm = remove eq_thm_prop; val merge_thms = merge eq_thm_prop; val full_rules = Item_Net.init eq_thm_prop (single o Thm.full_prop_of); val intro_rules = Item_Net.init eq_thm_prop (single o Thm.concl_of); val elim_rules = Item_Net.init eq_thm_prop (single o Thm.major_prem_of); (** declared hyps and sort hyps **) structure Hyps = Proof_Data ( type T = {checked_hyps: bool, hyps: Termtab.set, shyps: sort Ord_List.T}; fun init _ : T = {checked_hyps = true, hyps = Termtab.empty, shyps = []}; ); fun map_hyps f = Hyps.map (fn {checked_hyps, hyps, shyps} => let val (checked_hyps', hyps', shyps') = f (checked_hyps, hyps, shyps) in {checked_hyps = checked_hyps', hyps = hyps', shyps = shyps'} end); (* hyps *) fun declare_hyps raw_ct ctxt = ctxt |> map_hyps (fn (checked_hyps, hyps, shyps) => let val ct = Thm.transfer_cterm (Proof_Context.theory_of ctxt) raw_ct; val hyps' = Termtab.update (Thm.term_of ct, ()) hyps; in (checked_hyps, hyps', shyps) end); fun assume_hyps ct ctxt = (Thm.assume ct, declare_hyps ct ctxt); val unchecked_hyps = map_hyps (fn (_, hyps, shyps) => (false, hyps, shyps)); fun restore_hyps ctxt = map_hyps (fn (_, hyps, shyps) => (#checked_hyps (Hyps.get ctxt), hyps, shyps)); fun undeclared_hyps context th = Thm.hyps_of th |> filter_out (case context of Context.Theory _ => K false | Context.Proof ctxt => (case Hyps.get ctxt of {checked_hyps = false, ...} => K true | {hyps, ...} => Termtab.defined hyps)); fun check_hyps context th = (case undeclared_hyps context th of [] => th | undeclared => error (Pretty.string_of (Pretty.big_list "Undeclared hyps:" (map (Pretty.item o single o Syntax.pretty_term (Syntax.init_pretty context)) undeclared)))); (* shyps *) fun declare_term_sorts t = map_hyps (fn (checked_hyps, hyps, shyps) => (checked_hyps, hyps, Sorts.insert_term t shyps)); fun extra_shyps' ctxt th = Sorts.subtract (#shyps (Hyps.get ctxt)) (Thm.extra_shyps th); fun check_shyps ctxt raw_th = let val th = Thm.strip_shyps raw_th; val extra_shyps = extra_shyps' ctxt th; in if null extra_shyps then th else error (Pretty.string_of (Pretty.block (Pretty.str "Pending sort hypotheses:" :: Pretty.brk 1 :: Pretty.commas (map (Syntax.pretty_sort ctxt) extra_shyps)))) end; val weaken_sorts' = Thm.weaken_sorts o #shyps o Hyps.get; (** basic derived rules **) (*Elimination of implication A A \ B ------------ B *) fun elim_implies thA thAB = Thm.implies_elim thAB thA; (* forall_intr_name *) fun forall_intr_name (a, x) th = let val th' = Thm.forall_intr x th; val prop' = (case Thm.prop_of th' of all $ Abs (_, T, b) => all $ Abs (a, T, b)); in Thm.renamed_prop prop' th' end; (* forall_elim_var(s) *) local fun dest_all ct = (case Thm.term_of ct of Const ("Pure.all", _) $ Abs (a, _, _) => let val (x, ct') = Thm.dest_abs NONE (Thm.dest_arg ct) in SOME ((a, Thm.ctyp_of_cterm x), ct') end | _ => NONE); fun dest_all_list ct = (case dest_all ct of NONE => [] | SOME (v, ct') => v :: dest_all_list ct'); fun forall_elim_vars_list vars i th = let val used = (Thm.fold_terms o Term.fold_aterms) (fn Var ((x, j), _) => if i = j then insert (op =) x else I | _ => I) th []; val vars' = (Name.variant_list used (map #1 vars), vars) |> ListPair.map (fn (x, (_, T)) => Thm.var ((x, i), T)); in fold Thm.forall_elim vars' th end; in fun forall_elim_vars i th = forall_elim_vars_list (dest_all_list (Thm.cprop_of th)) i th; fun forall_elim_var i th = let val vars = (case dest_all (Thm.cprop_of th) of SOME (v, _) => [v] | NONE => raise THM ("forall_elim_var", i, [th])); in forall_elim_vars_list vars i th end; end; (* instantiate frees *) fun instantiate_frees ([], []) th = th | instantiate_frees (instT, inst) th = let val idx = Thm.maxidx_of th + 1; fun index ((a, A), b) = (((a, idx), A), b); val insts = (map index instT, map index inst); val frees = (map (#1 o #1) instT, map (#1 o #1) inst); val hyps = Thm.chyps_of th; val inst_cterm = Thm.generalize_cterm frees idx #> Thm.instantiate_cterm insts; in th |> fold_rev Thm.implies_intr hyps |> Thm.generalize frees idx |> Thm.instantiate insts |> fold (elim_implies o Thm.assume o inst_cterm) hyps end; (* instantiate by left-to-right occurrence of variables *) fun instantiate' cTs cts thm = let fun err msg = raise TYPE ("instantiate': " ^ msg, map_filter (Option.map Thm.typ_of) cTs, map_filter (Option.map Thm.term_of) cts); fun zip_vars xs ys = zip_options xs ys handle ListPair.UnequalLengths => err "more instantiations than variables in thm"; val thm' = Thm.instantiate ((zip_vars (rev (Thm.fold_terms Term.add_tvars thm [])) cTs), []) thm; val thm'' = Thm.instantiate ([], zip_vars (rev (Thm.fold_terms Term.add_vars thm' [])) cts) thm'; in thm'' end; (* forall_intr_frees: generalization over all suitable Free variables *) fun forall_intr_frees th = let val fixed = fold Term.add_frees (Thm.terms_of_tpairs (Thm.tpairs_of th) @ Thm.hyps_of th) []; val frees = Thm.fold_atomic_cterms (fn a => (case Thm.term_of a of Free v => not (member (op =) fixed v) ? insert (op aconvc) a | _ => I)) th []; in fold Thm.forall_intr frees th end; (* unvarify_global: global schematic variables *) fun unvarify_global thy th = let val prop = Thm.full_prop_of th; val _ = map Logic.unvarify_global (prop :: Thm.hyps_of th) handle TERM (msg, _) => raise THM (msg, 0, [th]); val instT = rev (Term.add_tvars prop []) |> map (fn v as ((a, _), S) => (v, TFree (a, S))); val inst = rev (Term.add_vars prop []) |> map (fn ((a, i), T) => let val T' = Term_Subst.instantiateT instT T in (((a, i), T'), Thm.global_cterm_of thy (Free ((a, T')))) end); in Thm.instantiate (map (apsnd (Thm.global_ctyp_of thy)) instT, inst) th end; fun unvarify_axiom thy = unvarify_global thy o Thm.axiom thy; (* user renaming of parameters in a subgoal *) (*The names, if distinct, are used for the innermost parameters of subgoal i; preceding parameters may be renamed to make all parameters distinct.*) fun rename_params_rule (names, i) st = let val (_, Bs, Bi, C) = Thm.dest_state (st, i); val params = map #1 (Logic.strip_params Bi); val short = length params - length names; val names' = if short < 0 then error "More names than parameters in subgoal!" else Name.variant_list names (take short params) @ names; val free_names = Term.fold_aterms (fn Free (x, _) => insert (op =) x | _ => I) Bi []; val Bi' = Logic.list_rename_params names' Bi; in (case duplicates (op =) names of a :: _ => (warning ("Can't rename. Bound variables not distinct: " ^ a); st) | [] => (case inter (op =) names free_names of a :: _ => (warning ("Can't rename. Bound/Free variable clash: " ^ a); st) | [] => Thm.renamed_prop (Logic.list_implies (Bs @ [Bi'], C)) st)) end; (* preservation of bound variable names *) fun rename_boundvars pat obj th = (case Term.rename_abs pat obj (Thm.prop_of th) of NONE => th | SOME prop' => Thm.renamed_prop prop' th); (** specification primitives **) (* rules *) fun stripped_sorts thy t = let val tfrees = rev (Term.add_tfrees t []); val tfrees' = map (fn a => (a, [])) (Name.variant_list [] (map #1 tfrees)); val recover = map2 (fn (a', S') => fn (a, S) => (((a', 0), S'), Thm.global_ctyp_of thy (TVar ((a, 0), S)))) tfrees' tfrees; val strip = map (apply2 TFree) (tfrees ~~ tfrees'); val t' = Term.map_types (Term.map_atyps (perhaps (AList.lookup (op =) strip))) t; in (strip, recover, t') end; fun add_axiom ctxt (b, prop) thy = let val _ = Sign.no_vars ctxt prop; val (strip, recover, prop') = stripped_sorts thy prop; val constraints = map (fn (TFree (_, S), T) => (T, S)) strip; val of_sorts = maps (fn (T as TFree (_, S), _) => of_sort (Thm.ctyp_of ctxt T, S)) strip; val thy' = thy |> Theory.add_axiom ctxt (b, Logic.list_implies (maps Logic.mk_of_sort constraints, prop')); val axm_name = Sign.full_name thy' b; val axm' = Thm.axiom thy' axm_name; val thm = Thm.instantiate (recover, []) axm' |> unvarify_global thy' |> fold elim_implies of_sorts; in ((axm_name, thm), thy') end; fun add_axiom_global arg thy = add_axiom (Syntax.init_pretty_global thy) arg thy; fun add_def (context as (ctxt, _)) unchecked overloaded (b, prop) thy = let val _ = Sign.no_vars ctxt prop; val prems = map (Thm.cterm_of ctxt) (Logic.strip_imp_prems prop); val (_, recover, concl') = stripped_sorts thy (Logic.strip_imp_concl prop); val thy' = Theory.add_def context unchecked overloaded (b, concl') thy; val axm_name = Sign.full_name thy' b; val axm' = Thm.axiom thy' axm_name; val thm = Thm.instantiate (recover, []) axm' |> unvarify_global thy' |> fold_rev Thm.implies_intr prems; in ((axm_name, thm), thy') end; fun add_def_global unchecked overloaded arg thy = add_def (Defs.global_context thy) unchecked overloaded arg thy; (** theorem tags **) (* add / delete tags *) fun tag_rule tg = Thm.map_tags (insert (op =) tg); fun untag_rule s = Thm.map_tags (filter_out (fn (s', _) => s = s')); (* free dummy thm -- for abstract closure *) val free_dummyN = "free_dummy"; fun is_free_dummy thm = Properties.defined (Thm.get_tags thm) free_dummyN; val tag_free_dummy = tag_rule (free_dummyN, ""); (* def_name *) fun def_name c = c ^ "_def"; fun def_name_optional c "" = def_name c | def_name_optional _ name = name; val def_binding = Binding.map_name def_name #> Binding.reset_pos; fun def_binding_optional b name = if Binding.is_empty name then def_binding b else name; fun make_def_binding cond b = if cond then def_binding b else Binding.empty; (* unofficial theorem names *) fun has_name_hint thm = AList.defined (op =) (Thm.get_tags thm) Markup.nameN; fun the_name_hint thm = the (AList.lookup (op =) (Thm.get_tags thm) Markup.nameN); fun get_name_hint thm = if has_name_hint thm then the_name_hint thm else "??.unknown"; fun put_name_hint name = untag_rule Markup.nameN #> tag_rule (Markup.nameN, name); (* theorem kinds *) val theoremK = "theorem"; fun legacy_get_kind thm = the_default "" (Properties.get (Thm.get_tags thm) Markup.kindN); fun kind_rule k = tag_rule (Markup.kindN, k) o untag_rule Markup.kindN; (** attributes **) (*attributes subsume any kind of rules or context modifiers*) type attribute = Context.generic * thm -> Context.generic option * thm option; type binding = binding * attribute list; fun rule_attribute ths f (x, th) = (NONE, (case find_first is_free_dummy (th :: ths) of SOME th' => SOME th' | NONE => SOME (f x th))); fun declaration_attribute f (x, th) = (if is_free_dummy th then NONE else SOME (f th x), NONE); fun mixed_attribute f (x, th) = let val (x', th') = f (x, th) in (SOME x', SOME th') end; fun apply_attribute (att: attribute) th x = let val (x', th') = att (x, check_hyps x (Thm.transfer'' x th)) in (the_default th th', the_default x x') end; fun attribute_declaration att th x = #2 (apply_attribute att th x); fun apply_attributes mk dest = let fun app [] th x = (th, x) | app (att :: atts) th x = apply_attribute att th (mk x) ||> dest |-> app atts; in app end; val theory_attributes = apply_attributes Context.Theory Context.the_theory; val proof_attributes = apply_attributes Context.Proof Context.the_proof; fun no_attributes x = (x, []); fun simple_fact x = [(x, [])]; fun tag tg = rule_attribute [] (K (tag_rule tg)); fun untag s = rule_attribute [] (K (untag_rule s)); fun kind k = rule_attribute [] (K (k <> "" ? kind_rule k)); -(** proof terms **) - -fun reconstruct_proof_of thm = - Proofterm.reconstruct_proof (Thm.theory_of_thm thm) (Thm.prop_of thm) (Thm.proof_of thm); - -fun standard_proof_of {full, expand_name} thm = - let val thy = Thm.theory_of_thm thm in - reconstruct_proof_of thm - |> Proofterm.expand_proof thy expand_name - |> Proofterm.rew_proof thy - |> Proofterm.no_thm_proofs - |> not full ? Proofterm.shrink_proof - end; - - - (** forked proofs **) structure Proofs = Theory_Data ( type T = thm list lazy list; val empty = []; fun extend _ = empty; fun merge _ = empty; ); fun register_proofs ths = (Proofs.map o cons) (Lazy.map_finished (map Thm.trim_context) ths); fun force_proofs thy = rev (Proofs.get thy) |> maps (map (Thm.transfer thy) o Lazy.force); val consolidate_theory = Thm.consolidate o force_proofs; fun expose_theory thy = if Proofterm.export_enabled () then Thm.expose_proofs thy (force_proofs thy) else (); (** print theorems **) (* options *) val show_consts = Config.declare_option_bool ("show_consts", \<^here>); val show_hyps = Config.declare_bool ("show_hyps", \<^here>) (K false); val show_tags = Config.declare_bool ("show_tags", \<^here>) (K false); (* pretty_thm etc. *) fun pretty_tag (name, arg) = Pretty.strs [name, quote arg]; val pretty_tags = Pretty.list "[" "]" o map pretty_tag; fun pretty_thm_raw ctxt {quote, show_hyps = show_hyps'} raw_th = let val show_tags = Config.get ctxt show_tags; val show_hyps = Config.get ctxt show_hyps; val th = raw_th |> perhaps (try (Thm.transfer' ctxt)) |> perhaps (try Thm.strip_shyps); val hyps = if show_hyps then Thm.hyps_of th else undeclared_hyps (Context.Proof ctxt) th; val extra_shyps = extra_shyps' ctxt th; val tags = Thm.get_tags th; val tpairs = Thm.tpairs_of th; val q = if quote then Pretty.quote else I; val prt_term = q o Syntax.pretty_term ctxt; val hlen = length extra_shyps + length hyps + length tpairs; val hsymbs = if hlen = 0 then [] else if show_hyps orelse show_hyps' then [Pretty.brk 2, Pretty.list "[" "]" (map (q o Syntax.pretty_flexpair ctxt) tpairs @ map prt_term hyps @ map (Syntax.pretty_sort ctxt) extra_shyps)] else [Pretty.brk 2, Pretty.str ("[" ^ replicate_string hlen "." ^ "]")]; val tsymbs = if null tags orelse not show_tags then [] else [Pretty.brk 1, pretty_tags tags]; in Pretty.block (prt_term (Thm.prop_of th) :: (hsymbs @ tsymbs)) end; fun pretty_thm ctxt = pretty_thm_raw ctxt {quote = false, show_hyps = true}; fun pretty_thm_item ctxt th = Pretty.item [pretty_thm ctxt th]; fun pretty_thm_global thy = pretty_thm_raw (Syntax.init_pretty_global thy) {quote = false, show_hyps = false}; val string_of_thm = Pretty.string_of oo pretty_thm; val string_of_thm_global = Pretty.string_of oo pretty_thm_global; open Thm; end; structure Basic_Thm: BASIC_THM = Thm; open Basic_Thm; diff --git a/src/Pure/thm.ML b/src/Pure/thm.ML --- a/src/Pure/thm.ML +++ b/src/Pure/thm.ML @@ -1,2341 +1,2345 @@ (* Title: Pure/thm.ML Author: Lawrence C Paulson, Cambridge University Computer Laboratory Author: Makarius The very core of Isabelle's Meta Logic: certified types and terms, derivations, theorems, inference rules (including lifting and resolution), oracles. *) infix 0 RS RSN; signature BASIC_THM = sig type ctyp type cterm exception CTERM of string * cterm list type thm type conv = cterm -> thm exception THM of string * int * thm list val RSN: thm * (int * thm) -> thm val RS: thm * thm -> thm end; signature THM = sig include BASIC_THM (*certified types*) val typ_of: ctyp -> typ val global_ctyp_of: theory -> typ -> ctyp val ctyp_of: Proof.context -> typ -> ctyp val dest_ctyp: ctyp -> ctyp list val dest_ctypN: int -> ctyp -> ctyp val dest_ctyp0: ctyp -> ctyp val dest_ctyp1: ctyp -> ctyp val make_ctyp: ctyp -> ctyp list -> ctyp (*certified terms*) val term_of: cterm -> term val typ_of_cterm: cterm -> typ val ctyp_of_cterm: cterm -> ctyp val maxidx_of_cterm: cterm -> int val global_cterm_of: theory -> term -> cterm val cterm_of: Proof.context -> term -> cterm val renamed_term: term -> cterm -> cterm val dest_comb: cterm -> cterm * cterm val dest_fun: cterm -> cterm val dest_arg: cterm -> cterm val dest_fun2: cterm -> cterm val dest_arg1: cterm -> cterm val dest_abs: string option -> cterm -> cterm * cterm val rename_tvar: indexname -> ctyp -> ctyp val var: indexname * ctyp -> cterm val apply: cterm -> cterm -> cterm val lambda_name: string * cterm -> cterm -> cterm val lambda: cterm -> cterm -> cterm val adjust_maxidx_cterm: int -> cterm -> cterm val incr_indexes_cterm: int -> cterm -> cterm val match: cterm * cterm -> ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list val first_order_match: cterm * cterm -> ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list (*theorems*) val fold_terms: (term -> 'a -> 'a) -> thm -> 'a -> 'a val fold_atomic_ctyps: (ctyp -> 'a -> 'a) -> thm -> 'a -> 'a val fold_atomic_cterms: (cterm -> 'a -> 'a) -> thm -> 'a -> 'a val terms_of_tpairs: (term * term) list -> term list val full_prop_of: thm -> term val theory_id: thm -> Context.theory_id val theory_name: thm -> string val maxidx_of: thm -> int val maxidx_thm: thm -> int -> int val shyps_of: thm -> sort Ord_List.T val hyps_of: thm -> term list val prop_of: thm -> term val tpairs_of: thm -> (term * term) list val concl_of: thm -> term val prems_of: thm -> term list val nprems_of: thm -> int val no_prems: thm -> bool val major_prem_of: thm -> term val cprop_of: thm -> cterm val cprem_of: thm -> int -> cterm val cconcl_of: thm -> cterm val cprems_of: thm -> cterm list val chyps_of: thm -> cterm list exception CONTEXT of string * ctyp list * cterm list * thm list * Context.generic option val theory_of_cterm: cterm -> theory val theory_of_thm: thm -> theory val trim_context_ctyp: ctyp -> ctyp val trim_context_cterm: cterm -> cterm val transfer_ctyp: theory -> ctyp -> ctyp val transfer_cterm: theory -> cterm -> cterm val transfer: theory -> thm -> thm val transfer': Proof.context -> thm -> thm val transfer'': Context.generic -> thm -> thm val join_transfer: theory -> thm -> thm val join_transfer_context: Proof.context * thm -> Proof.context * thm val renamed_prop: term -> thm -> thm val weaken: cterm -> thm -> thm val weaken_sorts: sort list -> cterm -> cterm val extra_shyps: thm -> sort list val proof_bodies_of: thm list -> proof_body list val proof_body_of: thm -> proof_body val proof_of: thm -> proof + val reconstruct_proof_of: thm -> Proofterm.proof val consolidate: thm list -> unit val expose_proofs: theory -> thm list -> unit val expose_proof: theory -> thm -> unit val future: thm future -> cterm -> thm val thm_deps: thm -> Proofterm.thm Ord_List.T val derivation_closed: thm -> bool val derivation_name: thm -> string val derivation_id: thm -> Proofterm.thm_id option val raw_derivation_name: thm -> string val expand_name: thm -> Proofterm.thm_header -> string option val name_derivation: string * Position.T -> thm -> thm val close_derivation: Position.T -> thm -> thm val trim_context: thm -> thm val axiom: theory -> string -> thm val all_axioms_of: theory -> (string * thm) list val get_tags: thm -> Properties.T val map_tags: (Properties.T -> Properties.T) -> thm -> thm val norm_proof: thm -> thm val adjust_maxidx_thm: int -> thm -> thm (*type classes*) val the_classrel: theory -> class * class -> thm val the_arity: theory -> string * sort list * class -> thm val classrel_proof: theory -> class * class -> proof val arity_proof: theory -> string * sort list * class -> proof (*oracles*) val add_oracle: binding * ('a -> cterm) -> theory -> (string * ('a -> thm)) * theory val oracle_space: theory -> Name_Space.T val pretty_oracle: Proof.context -> string -> Pretty.T val extern_oracles: bool -> Proof.context -> (Markup.T * xstring) list val check_oracle: Proof.context -> xstring * Position.T -> string (*inference rules*) val assume: cterm -> thm val implies_intr: cterm -> thm -> thm val implies_elim: thm -> thm -> thm val forall_intr: cterm -> thm -> thm val forall_elim: cterm -> thm -> thm val reflexive: cterm -> thm val symmetric: thm -> thm val transitive: thm -> thm -> thm val beta_conversion: bool -> conv val eta_conversion: conv val eta_long_conversion: conv val abstract_rule: string -> cterm -> thm -> thm val combination: thm -> thm -> thm val equal_intr: thm -> thm -> thm val equal_elim: thm -> thm -> thm val solve_constraints: thm -> thm val flexflex_rule: Proof.context option -> thm -> thm Seq.seq val generalize: string list * string list -> int -> thm -> thm val generalize_cterm: string list * string list -> int -> cterm -> cterm val generalize_ctyp: string list -> int -> ctyp -> ctyp val instantiate: ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list -> thm -> thm val instantiate_cterm: ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list -> cterm -> cterm val trivial: cterm -> thm val of_class: ctyp * class -> thm val strip_shyps: thm -> thm val unconstrainT: thm -> thm val varifyT_global': (string * sort) list -> thm -> ((string * sort) * indexname) list * thm val varifyT_global: thm -> thm val legacy_freezeT: thm -> thm val plain_prop_of: thm -> term val dest_state: thm * int -> (term * term) list * term list * term * term val lift_rule: cterm -> thm -> thm val incr_indexes: int -> thm -> thm val assumption: Proof.context option -> int -> thm -> thm Seq.seq val eq_assumption: int -> thm -> thm val rotate_rule: int -> int -> thm -> thm val permute_prems: int -> int -> thm -> thm val bicompose: Proof.context option -> {flatten: bool, match: bool, incremented: bool} -> bool * thm * int -> int -> thm -> thm Seq.seq val biresolution: Proof.context option -> bool -> (bool * thm) list -> int -> thm -> thm Seq.seq val thynames_of_arity: theory -> string * class -> string list val add_classrel: thm -> theory -> theory val add_arity: thm -> theory -> theory end; structure Thm: THM = struct (*** Certified terms and types ***) (** certified types **) datatype ctyp = Ctyp of {cert: Context.certificate, T: typ, maxidx: int, sorts: sort Ord_List.T}; fun typ_of (Ctyp {T, ...}) = T; fun global_ctyp_of thy raw_T = let val T = Sign.certify_typ thy raw_T; val maxidx = Term.maxidx_of_typ T; val sorts = Sorts.insert_typ T []; in Ctyp {cert = Context.Certificate thy, T = T, maxidx = maxidx, sorts = sorts} end; val ctyp_of = global_ctyp_of o Proof_Context.theory_of; fun dest_ctyp (Ctyp {cert, T = Type (_, Ts), maxidx, sorts}) = map (fn T => Ctyp {cert = cert, T = T, maxidx = maxidx, sorts = sorts}) Ts | dest_ctyp cT = raise TYPE ("dest_ctyp", [typ_of cT], []); fun dest_ctypN n (Ctyp {cert, T, maxidx, sorts}) = let fun err () = raise TYPE ("dest_ctypN", [T], []) in (case T of Type (_, Ts) => Ctyp {cert = cert, T = nth Ts n handle General.Subscript => err (), maxidx = maxidx, sorts = sorts} | _ => err ()) end; val dest_ctyp0 = dest_ctypN 0; val dest_ctyp1 = dest_ctypN 1; fun join_certificate_ctyp (Ctyp {cert, ...}) cert0 = Context.join_certificate (cert0, cert); fun union_sorts_ctyp (Ctyp {sorts, ...}) sorts0 = Sorts.union sorts0 sorts; fun maxidx_ctyp (Ctyp {maxidx, ...}) maxidx0 = Int.max (maxidx0, maxidx); fun make_ctyp (Ctyp {cert, T, maxidx = _, sorts = _}) cargs = let val As = map typ_of cargs; fun err () = raise TYPE ("make_ctyp", T :: As, []); in (case T of Type (a, args) => Ctyp { cert = fold join_certificate_ctyp cargs cert, maxidx = fold maxidx_ctyp cargs ~1, sorts = fold union_sorts_ctyp cargs [], T = if length args = length cargs then Type (a, As) else err ()} | _ => err ()) end; (** certified terms **) (*certified terms with checked typ, maxidx, and sorts*) datatype cterm = Cterm of {cert: Context.certificate, t: term, T: typ, maxidx: int, sorts: sort Ord_List.T}; exception CTERM of string * cterm list; fun term_of (Cterm {t, ...}) = t; fun typ_of_cterm (Cterm {T, ...}) = T; fun ctyp_of_cterm (Cterm {cert, T, maxidx, sorts, ...}) = Ctyp {cert = cert, T = T, maxidx = maxidx, sorts = sorts}; fun maxidx_of_cterm (Cterm {maxidx, ...}) = maxidx; fun global_cterm_of thy tm = let val (t, T, maxidx) = Sign.certify_term thy tm; val sorts = Sorts.insert_term t []; in Cterm {cert = Context.Certificate thy, t = t, T = T, maxidx = maxidx, sorts = sorts} end; val cterm_of = global_cterm_of o Proof_Context.theory_of; fun join_certificate0 (Cterm {cert = cert1, ...}, Cterm {cert = cert2, ...}) = Context.join_certificate (cert1, cert2); fun renamed_term t' (Cterm {cert, t, T, maxidx, sorts}) = if t aconv t' then Cterm {cert = cert, t = t', T = T, maxidx = maxidx, sorts = sorts} else raise TERM ("renamed_term: terms disagree", [t, t']); (* destructors *) fun dest_comb (Cterm {t = c $ a, T, cert, maxidx, sorts}) = let val A = Term.argument_type_of c 0 in (Cterm {t = c, T = A --> T, cert = cert, maxidx = maxidx, sorts = sorts}, Cterm {t = a, T = A, cert = cert, maxidx = maxidx, sorts = sorts}) end | dest_comb ct = raise CTERM ("dest_comb", [ct]); fun dest_fun (Cterm {t = c $ _, T, cert, maxidx, sorts}) = let val A = Term.argument_type_of c 0 in Cterm {t = c, T = A --> T, cert = cert, maxidx = maxidx, sorts = sorts} end | dest_fun ct = raise CTERM ("dest_fun", [ct]); fun dest_arg (Cterm {t = c $ a, T = _, cert, maxidx, sorts}) = let val A = Term.argument_type_of c 0 in Cterm {t = a, T = A, cert = cert, maxidx = maxidx, sorts = sorts} end | dest_arg ct = raise CTERM ("dest_arg", [ct]); fun dest_fun2 (Cterm {t = c $ _ $ _, T, cert, maxidx, sorts}) = let val A = Term.argument_type_of c 0; val B = Term.argument_type_of c 1; in Cterm {t = c, T = A --> B --> T, cert = cert, maxidx = maxidx, sorts = sorts} end | dest_fun2 ct = raise CTERM ("dest_fun2", [ct]); fun dest_arg1 (Cterm {t = c $ a $ _, T = _, cert, maxidx, sorts}) = let val A = Term.argument_type_of c 0 in Cterm {t = a, T = A, cert = cert, maxidx = maxidx, sorts = sorts} end | dest_arg1 ct = raise CTERM ("dest_arg1", [ct]); fun dest_abs a (Cterm {t = Abs (x, T, t), T = Type ("fun", [_, U]), cert, maxidx, sorts}) = let val (y', t') = Term.dest_abs (the_default x a, T, t) in (Cterm {t = Free (y', T), T = T, cert = cert, maxidx = maxidx, sorts = sorts}, Cterm {t = t', T = U, cert = cert, maxidx = maxidx, sorts = sorts}) end | dest_abs _ ct = raise CTERM ("dest_abs", [ct]); (* constructors *) fun rename_tvar (a, i) (Ctyp {cert, T, maxidx, sorts}) = let val S = (case T of TFree (_, S) => S | TVar (_, S) => S | _ => raise TYPE ("rename_tvar: no variable", [T], [])); val _ = if i < 0 then raise TYPE ("rename_tvar: bad index", [TVar ((a, i), S)], []) else (); in Ctyp {cert = cert, T = TVar ((a, i), S), maxidx = Int.max (i, maxidx), sorts = sorts} end; fun var ((x, i), Ctyp {cert, T, maxidx, sorts}) = if i < 0 then raise TERM ("var: bad index", [Var ((x, i), T)]) else Cterm {cert = cert, t = Var ((x, i), T), T = T, maxidx = Int.max (i, maxidx), sorts = sorts}; fun apply (cf as Cterm {t = f, T = Type ("fun", [dty, rty]), maxidx = maxidx1, sorts = sorts1, ...}) (cx as Cterm {t = x, T, maxidx = maxidx2, sorts = sorts2, ...}) = if T = dty then Cterm {cert = join_certificate0 (cf, cx), t = f $ x, T = rty, maxidx = Int.max (maxidx1, maxidx2), sorts = Sorts.union sorts1 sorts2} else raise CTERM ("apply: types don't agree", [cf, cx]) | apply cf cx = raise CTERM ("apply: first arg is not a function", [cf, cx]); fun lambda_name (x, ct1 as Cterm {t = t1, T = T1, maxidx = maxidx1, sorts = sorts1, ...}) (ct2 as Cterm {t = t2, T = T2, maxidx = maxidx2, sorts = sorts2, ...}) = let val t = Term.lambda_name (x, t1) t2 in Cterm {cert = join_certificate0 (ct1, ct2), t = t, T = T1 --> T2, maxidx = Int.max (maxidx1, maxidx2), sorts = Sorts.union sorts1 sorts2} end; fun lambda t u = lambda_name ("", t) u; (* indexes *) fun adjust_maxidx_cterm i (ct as Cterm {cert, t, T, maxidx, sorts}) = if maxidx = i then ct else if maxidx < i then Cterm {maxidx = i, cert = cert, t = t, T = T, sorts = sorts} else Cterm {maxidx = Int.max (maxidx_of_term t, i), cert = cert, t = t, T = T, sorts = sorts}; fun incr_indexes_cterm i (ct as Cterm {cert, t, T, maxidx, sorts}) = if i < 0 then raise CTERM ("negative increment", [ct]) else if i = 0 then ct else Cterm {cert = cert, t = Logic.incr_indexes ([], [], i) t, T = Logic.incr_tvar i T, maxidx = maxidx + i, sorts = sorts}; (*** Derivations and Theorems ***) (* sort constraints *) type constraint = {theory: theory, typ: typ, sort: sort}; local val constraint_ord : constraint ord = Context.theory_id_ord o apply2 (Context.theory_id o #theory) <<< Term_Ord.typ_ord o apply2 #typ <<< Term_Ord.sort_ord o apply2 #sort; val smash_atyps = map_atyps (fn TVar (_, S) => Term.aT S | TFree (_, S) => Term.aT S | T => T); in val union_constraints = Ord_List.union constraint_ord; fun insert_constraints thy (T, S) = let val ignored = S = [] orelse (case T of TFree (_, S') => S = S' | TVar (_, S') => S = S' | _ => false); in if ignored then I else Ord_List.insert constraint_ord {theory = thy, typ = smash_atyps T, sort = S} end; fun insert_constraints_env thy env = let val tyenv = Envir.type_env env; fun insert ([], _) = I | insert (S, T) = insert_constraints thy (Envir.norm_type tyenv T, S); in tyenv |> Vartab.fold (insert o #2) end; end; (* datatype thm *) datatype thm = Thm of deriv * (*derivation*) {cert: Context.certificate, (*background theory certificate*) tags: Properties.T, (*additional annotations/comments*) maxidx: int, (*maximum index of any Var or TVar*) constraints: constraint Ord_List.T, (*implicit proof obligations for sort constraints*) shyps: sort Ord_List.T, (*sort hypotheses*) hyps: term Ord_List.T, (*hypotheses*) tpairs: (term * term) list, (*flex-flex pairs*) prop: term} (*conclusion*) and deriv = Deriv of {promises: (serial * thm future) Ord_List.T, body: Proofterm.proof_body}; type conv = cterm -> thm; (*errors involving theorems*) exception THM of string * int * thm list; fun rep_thm (Thm (_, args)) = args; fun fold_terms f (Thm (_, {tpairs, prop, hyps, ...})) = fold (fn (t, u) => f t #> f u) tpairs #> f prop #> fold f hyps; fun fold_atomic_ctyps f (th as Thm (_, {cert, maxidx, shyps, ...})) = let fun ctyp T = Ctyp {cert = cert, T = T, maxidx = maxidx, sorts = shyps} in (fold_terms o fold_types o fold_atyps) (f o ctyp) th end; fun fold_atomic_cterms f (th as Thm (_, {cert, maxidx, shyps, ...})) = let fun cterm t T = Cterm {cert = cert, t = t, T = T, maxidx = maxidx, sorts = shyps} in (fold_terms o fold_aterms) (fn t as Const (_, T) => f (cterm t T) | t as Free (_, T) => f (cterm t T) | t as Var (_, T) => f (cterm t T) | _ => I) th end; fun terms_of_tpairs tpairs = fold_rev (fn (t, u) => cons t o cons u) tpairs []; fun eq_tpairs ((t, u), (t', u')) = t aconv t' andalso u aconv u'; fun union_tpairs ts us = Library.merge eq_tpairs (ts, us); val maxidx_tpairs = fold (fn (t, u) => Term.maxidx_term t #> Term.maxidx_term u); fun attach_tpairs tpairs prop = Logic.list_implies (map Logic.mk_equals tpairs, prop); fun full_prop_of (Thm (_, {tpairs, prop, ...})) = attach_tpairs tpairs prop; val union_hyps = Ord_List.union Term_Ord.fast_term_ord; val insert_hyps = Ord_List.insert Term_Ord.fast_term_ord; val remove_hyps = Ord_List.remove Term_Ord.fast_term_ord; fun join_certificate1 (Cterm {cert = cert1, ...}, Thm (_, {cert = cert2, ...})) = Context.join_certificate (cert1, cert2); fun join_certificate2 (Thm (_, {cert = cert1, ...}), Thm (_, {cert = cert2, ...})) = Context.join_certificate (cert1, cert2); (* basic components *) val cert_of = #cert o rep_thm; val theory_id = Context.certificate_theory_id o cert_of; val theory_name = Context.theory_id_name o theory_id; val maxidx_of = #maxidx o rep_thm; fun maxidx_thm th i = Int.max (maxidx_of th, i); val shyps_of = #shyps o rep_thm; val hyps_of = #hyps o rep_thm; val prop_of = #prop o rep_thm; val tpairs_of = #tpairs o rep_thm; val concl_of = Logic.strip_imp_concl o prop_of; val prems_of = Logic.strip_imp_prems o prop_of; val nprems_of = Logic.count_prems o prop_of; fun no_prems th = nprems_of th = 0; fun major_prem_of th = (case prems_of th of prem :: _ => Logic.strip_assums_concl prem | [] => raise THM ("major_prem_of: rule with no premises", 0, [th])); fun cprop_of (Thm (_, {cert, maxidx, shyps, prop, ...})) = Cterm {cert = cert, maxidx = maxidx, T = propT, t = prop, sorts = shyps}; fun cprem_of (th as Thm (_, {cert, maxidx, shyps, prop, ...})) i = Cterm {cert = cert, maxidx = maxidx, T = propT, sorts = shyps, t = Logic.nth_prem (i, prop) handle TERM _ => raise THM ("cprem_of", i, [th])}; fun cconcl_of (th as Thm (_, {cert, maxidx, shyps, ...})) = Cterm {cert = cert, maxidx = maxidx, T = propT, sorts = shyps, t = concl_of th}; fun cprems_of (th as Thm (_, {cert, maxidx, shyps, ...})) = map (fn t => Cterm {cert = cert, maxidx = maxidx, T = propT, sorts = shyps, t = t}) (prems_of th); fun chyps_of (Thm (_, {cert, shyps, hyps, ...})) = map (fn t => Cterm {cert = cert, maxidx = ~1, T = propT, sorts = shyps, t = t}) hyps; (* implicit theory context *) exception CONTEXT of string * ctyp list * cterm list * thm list * Context.generic option; fun theory_of_cterm (ct as Cterm {cert, ...}) = Context.certificate_theory cert handle ERROR msg => raise CONTEXT (msg, [], [ct], [], NONE); fun theory_of_thm th = Context.certificate_theory (cert_of th) handle ERROR msg => raise CONTEXT (msg, [], [], [th], NONE); fun trim_context_ctyp cT = (case cT of Ctyp {cert = Context.Certificate_Id _, ...} => cT | Ctyp {cert = Context.Certificate thy, T, maxidx, sorts} => Ctyp {cert = Context.Certificate_Id (Context.theory_id thy), T = T, maxidx = maxidx, sorts = sorts}); fun trim_context_cterm ct = (case ct of Cterm {cert = Context.Certificate_Id _, ...} => ct | Cterm {cert = Context.Certificate thy, t, T, maxidx, sorts} => Cterm {cert = Context.Certificate_Id (Context.theory_id thy), t = t, T = T, maxidx = maxidx, sorts = sorts}); fun trim_context_thm th = (case th of Thm (_, {constraints = _ :: _, ...}) => raise THM ("trim_context: pending sort constraints", 0, [th]) | Thm (_, {cert = Context.Certificate_Id _, ...}) => th | Thm (der, {cert = Context.Certificate thy, tags, maxidx, constraints = [], shyps, hyps, tpairs, prop}) => Thm (der, {cert = Context.Certificate_Id (Context.theory_id thy), tags = tags, maxidx = maxidx, constraints = [], shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop})); fun transfer_ctyp thy' cT = let val Ctyp {cert, T, maxidx, sorts} = cT; val _ = Context.subthy_id (Context.certificate_theory_id cert, Context.theory_id thy') orelse raise CONTEXT ("Cannot transfer: not a super theory", [cT], [], [], SOME (Context.Theory thy')); val cert' = Context.join_certificate (Context.Certificate thy', cert); in if Context.eq_certificate (cert, cert') then cT else Ctyp {cert = cert', T = T, maxidx = maxidx, sorts = sorts} end; fun transfer_cterm thy' ct = let val Cterm {cert, t, T, maxidx, sorts} = ct; val _ = Context.subthy_id (Context.certificate_theory_id cert, Context.theory_id thy') orelse raise CONTEXT ("Cannot transfer: not a super theory", [], [ct], [], SOME (Context.Theory thy')); val cert' = Context.join_certificate (Context.Certificate thy', cert); in if Context.eq_certificate (cert, cert') then ct else Cterm {cert = cert', t = t, T = T, maxidx = maxidx, sorts = sorts} end; fun transfer thy' th = let val Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop}) = th; val _ = Context.subthy_id (Context.certificate_theory_id cert, Context.theory_id thy') orelse raise CONTEXT ("Cannot transfer: not a super theory", [], [], [th], SOME (Context.Theory thy')); val cert' = Context.join_certificate (Context.Certificate thy', cert); in if Context.eq_certificate (cert, cert') then th else Thm (der, {cert = cert', tags = tags, maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop}) end; val transfer' = transfer o Proof_Context.theory_of; val transfer'' = transfer o Context.theory_of; fun join_transfer thy th = if Context.subthy_id (Context.theory_id thy, theory_id th) then th else transfer thy th; fun join_transfer_context (ctxt, th) = if Context.subthy_id (Context.theory_id (Proof_Context.theory_of ctxt), theory_id th) then (Context.raw_transfer (theory_of_thm th) ctxt, th) else (ctxt, transfer' ctxt th); (* matching *) local fun gen_match match (ct1 as Cterm {t = t1, sorts = sorts1, ...}, ct2 as Cterm {t = t2, sorts = sorts2, maxidx = maxidx2, ...}) = let val cert = join_certificate0 (ct1, ct2); val thy = Context.certificate_theory cert handle ERROR msg => raise CONTEXT (msg, [], [ct1, ct2], [], NONE); val (Tinsts, tinsts) = match thy (t1, t2) (Vartab.empty, Vartab.empty); val sorts = Sorts.union sorts1 sorts2; fun mk_cTinst ((a, i), (S, T)) = (((a, i), S), Ctyp {T = T, cert = cert, maxidx = maxidx2, sorts = sorts}); fun mk_ctinst ((x, i), (U, t)) = let val T = Envir.subst_type Tinsts U in (((x, i), T), Cterm {t = t, T = T, cert = cert, maxidx = maxidx2, sorts = sorts}) end; in (Vartab.fold (cons o mk_cTinst) Tinsts [], Vartab.fold (cons o mk_ctinst) tinsts []) end; in val match = gen_match Pattern.match; val first_order_match = gen_match Pattern.first_order_match; end; (*implicit alpha-conversion*) fun renamed_prop prop' (Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop})) = if prop aconv prop' then Thm (der, {cert = cert, tags = tags, maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop'}) else raise TERM ("renamed_prop: props disagree", [prop, prop']); fun make_context ths NONE cert = (Context.Theory (Context.certificate_theory cert) handle ERROR msg => raise CONTEXT (msg, [], [], ths, NONE)) | make_context ths (SOME ctxt) cert = let val thy_id = Context.certificate_theory_id cert; val thy_id' = Context.theory_id (Proof_Context.theory_of ctxt); in if Context.subthy_id (thy_id, thy_id') then Context.Proof ctxt else raise CONTEXT ("Bad context", [], [], ths, SOME (Context.Proof ctxt)) end; fun make_context_certificate ths opt_ctxt cert = let val context = make_context ths opt_ctxt cert; val cert' = Context.Certificate (Context.theory_of context); in (context, cert') end; (*explicit weakening: maps |- B to A |- B*) fun weaken raw_ct th = let val ct as Cterm {t = A, T, sorts, maxidx = maxidxA, ...} = adjust_maxidx_cterm ~1 raw_ct; val Thm (der, {tags, maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = th; in if T <> propT then raise THM ("weaken: assumptions must have type prop", 0, []) else if maxidxA <> ~1 then raise THM ("weaken: assumptions may not contain schematic variables", maxidxA, []) else Thm (der, {cert = join_certificate1 (ct, th), tags = tags, maxidx = maxidx, constraints = constraints, shyps = Sorts.union sorts shyps, hyps = insert_hyps A hyps, tpairs = tpairs, prop = prop}) end; fun weaken_sorts raw_sorts ct = let val Cterm {cert, t, T, maxidx, sorts} = ct; val thy = theory_of_cterm ct; val more_sorts = Sorts.make (map (Sign.certify_sort thy) raw_sorts); val sorts' = Sorts.union sorts more_sorts; in Cterm {cert = cert, t = t, T = T, maxidx = maxidx, sorts = sorts'} end; (*dangling sort constraints of a thm*) fun extra_shyps (th as Thm (_, {shyps, ...})) = Sorts.subtract (fold_terms Sorts.insert_term th []) shyps; (** derivations and promised proofs **) fun make_deriv promises oracles thms proof = Deriv {promises = promises, body = PBody {oracles = oracles, thms = thms, proof = proof}}; val empty_deriv = make_deriv [] [] [] MinProof; (* inference rules *) val promise_ord: (serial * thm future) ord = fn ((i, _), (j, _)) => int_ord (j, i); fun bad_proofs i = error ("Illegal level of detail for proof objects: " ^ string_of_int i); fun deriv_rule2 f (Deriv {promises = ps1, body = PBody {oracles = oracles1, thms = thms1, proof = prf1}}) (Deriv {promises = ps2, body = PBody {oracles = oracles2, thms = thms2, proof = prf2}}) = let val ps = Ord_List.union promise_ord ps1 ps2; val oracles = Proofterm.unions_oracles [oracles1, oracles2]; val thms = Proofterm.unions_thms [thms1, thms2]; val prf = (case ! Proofterm.proofs of 2 => f prf1 prf2 | 1 => MinProof | 0 => MinProof | i => bad_proofs i); in make_deriv ps oracles thms prf end; fun deriv_rule1 f = deriv_rule2 (K f) empty_deriv; fun deriv_rule0 make_prf = if ! Proofterm.proofs <= 1 then empty_deriv else deriv_rule1 I (make_deriv [] [] [] (make_prf ())); fun deriv_rule_unconditional f (Deriv {promises, body = PBody {oracles, thms, proof}}) = make_deriv promises oracles thms (f proof); (* fulfilled proofs *) fun raw_promises_of (Thm (Deriv {promises, ...}, _)) = promises; fun join_promises [] = () | join_promises promises = join_promises_of (Future.joins (map snd promises)) and join_promises_of thms = join_promises (Ord_List.make promise_ord (maps raw_promises_of thms)); fun fulfill_body (th as Thm (Deriv {promises, body}, _)) = let val fulfilled_promises = map #1 promises ~~ map fulfill_body (Future.joins (map #2 promises)) in Proofterm.fulfill_norm_proof (theory_of_thm th) fulfilled_promises body end; fun proof_bodies_of thms = (join_promises_of thms; map fulfill_body thms); val proof_body_of = singleton proof_bodies_of; val proof_of = Proofterm.proof_of o proof_body_of; +fun reconstruct_proof_of thm = + Proofterm.reconstruct_proof (theory_of_thm thm) (prop_of thm) (proof_of thm); + val consolidate = ignore o proof_bodies_of; fun expose_proofs thy thms = if Proofterm.export_proof_boxes_required thy then Proofterm.export_proof_boxes (proof_bodies_of (map (transfer thy) thms)) else (); fun expose_proof thy = expose_proofs thy o single; (* future rule *) fun future_result i orig_cert orig_shyps orig_prop thm = let fun err msg = raise THM ("future_result: " ^ msg, 0, [thm]); val Thm (Deriv {promises, ...}, {cert, constraints, shyps, hyps, tpairs, prop, ...}) = thm; val _ = Context.eq_certificate (cert, orig_cert) orelse err "bad theory"; val _ = prop aconv orig_prop orelse err "bad prop"; val _ = null constraints orelse err "bad sort constraints"; val _ = null tpairs orelse err "bad flex-flex constraints"; val _ = null hyps orelse err "bad hyps"; val _ = Sorts.subset (shyps, orig_shyps) orelse err "bad shyps"; val _ = forall (fn (j, _) => i <> j) promises orelse err "bad dependencies"; val _ = join_promises promises; in thm end; fun future future_thm ct = let val Cterm {cert = cert, t = prop, T, maxidx, sorts} = ct; val _ = T <> propT andalso raise CTERM ("future: prop expected", [ct]); val _ = if Proofterm.proofs_enabled () then raise CTERM ("future: proof terms enabled", [ct]) else (); val i = serial (); val future = future_thm |> Future.map (future_result i cert sorts prop); in Thm (make_deriv [(i, future)] [] [] MinProof, {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = prop}) end; (** Axioms **) fun axiom thy name = (case Name_Space.lookup (Theory.axiom_table thy) name of SOME prop => let val der = deriv_rule0 (fn () => Proofterm.axm_proof name prop); val cert = Context.Certificate thy; val maxidx = maxidx_of_term prop; val shyps = Sorts.insert_term prop []; in Thm (der, {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = shyps, hyps = [], tpairs = [], prop = prop}) end | NONE => raise THEORY ("No axiom " ^ quote name, [thy])); fun all_axioms_of thy = map (fn (name, _) => (name, axiom thy name)) (Theory.all_axioms_of thy); (* tags *) val get_tags = #tags o rep_thm; fun map_tags f (Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop})) = Thm (der, {cert = cert, tags = f tags, maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop}); (* technical adjustments *) fun norm_proof (th as Thm (der, args)) = Thm (deriv_rule1 (Proofterm.rew_proof (theory_of_thm th)) der, args); fun adjust_maxidx_thm i (th as Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop})) = if maxidx = i then th else if maxidx < i then Thm (der, {maxidx = i, cert = cert, tags = tags, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop}) else Thm (der, {maxidx = Int.max (maxidx_tpairs tpairs (maxidx_of_term prop), i), cert = cert, tags = tags, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop}); (*** Theory data ***) (* type classes *) structure Aritytab = Table( type key = string * sort list * class; val ord = fast_string_ord o apply2 #1 <<< fast_string_ord o apply2 #3 <<< list_ord Term_Ord.sort_ord o apply2 #2; ); datatype classes = Classes of {classrels: thm Symreltab.table, arities: (thm * string * serial) Aritytab.table}; fun make_classes (classrels, arities) = Classes {classrels = classrels, arities = arities}; val empty_classes = make_classes (Symreltab.empty, Aritytab.empty); (*see Theory.at_begin hook for transitive closure of classrels and arity completion*) fun merge_classes (Classes {classrels = classrels1, arities = arities1}, Classes {classrels = classrels2, arities = arities2}) = let val classrels' = Symreltab.merge (K true) (classrels1, classrels2); val arities' = Aritytab.merge (K true) (arities1, arities2); in make_classes (classrels', arities') end; (* data *) structure Data = Theory_Data ( type T = unit Name_Space.table * (*oracles: authentic derivation names*) classes; (*type classes within the logic*) val empty : T = (Name_Space.empty_table "oracle", empty_classes); val extend = I; fun merge ((oracles1, sorts1), (oracles2, sorts2)) : T = (Name_Space.merge_tables (oracles1, oracles2), merge_classes (sorts1, sorts2)); ); val get_oracles = #1 o Data.get; val map_oracles = Data.map o apfst; val get_classes = (fn (_, Classes args) => args) o Data.get; val get_classrels = #classrels o get_classes; val get_arities = #arities o get_classes; fun map_classes f = (Data.map o apsnd) (fn Classes {classrels, arities} => make_classes (f (classrels, arities))); fun map_classrels f = map_classes (fn (classrels, arities) => (f classrels, arities)); fun map_arities f = map_classes (fn (classrels, arities) => (classrels, f arities)); (* type classes *) fun the_classrel thy (c1, c2) = (case Symreltab.lookup (get_classrels thy) (c1, c2) of SOME thm => transfer thy thm | NONE => error ("Unproven class relation " ^ Syntax.string_of_classrel (Proof_Context.init_global thy) [c1, c2])); fun the_arity thy (a, Ss, c) = (case Aritytab.lookup (get_arities thy) (a, Ss, c) of SOME (thm, _, _) => transfer thy thm | NONE => error ("Unproven type arity " ^ Syntax.string_of_arity (Proof_Context.init_global thy) (a, Ss, [c]))); val classrel_proof = proof_of oo the_classrel; val arity_proof = proof_of oo the_arity; (* solve sort constraints by pro-forma proof *) local fun union_digest (oracles1, thms1) (oracles2, thms2) = (Proofterm.unions_oracles [oracles1, oracles2], Proofterm.unions_thms [thms1, thms2]); fun thm_digest (Thm (Deriv {body = PBody {oracles, thms, ...}, ...}, _)) = (oracles, thms); fun constraint_digest ({theory = thy, typ, sort, ...}: constraint) = Sorts.of_sort_derivation (Sign.classes_of thy) {class_relation = fn _ => fn _ => fn (digest, c1) => fn c2 => if c1 = c2 then ([], []) else union_digest digest (thm_digest (the_classrel thy (c1, c2))), type_constructor = fn (a, _) => fn dom => fn c => let val arity_digest = thm_digest (the_arity thy (a, (map o map) #2 dom, c)) in (fold o fold) (union_digest o #1) dom arity_digest end, type_variable = fn T => map (pair ([], [])) (Type.sort_of_atyp T)} (typ, sort); in fun solve_constraints (thm as Thm (_, {constraints = [], ...})) = thm | solve_constraints (thm as Thm (der, args)) = let val {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop} = args; val thy = Context.certificate_theory cert; val bad_thys = constraints |> map_filter (fn {theory = thy', ...} => if Context.eq_thy (thy, thy') then NONE else SOME thy'); val () = if null bad_thys then () else raise THEORY ("solve_constraints: bad theories for theorem\n" ^ Syntax.string_of_term_global thy (prop_of thm), thy :: bad_thys); val Deriv {promises, body = PBody {oracles, thms, proof}} = der; val (oracles', thms') = (oracles, thms) |> fold (fold union_digest o constraint_digest) constraints; val body' = PBody {oracles = oracles', thms = thms', proof = proof}; in Thm (Deriv {promises = promises, body = body'}, {constraints = [], cert = cert, tags = tags, maxidx = maxidx, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop}) end; end; (*** Closed theorems with official name ***) (*non-deterministic, depends on unknown promises*) fun derivation_closed (Thm (Deriv {body, ...}, _)) = Proofterm.compact_proof (Proofterm.proof_of body); (*non-deterministic, depends on unknown promises*) fun raw_derivation_name (Thm (Deriv {body, ...}, {shyps, hyps, prop, ...})) = Proofterm.get_approximative_name shyps hyps prop (Proofterm.proof_of body); fun expand_name (Thm (Deriv {body, ...}, {shyps, hyps, prop, ...})) = let val self_id = (case Proofterm.get_identity shyps hyps prop (Proofterm.proof_of body) of NONE => K false | SOME {serial, ...} => fn (header: Proofterm.thm_header) => serial = #serial header); fun expand header = if self_id header orelse #name header = "" then SOME "" else NONE; in expand end; (*deterministic name of finished proof*) fun derivation_name (thm as Thm (_, {shyps, hyps, prop, ...})) = Proofterm.get_approximative_name shyps hyps prop (proof_of thm); (*identified PThm node*) fun derivation_id (thm as Thm (_, {shyps, hyps, prop, ...})) = Proofterm.get_id shyps hyps prop (proof_of thm); (*dependencies of PThm node*) fun thm_deps (thm as Thm (Deriv {promises = [], body = PBody {thms, ...}, ...}, _)) = (case (derivation_id thm, thms) of (SOME {serial = i, ...}, [(j, thm_node)]) => if i = j then Proofterm.thm_node_thms thm_node else thms | _ => thms) | thm_deps thm = raise THM ("thm_deps: bad promises", 0, [thm]); fun name_derivation name_pos = solve_constraints #> (fn thm as Thm (der, args) => let val thy = theory_of_thm thm; val Deriv {promises, body} = der; val {shyps, hyps, prop, tpairs, ...} = args; val _ = null tpairs orelse raise THM ("name_derivation: bad flex-flex constraints", 0, [thm]); val ps = map (apsnd (Future.map fulfill_body)) promises; val (pthm, proof) = Proofterm.thm_proof thy (classrel_proof thy) (arity_proof thy) name_pos shyps hyps prop ps body; val der' = make_deriv [] [] [pthm] proof; in Thm (der', args) end); fun close_derivation pos = solve_constraints #> (fn thm => if not (null (tpairs_of thm)) orelse derivation_closed thm then thm else name_derivation ("", pos) thm); val trim_context = solve_constraints #> trim_context_thm; (*** Oracles ***) fun add_oracle (b, oracle_fn) thy = let val (name, oracles') = Name_Space.define (Context.Theory thy) true (b, ()) (get_oracles thy); val thy' = map_oracles (K oracles') thy; fun invoke_oracle arg = let val Cterm {cert = cert2, t = prop, T, maxidx, sorts} = oracle_fn arg in if T <> propT then raise THM ("Oracle's result must have type prop: " ^ name, 0, []) else let val (oracle, prf) = (case ! Proofterm.proofs of 2 => ((name, SOME prop), Proofterm.oracle_proof name prop) | 1 => ((name, SOME prop), MinProof) | 0 => ((name, NONE), MinProof) | i => bad_proofs i); in Thm (make_deriv [] [oracle] [] prf, {cert = Context.join_certificate (Context.Certificate thy', cert2), tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = prop}) end end; in ((name, invoke_oracle), thy') end; val oracle_space = Name_Space.space_of_table o get_oracles; fun pretty_oracle ctxt = Name_Space.pretty ctxt (oracle_space (Proof_Context.theory_of ctxt)); fun extern_oracles verbose ctxt = map #1 (Name_Space.markup_table verbose ctxt (get_oracles (Proof_Context.theory_of ctxt))); fun check_oracle ctxt = Name_Space.check (Context.Proof ctxt) (get_oracles (Proof_Context.theory_of ctxt)) #> #1; (*** Meta rules ***) (** primitive rules **) (*The assumption rule A |- A*) fun assume raw_ct = let val Cterm {cert, t = prop, T, maxidx, sorts} = adjust_maxidx_cterm ~1 raw_ct in if T <> propT then raise THM ("assume: prop", 0, []) else if maxidx <> ~1 then raise THM ("assume: variables", maxidx, []) else Thm (deriv_rule0 (fn () => Proofterm.Hyp prop), {cert = cert, tags = [], maxidx = ~1, constraints = [], shyps = sorts, hyps = [prop], tpairs = [], prop = prop}) end; (*Implication introduction [A] : B ------- A \ B *) fun implies_intr (ct as Cterm {t = A, T, maxidx = maxidx1, sorts, ...}) (th as Thm (der, {maxidx = maxidx2, hyps, constraints, shyps, tpairs, prop, ...})) = if T <> propT then raise THM ("implies_intr: assumptions must have type prop", 0, [th]) else Thm (deriv_rule1 (Proofterm.implies_intr_proof A) der, {cert = join_certificate1 (ct, th), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = constraints, shyps = Sorts.union sorts shyps, hyps = remove_hyps A hyps, tpairs = tpairs, prop = Logic.mk_implies (A, prop)}); (*Implication elimination A \ B A ------------ B *) fun implies_elim thAB thA = let val Thm (derA, {maxidx = maxidx1, hyps = hypsA, constraints = constraintsA, shyps = shypsA, tpairs = tpairsA, prop = propA, ...}) = thA and Thm (der, {maxidx = maxidx2, hyps, constraints, shyps, tpairs, prop, ...}) = thAB; fun err () = raise THM ("implies_elim: major premise", 0, [thAB, thA]); in (case prop of Const ("Pure.imp", _) $ A $ B => if A aconv propA then Thm (deriv_rule2 (curry Proofterm.%%) der derA, {cert = join_certificate2 (thAB, thA), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = union_constraints constraintsA constraints, shyps = Sorts.union shypsA shyps, hyps = union_hyps hypsA hyps, tpairs = union_tpairs tpairsA tpairs, prop = B}) else err () | _ => err ()) end; (*Forall introduction. The Free or Var x must not be free in the hypotheses. [x] : A ------ \x. A *) fun forall_intr (ct as Cterm {maxidx = maxidx1, t = x, T, sorts, ...}) (th as Thm (der, {maxidx = maxidx2, constraints, shyps, hyps, tpairs, prop, ...})) = let fun result a = Thm (deriv_rule1 (Proofterm.forall_intr_proof (a, x) NONE) der, {cert = join_certificate1 (ct, th), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = constraints, shyps = Sorts.union sorts shyps, hyps = hyps, tpairs = tpairs, prop = Logic.all_const T $ Abs (a, T, abstract_over (x, prop))}); fun check_occs a x ts = if exists (fn t => Logic.occs (x, t)) ts then raise THM ("forall_intr: variable " ^ quote a ^ " free in assumptions", 0, [th]) else (); in (case x of Free (a, _) => (check_occs a x hyps; check_occs a x (terms_of_tpairs tpairs); result a) | Var ((a, _), _) => (check_occs a x (terms_of_tpairs tpairs); result a) | _ => raise THM ("forall_intr: not a variable", 0, [th])) end; (*Forall elimination \x. A ------ A[t/x] *) fun forall_elim (ct as Cterm {t, T, maxidx = maxidx1, sorts, ...}) (th as Thm (der, {maxidx = maxidx2, constraints, shyps, hyps, tpairs, prop, ...})) = (case prop of Const ("Pure.all", Type ("fun", [Type ("fun", [qary, _]), _])) $ A => if T <> qary then raise THM ("forall_elim: type mismatch", 0, [th]) else Thm (deriv_rule1 (Proofterm.% o rpair (SOME t)) der, {cert = join_certificate1 (ct, th), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = constraints, shyps = Sorts.union sorts shyps, hyps = hyps, tpairs = tpairs, prop = Term.betapply (A, t)}) | _ => raise THM ("forall_elim: not quantified", 0, [th])); (* Equality *) (*Reflexivity t \ t *) fun reflexive (Cterm {cert, t, T = _, maxidx, sorts}) = Thm (deriv_rule0 (fn () => Proofterm.reflexive_proof), {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = Logic.mk_equals (t, t)}); (*Symmetry t \ u ------ u \ t *) fun symmetric (th as Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...})) = (case prop of (eq as Const ("Pure.eq", _)) $ t $ u => Thm (deriv_rule1 Proofterm.symmetric_proof der, {cert = cert, tags = [], maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = eq $ u $ t}) | _ => raise THM ("symmetric", 0, [th])); (*Transitivity t1 \ u u \ t2 ------------------ t1 \ t2 *) fun transitive th1 th2 = let val Thm (der1, {maxidx = maxidx1, hyps = hyps1, constraints = constraints1, shyps = shyps1, tpairs = tpairs1, prop = prop1, ...}) = th1 and Thm (der2, {maxidx = maxidx2, hyps = hyps2, constraints = constraints2, shyps = shyps2, tpairs = tpairs2, prop = prop2, ...}) = th2; fun err msg = raise THM ("transitive: " ^ msg, 0, [th1, th2]); in case (prop1, prop2) of ((eq as Const ("Pure.eq", Type (_, [U, _]))) $ t1 $ u, Const ("Pure.eq", _) $ u' $ t2) => if not (u aconv u') then err "middle term" else Thm (deriv_rule2 (Proofterm.transitive_proof U u) der1 der2, {cert = join_certificate2 (th1, th2), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = union_constraints constraints1 constraints2, shyps = Sorts.union shyps1 shyps2, hyps = union_hyps hyps1 hyps2, tpairs = union_tpairs tpairs1 tpairs2, prop = eq $ t1 $ t2}) | _ => err "premises" end; (*Beta-conversion (\x. t) u \ t[u/x] fully beta-reduces the term if full = true *) fun beta_conversion full (Cterm {cert, t, T = _, maxidx, sorts}) = let val t' = if full then Envir.beta_norm t else (case t of Abs (_, _, bodt) $ u => subst_bound (u, bodt) | _ => raise THM ("beta_conversion: not a redex", 0, [])); in Thm (deriv_rule0 (fn () => Proofterm.reflexive_proof), {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = Logic.mk_equals (t, t')}) end; fun eta_conversion (Cterm {cert, t, T = _, maxidx, sorts}) = Thm (deriv_rule0 (fn () => Proofterm.reflexive_proof), {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = Logic.mk_equals (t, Envir.eta_contract t)}); fun eta_long_conversion (Cterm {cert, t, T = _, maxidx, sorts}) = Thm (deriv_rule0 (fn () => Proofterm.reflexive_proof), {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = Logic.mk_equals (t, Envir.eta_long [] t)}); (*The abstraction rule. The Free or Var x must not be free in the hypotheses. The bound variable will be named "a" (since x will be something like x320) t \ u -------------- \x. t \ \x. u *) fun abstract_rule a (Cterm {t = x, T, sorts, ...}) (th as Thm (der, {cert, maxidx, hyps, constraints, shyps, tpairs, prop, ...})) = let val (t, u) = Logic.dest_equals prop handle TERM _ => raise THM ("abstract_rule: premise not an equality", 0, [th]); val result = Thm (deriv_rule1 (Proofterm.abstract_rule_proof (a, x)) der, {cert = cert, tags = [], maxidx = maxidx, constraints = constraints, shyps = Sorts.union sorts shyps, hyps = hyps, tpairs = tpairs, prop = Logic.mk_equals (Abs (a, T, abstract_over (x, t)), Abs (a, T, abstract_over (x, u)))}); fun check_occs a x ts = if exists (fn t => Logic.occs (x, t)) ts then raise THM ("abstract_rule: variable " ^ quote a ^ " free in assumptions", 0, [th]) else (); in (case x of Free (a, _) => (check_occs a x hyps; check_occs a x (terms_of_tpairs tpairs); result) | Var ((a, _), _) => (check_occs a x (terms_of_tpairs tpairs); result) | _ => raise THM ("abstract_rule: not a variable", 0, [th])) end; (*The combination rule f \ g t \ u ------------- f t \ g u *) fun combination th1 th2 = let val Thm (der1, {maxidx = maxidx1, constraints = constraints1, shyps = shyps1, hyps = hyps1, tpairs = tpairs1, prop = prop1, ...}) = th1 and Thm (der2, {maxidx = maxidx2, constraints = constraints2, shyps = shyps2, hyps = hyps2, tpairs = tpairs2, prop = prop2, ...}) = th2; fun chktypes fT tT = (case fT of Type ("fun", [T1, _]) => if T1 <> tT then raise THM ("combination: types", 0, [th1, th2]) else () | _ => raise THM ("combination: not function type", 0, [th1, th2])); in (case (prop1, prop2) of (Const ("Pure.eq", Type ("fun", [fT, _])) $ f $ g, Const ("Pure.eq", Type ("fun", [tT, _])) $ t $ u) => (chktypes fT tT; Thm (deriv_rule2 (Proofterm.combination_proof f g t u) der1 der2, {cert = join_certificate2 (th1, th2), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = union_constraints constraints1 constraints2, shyps = Sorts.union shyps1 shyps2, hyps = union_hyps hyps1 hyps2, tpairs = union_tpairs tpairs1 tpairs2, prop = Logic.mk_equals (f $ t, g $ u)})) | _ => raise THM ("combination: premises", 0, [th1, th2])) end; (*Equality introduction A \ B B \ A ---------------- A \ B *) fun equal_intr th1 th2 = let val Thm (der1, {maxidx = maxidx1, constraints = constraints1, shyps = shyps1, hyps = hyps1, tpairs = tpairs1, prop = prop1, ...}) = th1 and Thm (der2, {maxidx = maxidx2, constraints = constraints2, shyps = shyps2, hyps = hyps2, tpairs = tpairs2, prop = prop2, ...}) = th2; fun err msg = raise THM ("equal_intr: " ^ msg, 0, [th1, th2]); in (case (prop1, prop2) of (Const("Pure.imp", _) $ A $ B, Const("Pure.imp", _) $ B' $ A') => if A aconv A' andalso B aconv B' then Thm (deriv_rule2 (Proofterm.equal_intr_proof A B) der1 der2, {cert = join_certificate2 (th1, th2), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = union_constraints constraints1 constraints2, shyps = Sorts.union shyps1 shyps2, hyps = union_hyps hyps1 hyps2, tpairs = union_tpairs tpairs1 tpairs2, prop = Logic.mk_equals (A, B)}) else err "not equal" | _ => err "premises") end; (*The equal propositions rule A \ B A --------- B *) fun equal_elim th1 th2 = let val Thm (der1, {maxidx = maxidx1, constraints = constraints1, shyps = shyps1, hyps = hyps1, tpairs = tpairs1, prop = prop1, ...}) = th1 and Thm (der2, {maxidx = maxidx2, constraints = constraints2, shyps = shyps2, hyps = hyps2, tpairs = tpairs2, prop = prop2, ...}) = th2; fun err msg = raise THM ("equal_elim: " ^ msg, 0, [th1, th2]); in (case prop1 of Const ("Pure.eq", _) $ A $ B => if prop2 aconv A then Thm (deriv_rule2 (Proofterm.equal_elim_proof A B) der1 der2, {cert = join_certificate2 (th1, th2), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = union_constraints constraints1 constraints2, shyps = Sorts.union shyps1 shyps2, hyps = union_hyps hyps1 hyps2, tpairs = union_tpairs tpairs1 tpairs2, prop = B}) else err "not equal" | _ => err "major premise") end; (**** Derived rules ****) (*Smash unifies the list of term pairs leaving no flex-flex pairs. Instantiates the theorem and deletes trivial tpairs. Resulting sequence may contain multiple elements if the tpairs are not all flex-flex.*) fun flexflex_rule opt_ctxt = solve_constraints #> (fn th => let val Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = th; val (context, cert') = make_context_certificate [th] opt_ctxt cert; in Unify.smash_unifiers context tpairs (Envir.empty maxidx) |> Seq.map (fn env => if Envir.is_empty env then th else let val tpairs' = tpairs |> map (apply2 (Envir.norm_term env)) (*remove trivial tpairs, of the form t \ t*) |> filter_out (op aconv); val der' = deriv_rule1 (Proofterm.norm_proof' env) der; val constraints' = insert_constraints_env (Context.certificate_theory cert') env constraints; val prop' = Envir.norm_term env prop; val maxidx = maxidx_tpairs tpairs' (maxidx_of_term prop'); val shyps = Envir.insert_sorts env shyps; in Thm (der', {cert = cert', tags = [], maxidx = maxidx, constraints = constraints', shyps = shyps, hyps = hyps, tpairs = tpairs', prop = prop'}) end) end); (*Generalization of fixed variables A -------------------- A[?'a/'a, ?x/x, ...] *) fun generalize ([], []) _ th = th | generalize (tfrees, frees) idx th = let val Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = th; val _ = idx <= maxidx andalso raise THM ("generalize: bad index", idx, [th]); val bad_type = if null tfrees then K false else Term.exists_subtype (fn TFree (a, _) => member (op =) tfrees a | _ => false); fun bad_term (Free (x, T)) = bad_type T orelse member (op =) frees x | bad_term (Var (_, T)) = bad_type T | bad_term (Const (_, T)) = bad_type T | bad_term (Abs (_, T, t)) = bad_type T orelse bad_term t | bad_term (t $ u) = bad_term t orelse bad_term u | bad_term (Bound _) = false; val _ = exists bad_term hyps andalso raise THM ("generalize: variable free in assumptions", 0, [th]); val generalize = Term_Subst.generalize (tfrees, frees) idx; val prop' = generalize prop; val tpairs' = map (apply2 generalize) tpairs; val maxidx' = maxidx_tpairs tpairs' (maxidx_of_term prop'); in Thm (deriv_rule1 (Proofterm.generalize_proof (tfrees, frees) idx prop) der, {cert = cert, tags = [], maxidx = maxidx', constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs', prop = prop'}) end; fun generalize_cterm ([], []) _ ct = ct | generalize_cterm (tfrees, frees) idx (ct as Cterm {cert, t, T, maxidx, sorts}) = if idx <= maxidx then raise CTERM ("generalize_cterm: bad index", [ct]) else Cterm {cert = cert, sorts = sorts, T = Term_Subst.generalizeT tfrees idx T, t = Term_Subst.generalize (tfrees, frees) idx t, maxidx = Int.max (maxidx, idx)}; fun generalize_ctyp [] _ cT = cT | generalize_ctyp tfrees idx (Ctyp {cert, T, maxidx, sorts}) = if idx <= maxidx then raise CTERM ("generalize_ctyp: bad index", []) else Ctyp {cert = cert, sorts = sorts, T = Term_Subst.generalizeT tfrees idx T, maxidx = Int.max (maxidx, idx)}; (*Instantiation of schematic variables A -------------------- A[t1/v1, ..., tn/vn] *) local fun pretty_typing thy t T = Pretty.block [Syntax.pretty_term_global thy t, Pretty.str " ::", Pretty.brk 1, Syntax.pretty_typ_global thy T]; fun add_inst (v as (_, T), cu) (cert, sorts) = let val Cterm {t = u, T = U, cert = cert2, sorts = sorts_u, maxidx = maxidx_u, ...} = cu; val cert' = Context.join_certificate (cert, cert2); val sorts' = Sorts.union sorts_u sorts; in if T = U then ((v, (u, maxidx_u)), (cert', sorts')) else let val msg = (case cert' of Context.Certificate thy' => Pretty.string_of (Pretty.block [Pretty.str "instantiate: type conflict", Pretty.fbrk, pretty_typing thy' (Var v) T, Pretty.fbrk, pretty_typing thy' u U]) | Context.Certificate_Id _ => "instantiate: type conflict"); in raise TYPE (msg, [T, U], [Var v, u]) end end; fun add_instT (v as (_, S), cU) (cert, sorts) = let val Ctyp {T = U, cert = cert2, sorts = sorts_U, maxidx = maxidx_U, ...} = cU; val cert' = Context.join_certificate (cert, cert2); val thy' = Context.certificate_theory cert' handle ERROR msg => raise CONTEXT (msg, [cU], [], [], NONE); val sorts' = Sorts.union sorts_U sorts; in if Sign.of_sort thy' (U, S) then ((v, (U, maxidx_U)), (cert', sorts')) else raise TYPE ("Type not of sort " ^ Syntax.string_of_sort_global thy' S, [U], []) end; in (*Left-to-right replacements: ctpairs = [..., (vi, ti), ...]. Instantiates distinct Vars by terms of same type. Does NOT normalize the resulting theorem!*) fun instantiate ([], []) th = th | instantiate (instT, inst) th = let val Thm (der, {cert, hyps, constraints, shyps, tpairs, prop, ...}) = th; val (inst', (instT', (cert', shyps'))) = (cert, shyps) |> fold_map add_inst inst ||> fold_map add_instT instT handle CONTEXT (msg, cTs, cts, ths, context) => raise CONTEXT (msg, cTs, cts, th :: ths, context); val subst = Term_Subst.instantiate_maxidx (instT', inst'); val (prop', maxidx1) = subst prop ~1; val (tpairs', maxidx') = fold_map (fn (t, u) => fn i => subst t i ||>> subst u) tpairs maxidx1; val thy' = Context.certificate_theory cert'; val constraints' = fold (fn ((_, S), (T, _)) => insert_constraints thy' (T, S)) instT' constraints; in Thm (deriv_rule1 (fn d => Proofterm.instantiate (map (apsnd #1) instT', map (apsnd #1) inst') d) der, {cert = cert', tags = [], maxidx = maxidx', constraints = constraints', shyps = shyps', hyps = hyps, tpairs = tpairs', prop = prop'}) |> solve_constraints end handle TYPE (msg, _, _) => raise THM (msg, 0, [th]); fun instantiate_cterm ([], []) ct = ct | instantiate_cterm (instT, inst) ct = let val Cterm {cert, t, T, sorts, ...} = ct; val (inst', (instT', (cert', sorts'))) = (cert, sorts) |> fold_map add_inst inst ||> fold_map add_instT instT; val subst = Term_Subst.instantiate_maxidx (instT', inst'); val substT = Term_Subst.instantiateT_maxidx instT'; val (t', maxidx1) = subst t ~1; val (T', maxidx') = substT T maxidx1; in Cterm {cert = cert', t = t', T = T', sorts = sorts', maxidx = maxidx'} end handle TYPE (msg, _, _) => raise CTERM (msg, [ct]); end; (*The trivial implication A \ A, justified by assume and forall rules. A can contain Vars, not so for assume!*) fun trivial (Cterm {cert, t = A, T, maxidx, sorts}) = if T <> propT then raise THM ("trivial: the term must have type prop", 0, []) else Thm (deriv_rule0 (fn () => Proofterm.trivial_proof), {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = Logic.mk_implies (A, A)}); (*Axiom-scheme reflecting signature contents T :: c ------------------- OFCLASS(T, c_class) *) fun of_class (cT, raw_c) = let val Ctyp {cert, T, ...} = cT; val thy = Context.certificate_theory cert handle ERROR msg => raise CONTEXT (msg, [cT], [], [], NONE); val c = Sign.certify_class thy raw_c; val Cterm {t = prop, maxidx, sorts, ...} = global_cterm_of thy (Logic.mk_of_class (T, c)); in if Sign.of_sort thy (T, [c]) then Thm (deriv_rule0 (fn () => Proofterm.OfClass (T, c)), {cert = cert, tags = [], maxidx = maxidx, constraints = insert_constraints thy (T, [c]) [], shyps = sorts, hyps = [], tpairs = [], prop = prop}) else raise THM ("of_class: type not of class " ^ Syntax.string_of_sort_global thy [c], 0, []) end |> solve_constraints; (*Remove extra sorts that are witnessed by type signature information*) fun strip_shyps (thm as Thm (_, {shyps = [], ...})) = thm | strip_shyps (thm as Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop})) = let val thy = theory_of_thm thm; val algebra = Sign.classes_of thy; val present = (fold_terms o fold_types o fold_atyps_sorts) (insert (eq_fst op =)) thm []; val extra = fold (Sorts.remove_sort o #2) present shyps; val witnessed = Sign.witness_sorts thy present extra; val extra' = fold (Sorts.remove_sort o #2) witnessed extra |> Sorts.minimal_sorts algebra; val shyps' = fold (Sorts.insert_sort o #2) present extra'; in Thm (deriv_rule_unconditional (Proofterm.strip_shyps_proof algebra present witnessed extra') der, {cert = cert, tags = tags, maxidx = maxidx, constraints = constraints, shyps = shyps', hyps = hyps, tpairs = tpairs, prop = prop}) end; (*Internalize sort constraints of type variables*) val unconstrainT = solve_constraints #> (fn thm as Thm (der, args) => let val Deriv {promises, body} = der; val {cert, shyps, hyps, tpairs, prop, ...} = args; val thy = theory_of_thm thm; fun err msg = raise THM ("unconstrainT: " ^ msg, 0, [thm]); val _ = null hyps orelse err "bad hyps"; val _ = null tpairs orelse err "bad flex-flex constraints"; val tfrees = rev (Term.add_tfree_names prop []); val _ = null tfrees orelse err ("illegal free type variables " ^ commas_quote tfrees); val ps = map (apsnd (Future.map fulfill_body)) promises; val (pthm, proof) = Proofterm.unconstrain_thm_proof thy (classrel_proof thy) (arity_proof thy) shyps prop ps body; val der' = make_deriv [] [] [pthm] proof; val prop' = Proofterm.thm_node_prop (#2 pthm); in Thm (der', {cert = cert, tags = [], maxidx = maxidx_of_term prop', constraints = [], shyps = [[]], (*potentially redundant*) hyps = [], tpairs = [], prop = prop'}) end); (*Replace all TFrees not fixed or in the hyps by new TVars*) fun varifyT_global' fixed (Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...})) = let val tfrees = fold Term.add_tfrees hyps fixed; val prop1 = attach_tpairs tpairs prop; val (al, prop2) = Type.varify_global tfrees prop1; val (ts, prop3) = Logic.strip_prems (length tpairs, [], prop2); in (al, Thm (deriv_rule1 (Proofterm.varify_proof prop tfrees) der, {cert = cert, tags = [], maxidx = Int.max (0, maxidx), constraints = constraints, shyps = shyps, hyps = hyps, tpairs = rev (map Logic.dest_equals ts), prop = prop3})) end; val varifyT_global = #2 o varifyT_global' []; (*Replace all TVars by TFrees that are often new*) fun legacy_freezeT (Thm (der, {cert, constraints, shyps, hyps, tpairs, prop, ...})) = let val prop1 = attach_tpairs tpairs prop; val prop2 = Type.legacy_freeze prop1; val (ts, prop3) = Logic.strip_prems (length tpairs, [], prop2); in Thm (deriv_rule1 (Proofterm.legacy_freezeT prop1) der, {cert = cert, tags = [], maxidx = maxidx_of_term prop2, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = rev (map Logic.dest_equals ts), prop = prop3}) end; fun plain_prop_of raw_thm = let val thm = strip_shyps raw_thm; fun err msg = raise THM ("plain_prop_of: " ^ msg, 0, [thm]); in if not (null (hyps_of thm)) then err "theorem may not contain hypotheses" else if not (null (extra_shyps thm)) then err "theorem may not contain sort hypotheses" else if not (null (tpairs_of thm)) then err "theorem may not contain flex-flex pairs" else prop_of thm end; (*** Inference rules for tactics ***) (*Destruct proof state into constraints, other goals, goal(i), rest *) fun dest_state (state as Thm (_, {prop,tpairs,...}), i) = (case Logic.strip_prems(i, [], prop) of (B::rBs, C) => (tpairs, rev rBs, B, C) | _ => raise THM("dest_state", i, [state])) handle TERM _ => raise THM("dest_state", i, [state]); (*Prepare orule for resolution by lifting it over the parameters and assumptions of goal.*) fun lift_rule goal orule = let val Cterm {t = gprop, T, maxidx = gmax, sorts, ...} = goal; val inc = gmax + 1; val lift_abs = Logic.lift_abs inc gprop; val lift_all = Logic.lift_all inc gprop; val Thm (der, {maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = orule; val (As, B) = Logic.strip_horn prop; in if T <> propT then raise THM ("lift_rule: the term must have type prop", 0, []) else Thm (deriv_rule1 (Proofterm.lift_proof gprop inc prop) der, {cert = join_certificate1 (goal, orule), tags = [], maxidx = maxidx + inc, constraints = constraints, shyps = Sorts.union shyps sorts, (*sic!*) hyps = hyps, tpairs = map (apply2 lift_abs) tpairs, prop = Logic.list_implies (map lift_all As, lift_all B)}) end; fun incr_indexes i (thm as Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...})) = if i < 0 then raise THM ("negative increment", 0, [thm]) else if i = 0 then thm else Thm (deriv_rule1 (Proofterm.incr_indexes i) der, {cert = cert, tags = [], maxidx = maxidx + i, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = map (apply2 (Logic.incr_indexes ([], [], i))) tpairs, prop = Logic.incr_indexes ([], [], i) prop}); (*Solve subgoal Bi of proof state B1...Bn/C by assumption. *) fun assumption opt_ctxt i state = let val Thm (der, {cert, maxidx, constraints, shyps, hyps, ...}) = state; val (context, cert') = make_context_certificate [state] opt_ctxt cert; val (tpairs, Bs, Bi, C) = dest_state (state, i); fun newth n (env, tpairs) = let val normt = Envir.norm_term env; fun assumption_proof prf = Proofterm.assumption_proof (map normt Bs) (normt Bi) n prf; in Thm (deriv_rule1 (assumption_proof #> not (Envir.is_empty env) ? Proofterm.norm_proof' env) der, {tags = [], maxidx = Envir.maxidx_of env, constraints = insert_constraints_env (Context.certificate_theory cert') env constraints, shyps = Envir.insert_sorts env shyps, hyps = hyps, tpairs = if Envir.is_empty env then tpairs else map (apply2 normt) tpairs, prop = if Envir.is_empty env then Logic.list_implies (Bs, C) (*avoid wasted normalizations*) else normt (Logic.list_implies (Bs, C)) (*normalize the new rule fully*), cert = cert'}) end; val (close, asms, concl) = Logic.assum_problems (~1, Bi); val concl' = close concl; fun addprfs [] _ = Seq.empty | addprfs (asm :: rest) n = Seq.make (fn () => Seq.pull (Seq.mapp (newth n) (if Term.could_unify (asm, concl) then (Unify.unifiers (context, Envir.empty maxidx, (close asm, concl') :: tpairs)) else Seq.empty) (addprfs rest (n + 1)))) in addprfs asms 1 end; (*Solve subgoal Bi of proof state B1...Bn/C by assumption. Checks if Bi's conclusion is alpha/eta-convertible to one of its assumptions*) fun eq_assumption i state = let val Thm (der, {cert, maxidx, constraints, shyps, hyps, ...}) = state; val (tpairs, Bs, Bi, C) = dest_state (state, i); val (_, asms, concl) = Logic.assum_problems (~1, Bi); in (case find_index (fn asm => Envir.aeconv (asm, concl)) asms of ~1 => raise THM ("eq_assumption", 0, [state]) | n => Thm (deriv_rule1 (Proofterm.assumption_proof Bs Bi (n + 1)) der, {cert = cert, tags = [], maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = Logic.list_implies (Bs, C)})) end; (*For rotate_tac: fast rotation of assumptions of subgoal i*) fun rotate_rule k i state = let val Thm (der, {cert, maxidx, constraints, shyps, hyps, ...}) = state; val (tpairs, Bs, Bi, C) = dest_state (state, i); val params = Term.strip_all_vars Bi; val rest = Term.strip_all_body Bi; val asms = Logic.strip_imp_prems rest val concl = Logic.strip_imp_concl rest; val n = length asms; val m = if k < 0 then n + k else k; val Bi' = if 0 = m orelse m = n then Bi else if 0 < m andalso m < n then let val (ps, qs) = chop m asms in Logic.list_all (params, Logic.list_implies (qs @ ps, concl)) end else raise THM ("rotate_rule", k, [state]); in Thm (deriv_rule1 (Proofterm.rotate_proof Bs Bi' params asms m) der, {cert = cert, tags = [], maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = Logic.list_implies (Bs @ [Bi'], C)}) end; (*Rotates a rule's premises to the left by k, leaving the first j premises unchanged. Does nothing if k=0 or if k equals n-j, where n is the number of premises. Useful with eresolve_tac and underlies defer_tac*) fun permute_prems j k rl = let val Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = rl; val prems = Logic.strip_imp_prems prop and concl = Logic.strip_imp_concl prop; val moved_prems = List.drop (prems, j) and fixed_prems = List.take (prems, j) handle General.Subscript => raise THM ("permute_prems: j", j, [rl]); val n_j = length moved_prems; val m = if k < 0 then n_j + k else k; val (prems', prop') = if 0 = m orelse m = n_j then (prems, prop) else if 0 < m andalso m < n_j then let val (ps, qs) = chop m moved_prems; val prems' = fixed_prems @ qs @ ps; in (prems', Logic.list_implies (prems', concl)) end else raise THM ("permute_prems: k", k, [rl]); in Thm (deriv_rule1 (Proofterm.permute_prems_proof prems' j m) der, {cert = cert, tags = [], maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop'}) end; (* strip_apply f B A strips off all assumptions/parameters from A introduced by lifting over B, and applies f to remaining part of A*) fun strip_apply f = let fun strip (Const ("Pure.imp", _) $ _ $ B1) (Const ("Pure.imp", _) $ A2 $ B2) = Logic.mk_implies (A2, strip B1 B2) | strip ((c as Const ("Pure.all", _)) $ Abs (_, _, t1)) ( Const ("Pure.all", _) $ Abs (a, T, t2)) = c $ Abs (a, T, strip t1 t2) | strip _ A = f A in strip end; fun strip_lifted (Const ("Pure.imp", _) $ _ $ B1) (Const ("Pure.imp", _) $ _ $ B2) = strip_lifted B1 B2 | strip_lifted (Const ("Pure.all", _) $ Abs (_, _, t1)) (Const ("Pure.all", _) $ Abs (_, _, t2)) = strip_lifted t1 t2 | strip_lifted _ A = A; (*Use the alist to rename all bound variables and some unknowns in a term dpairs = current disagreement pairs; tpairs = permanent ones (flexflex); Preserves unknowns in tpairs and on lhs of dpairs. *) fun rename_bvs [] _ _ _ _ = K I | rename_bvs al dpairs tpairs B As = let val add_var = fold_aterms (fn Var ((x, _), _) => insert (op =) x | _ => I); val vids = [] |> fold (add_var o fst) dpairs |> fold (add_var o fst) tpairs |> fold (add_var o snd) tpairs; val vids' = fold (add_var o strip_lifted B) As []; (*unknowns appearing elsewhere be preserved!*) val al' = distinct ((op =) o apply2 fst) (filter_out (fn (x, y) => not (member (op =) vids' x) orelse member (op =) vids x orelse member (op =) vids y) al); val unchanged = filter_out (AList.defined (op =) al') vids'; fun del_clashing clash xs _ [] qs = if clash then del_clashing false xs xs qs [] else qs | del_clashing clash xs ys ((p as (x, y)) :: ps) qs = if member (op =) ys y then del_clashing true (x :: xs) (x :: ys) ps qs else del_clashing clash xs (y :: ys) ps (p :: qs); val al'' = del_clashing false unchanged unchanged al' []; fun rename (t as Var ((x, i), T)) = (case AList.lookup (op =) al'' x of SOME y => Var ((y, i), T) | NONE => t) | rename (Abs (x, T, t)) = Abs (the_default x (AList.lookup (op =) al x), T, rename t) | rename (f $ t) = rename f $ rename t | rename t = t; fun strip_ren f Ai = f rename B Ai in strip_ren end; (*Function to rename bounds/unknowns in the argument, lifted over B*) fun rename_bvars dpairs = rename_bvs (fold_rev Term.match_bvars dpairs []) dpairs; (*** RESOLUTION ***) (** Lifting optimizations **) (*strip off pairs of assumptions/parameters in parallel -- they are identical because of lifting*) fun strip_assums2 (Const("Pure.imp", _) $ _ $ B1, Const("Pure.imp", _) $ _ $ B2) = strip_assums2 (B1,B2) | strip_assums2 (Const("Pure.all",_)$Abs(a,T,t1), Const("Pure.all",_)$Abs(_,_,t2)) = let val (B1,B2) = strip_assums2 (t1,t2) in (Abs(a,T,B1), Abs(a,T,B2)) end | strip_assums2 BB = BB; (*Faster normalization: skip assumptions that were lifted over*) fun norm_term_skip env 0 t = Envir.norm_term env t | norm_term_skip env n (Const ("Pure.all", _) $ Abs (a, T, t)) = let val T' = Envir.norm_type (Envir.type_env env) T (*Must instantiate types of parameters because they are flattened; this could be a NEW parameter*) in Logic.all_const T' $ Abs (a, T', norm_term_skip env n t) end | norm_term_skip env n (Const ("Pure.imp", _) $ A $ B) = Logic.mk_implies (A, norm_term_skip env (n - 1) B) | norm_term_skip _ _ _ = error "norm_term_skip: too few assumptions??"; (*unify types of schematic variables (non-lifted case)*) fun unify_var_types context (th1, th2) env = let fun unify_vars (T :: Us) = fold (fn U => Pattern.unify_types context (T, U)) Us | unify_vars _ = I; val add_vars = full_prop_of #> fold_aterms (fn Var v => Vartab.insert_list (op =) v | _ => I); val vars = Vartab.empty |> add_vars th1 |> add_vars th2; in SOME (Vartab.fold (unify_vars o #2) vars env) end handle Pattern.Unif => NONE; (*Composition of object rule r=(A1...Am/B) with proof state s=(B1...Bn/C) Unifies B with Bi, replacing subgoal i (1 <= i <= n) If match then forbid instantiations in proof state If lifted then shorten the dpair using strip_assums2. If eres_flg then simultaneously proves A1 by assumption. nsubgoal is the number of new subgoals (written m above). Curried so that resolution calls dest_state only once. *) local exception COMPOSE in fun bicompose_aux opt_ctxt {flatten, match, incremented} (state, (stpairs, Bs, Bi, C), lifted) (eres_flg, orule, nsubgoal) = let val Thm (sder, {maxidx=smax, constraints = constraints2, shyps = shyps2, hyps = hyps2, ...}) = state and Thm (rder, {maxidx=rmax, constraints = constraints1, shyps = shyps1, hyps = hyps1, tpairs=rtpairs, prop=rprop,...}) = orule (*How many hyps to skip over during normalization*) and nlift = Logic.count_prems (strip_all_body Bi) + (if eres_flg then ~1 else 0) val (context, cert) = make_context_certificate [state, orule] opt_ctxt (join_certificate2 (state, orule)); (*Add new theorem with prop = "\Bs; As\ \ C" to thq*) fun addth A (As, oldAs, rder', n) ((env, tpairs), thq) = let val normt = Envir.norm_term env; (*perform minimal copying here by examining env*) val (ntpairs, normp) = if Envir.is_empty env then (tpairs, (Bs @ As, C)) else let val ntps = map (apply2 normt) tpairs in if Envir.above env smax then (*no assignments in state; normalize the rule only*) if lifted then (ntps, (Bs @ map (norm_term_skip env nlift) As, C)) else (ntps, (Bs @ map normt As, C)) else if match then raise COMPOSE else (*normalize the new rule fully*) (ntps, (map normt (Bs @ As), normt C)) end val constraints' = union_constraints constraints1 constraints2 |> insert_constraints_env (Context.certificate_theory cert) env; fun bicompose_proof prf1 prf2 = Proofterm.bicompose_proof flatten (map normt Bs) (map normt As) A oldAs n (nlift+1) prf1 prf2 val th = Thm (deriv_rule2 (if Envir.is_empty env then bicompose_proof else if Envir.above env smax then bicompose_proof o Proofterm.norm_proof' env else Proofterm.norm_proof' env oo bicompose_proof) rder' sder, {tags = [], maxidx = Envir.maxidx_of env, constraints = constraints', shyps = Envir.insert_sorts env (Sorts.union shyps1 shyps2), hyps = union_hyps hyps1 hyps2, tpairs = ntpairs, prop = Logic.list_implies normp, cert = cert}) in Seq.cons th thq end handle COMPOSE => thq; val (rAs,B) = Logic.strip_prems(nsubgoal, [], rprop) handle TERM _ => raise THM("bicompose: rule", 0, [orule,state]); (*Modify assumptions, deleting n-th if n>0 for e-resolution*) fun newAs(As0, n, dpairs, tpairs) = let val (As1, rder') = if not lifted then (As0, rder) else let val rename = rename_bvars dpairs tpairs B As0 in (map (rename strip_apply) As0, deriv_rule1 (Proofterm.map_proof_terms (rename K) I) rder) end; in (map (if flatten then (Logic.flatten_params n) else I) As1, As1, rder', n) handle TERM _ => raise THM("bicompose: 1st premise", 0, [orule]) end; val BBi = if lifted then strip_assums2(B,Bi) else (B,Bi); val dpairs = BBi :: (rtpairs@stpairs); (*elim-resolution: try each assumption in turn*) fun eres _ [] = raise THM ("bicompose: no premises", 0, [orule, state]) | eres env (A1 :: As) = let val A = SOME A1; val (close, asms, concl) = Logic.assum_problems (nlift + 1, A1); val concl' = close concl; fun tryasms [] _ = Seq.empty | tryasms (asm :: rest) n = if Term.could_unify (asm, concl) then let val asm' = close asm in (case Seq.pull (Unify.unifiers (context, env, (asm', concl') :: dpairs)) of NONE => tryasms rest (n + 1) | cell as SOME ((_, tpairs), _) => Seq.it_right (addth A (newAs (As, n, [BBi, (concl', asm')], tpairs))) (Seq.make (fn () => cell), Seq.make (fn () => Seq.pull (tryasms rest (n + 1))))) end else tryasms rest (n + 1); in tryasms asms 1 end; (*ordinary resolution*) fun res env = (case Seq.pull (Unify.unifiers (context, env, dpairs)) of NONE => Seq.empty | cell as SOME ((_, tpairs), _) => Seq.it_right (addth NONE (newAs (rev rAs, 0, [BBi], tpairs))) (Seq.make (fn () => cell), Seq.empty)); val env0 = Envir.empty (Int.max (rmax, smax)); in (case if incremented then SOME env0 else unify_var_types context (state, orule) env0 of NONE => Seq.empty | SOME env => if eres_flg then eres env (rev rAs) else res env) end; end; fun bicompose opt_ctxt flags arg i state = bicompose_aux opt_ctxt flags (state, dest_state (state,i), false) arg; (*Quick test whether rule is resolvable with the subgoal with hyps Hs and conclusion B. If eres_flg then checks 1st premise of rule also*) fun could_bires (Hs, B, eres_flg, rule) = let fun could_reshyp (A1::_) = exists (fn H => Term.could_unify (A1, H)) Hs | could_reshyp [] = false; (*no premise -- illegal*) in Term.could_unify(concl_of rule, B) andalso (not eres_flg orelse could_reshyp (prems_of rule)) end; (*Bi-resolution of a state with a list of (flag,rule) pairs. Puts the rule above: rule/state. Renames vars in the rules. *) fun biresolution opt_ctxt match brules i state = let val (stpairs, Bs, Bi, C) = dest_state(state,i); val lift = lift_rule (cprem_of state i); val B = Logic.strip_assums_concl Bi; val Hs = Logic.strip_assums_hyp Bi; val compose = bicompose_aux opt_ctxt {flatten = true, match = match, incremented = true} (state, (stpairs, Bs, Bi, C), true); fun res [] = Seq.empty | res ((eres_flg, rule)::brules) = if Config.get_generic (make_context [state] opt_ctxt (cert_of state)) Pattern.unify_trace_failure orelse could_bires (Hs, B, eres_flg, rule) then Seq.make (*delay processing remainder till needed*) (fn()=> SOME(compose (eres_flg, lift rule, nprems_of rule), res brules)) else res brules in Seq.flat (res brules) end; (*Resolution: exactly one resolvent must be produced*) fun tha RSN (i, thb) = (case Seq.chop 2 (biresolution NONE false [(false, tha)] i thb) of ([th], _) => solve_constraints th | ([], _) => raise THM ("RSN: no unifiers", i, [tha, thb]) | _ => raise THM ("RSN: multiple unifiers", i, [tha, thb])); (*Resolution: P \ Q, Q \ R gives P \ R*) fun tha RS thb = tha RSN (1,thb); (**** Type classes ****) fun standard_tvars thm = let val thy = theory_of_thm thm; val tvars = rev (Term.add_tvars (prop_of thm) []); val names = Name.invent Name.context Name.aT (length tvars); val tinst = map2 (fn (ai, S) => fn b => ((ai, S), global_ctyp_of thy (TVar ((b, 0), S)))) tvars names; in instantiate (tinst, []) thm end (* class relations *) val is_classrel = Symreltab.defined o get_classrels; fun complete_classrels thy = let fun complete (c, (_, (all_preds, all_succs))) (finished1, thy1) = let fun compl c1 c2 (finished2, thy2) = if is_classrel thy2 (c1, c2) then (finished2, thy2) else (false, thy2 |> (map_classrels o Symreltab.update) ((c1, c2), (the_classrel thy2 (c1, c) RS the_classrel thy2 (c, c2)) |> standard_tvars |> close_derivation \<^here> |> tap (expose_proof thy2) |> trim_context)); val proven = is_classrel thy1; val preds = Graph.Keys.fold (fn c1 => proven (c1, c) ? cons c1) all_preds []; val succs = Graph.Keys.fold (fn c2 => proven (c, c2) ? cons c2) all_succs []; in fold_product compl preds succs (finished1, thy1) end; in (case Graph.fold complete (Sorts.classes_of (Sign.classes_of thy)) (true, thy) of (true, _) => NONE | (_, thy') => SOME thy') end; (* type arities *) fun thynames_of_arity thy (a, c) = (get_arities thy, []) |-> Aritytab.fold (fn ((a', _, c'), (_, name, ser)) => (a = a' andalso c = c') ? cons (name, ser)) |> sort (int_ord o apply2 #2) |> map #1; fun insert_arity_completions thy ((t, Ss, c), (th, thy_name, ser)) (finished, arities) = let val completions = Sign.super_classes thy c |> map_filter (fn c1 => if Aritytab.defined arities (t, Ss, c1) then NONE else let val th1 = (th RS the_classrel thy (c, c1)) |> standard_tvars |> close_derivation \<^here> |> tap (expose_proof thy) |> trim_context; in SOME ((t, Ss, c1), (th1, thy_name, ser)) end); val finished' = finished andalso null completions; val arities' = fold Aritytab.update completions arities; in (finished', arities') end; fun complete_arities thy = let val arities = get_arities thy; val (finished, arities') = Aritytab.fold (insert_arity_completions thy) arities (true, get_arities thy); in if finished then NONE else SOME (map_arities (K arities') thy) end; val _ = Theory.setup (Theory.at_begin complete_classrels #> Theory.at_begin complete_arities); (* primitive rules *) fun add_classrel raw_th thy = let val th = strip_shyps (transfer thy raw_th); val th' = th |> unconstrainT |> tap (expose_proof thy) |> trim_context; val prop = plain_prop_of th; val (c1, c2) = Logic.dest_classrel prop; in thy |> Sign.primitive_classrel (c1, c2) |> map_classrels (Symreltab.update ((c1, c2), th')) |> perhaps complete_classrels |> perhaps complete_arities end; fun add_arity raw_th thy = let val th = strip_shyps (transfer thy raw_th); val th' = th |> unconstrainT |> tap (expose_proof thy) |> trim_context; val prop = plain_prop_of th; val (t, Ss, c) = Logic.dest_arity prop; val ar = ((t, Ss, c), (th', Context.theory_name thy, serial ())); in thy |> Sign.primitive_arity (t, Ss, [c]) |> map_arities (Aritytab.update ar #> curry (insert_arity_completions thy ar) true #> #2) end; end; structure Basic_Thm: BASIC_THM = Thm; open Basic_Thm; diff --git a/src/ZF/List.thy b/src/ZF/List.thy --- a/src/ZF/List.thy +++ b/src/ZF/List.thy @@ -1,1271 +1,1270 @@ (* Title: ZF/List.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge *) section\Lists in Zermelo-Fraenkel Set Theory\ theory List imports Datatype ArithSimp begin consts list :: "i=>i" datatype "list(A)" = Nil | Cons ("a \ A", "l \ list(A)") syntax "_Nil" :: i (\[]\) "_List" :: "is => i" (\[(_)]\) translations "[x, xs]" == "CONST Cons(x, [xs])" "[x]" == "CONST Cons(x, [])" "[]" == "CONST Nil" consts length :: "i=>i" hd :: "i=>i" tl :: "i=>i" primrec "length([]) = 0" "length(Cons(a,l)) = succ(length(l))" primrec "hd([]) = 0" "hd(Cons(a,l)) = a" primrec "tl([]) = []" "tl(Cons(a,l)) = l" consts map :: "[i=>i, i] => i" set_of_list :: "i=>i" app :: "[i,i]=>i" (infixr \@\ 60) (*map is a binding operator -- it applies to meta-level functions, not object-level functions. This simplifies the final form of term_rec_conv, although complicating its derivation.*) primrec "map(f,[]) = []" "map(f,Cons(a,l)) = Cons(f(a), map(f,l))" primrec "set_of_list([]) = 0" "set_of_list(Cons(a,l)) = cons(a, set_of_list(l))" primrec app_Nil: "[] @ ys = ys" app_Cons: "(Cons(a,l)) @ ys = Cons(a, l @ ys)" consts rev :: "i=>i" flat :: "i=>i" list_add :: "i=>i" primrec "rev([]) = []" "rev(Cons(a,l)) = rev(l) @ [a]" primrec "flat([]) = []" "flat(Cons(l,ls)) = l @ flat(ls)" primrec "list_add([]) = 0" "list_add(Cons(a,l)) = a #+ list_add(l)" consts drop :: "[i,i]=>i" primrec drop_0: "drop(0,l) = l" drop_succ: "drop(succ(i), l) = tl (drop(i,l))" (*** Thanks to Sidi Ehmety for the following ***) definition (* Function `take' returns the first n elements of a list *) take :: "[i,i]=>i" where "take(n, as) == list_rec(\n\nat. [], %a l r. \n\nat. nat_case([], %m. Cons(a, r`m), n), as)`n" definition nth :: "[i, i]=>i" where \ \returns the (n+1)th element of a list, or 0 if the list is too short.\ "nth(n, as) == list_rec(\n\nat. 0, %a l r. \n\nat. nat_case(a, %m. r`m, n), as) ` n" definition list_update :: "[i, i, i]=>i" where "list_update(xs, i, v) == list_rec(\n\nat. Nil, %u us vs. \n\nat. nat_case(Cons(v, us), %m. Cons(u, vs`m), n), xs)`i" consts filter :: "[i=>o, i] => i" upt :: "[i, i] =>i" primrec "filter(P, Nil) = Nil" "filter(P, Cons(x, xs)) = (if P(x) then Cons(x, filter(P, xs)) else filter(P, xs))" primrec "upt(i, 0) = Nil" "upt(i, succ(j)) = (if i \ j then upt(i, j)@[j] else Nil)" definition min :: "[i,i] =>i" where "min(x, y) == (if x \ y then x else y)" definition max :: "[i, i] =>i" where "max(x, y) == (if x \ y then y else x)" (*** Aspects of the datatype definition ***) declare list.intros [simp,TC] (*An elimination rule, for type-checking*) inductive_cases ConsE: "Cons(a,l) \ list(A)" lemma Cons_type_iff [simp]: "Cons(a,l) \ list(A) \ a \ A & l \ list(A)" by (blast elim: ConsE) (*Proving freeness results*) lemma Cons_iff: "Cons(a,l)=Cons(a',l') \ a=a' & l=l'" by auto lemma Nil_Cons_iff: "~ Nil=Cons(a,l)" by auto lemma list_unfold: "list(A) = {0} + (A * list(A))" by (blast intro!: list.intros [unfolded list.con_defs] elim: list.cases [unfolded list.con_defs]) (** Lemmas to justify using "list" in other recursive type definitions **) lemma list_mono: "A<=B ==> list(A) \ list(B)" apply (unfold list.defs ) apply (rule lfp_mono) apply (simp_all add: list.bnd_mono) apply (assumption | rule univ_mono basic_monos)+ done (*There is a similar proof by list induction.*) lemma list_univ: "list(univ(A)) \ univ(A)" apply (unfold list.defs list.con_defs) apply (rule lfp_lowerbound) apply (rule_tac [2] A_subset_univ [THEN univ_mono]) apply (blast intro!: zero_in_univ Inl_in_univ Inr_in_univ Pair_in_univ) done (*These two theorems justify datatypes involving list(nat), list(A), ...*) lemmas list_subset_univ = subset_trans [OF list_mono list_univ] lemma list_into_univ: "[| l \ list(A); A \ univ(B) |] ==> l \ univ(B)" by (blast intro: list_subset_univ [THEN subsetD]) lemma list_case_type: "[| l \ list(A); c \ C(Nil); !!x y. [| x \ A; y \ list(A) |] ==> h(x,y): C(Cons(x,y)) |] ==> list_case(c,h,l) \ C(l)" by (erule list.induct, auto) lemma list_0_triv: "list(0) = {Nil}" apply (rule equalityI, auto) apply (induct_tac x, auto) done (*** List functions ***) lemma tl_type: "l \ list(A) ==> tl(l) \ list(A)" apply (induct_tac "l") apply (simp_all (no_asm_simp) add: list.intros) done (** drop **) lemma drop_Nil [simp]: "i \ nat ==> drop(i, Nil) = Nil" apply (induct_tac "i") apply (simp_all (no_asm_simp)) done lemma drop_succ_Cons [simp]: "i \ nat ==> drop(succ(i), Cons(a,l)) = drop(i,l)" apply (rule sym) apply (induct_tac "i") apply (simp (no_asm)) apply (simp (no_asm_simp)) done lemma drop_type [simp,TC]: "[| i \ nat; l \ list(A) |] ==> drop(i,l) \ list(A)" apply (induct_tac "i") apply (simp_all (no_asm_simp) add: tl_type) done declare drop_succ [simp del] (** Type checking -- proved by induction, as usual **) lemma list_rec_type [TC]: "[| l \ list(A); c \ C(Nil); !!x y r. [| x \ A; y \ list(A); r \ C(y) |] ==> h(x,y,r): C(Cons(x,y)) |] ==> list_rec(c,h,l) \ C(l)" by (induct_tac "l", auto) (** map **) lemma map_type [TC]: "[| l \ list(A); !!x. x \ A ==> h(x): B |] ==> map(h,l) \ list(B)" apply (simp add: map_list_def) apply (typecheck add: list.intros list_rec_type, blast) done lemma map_type2 [TC]: "l \ list(A) ==> map(h,l) \ list({h(u). u \ A})" apply (erule map_type) apply (erule RepFunI) done (** length **) lemma length_type [TC]: "l \ list(A) ==> length(l) \ nat" by (simp add: length_list_def) lemma lt_length_in_nat: "[|x < length(xs); xs \ list(A)|] ==> x \ nat" by (frule lt_nat_in_nat, typecheck) (** app **) lemma app_type [TC]: "[| xs: list(A); ys: list(A) |] ==> xs@ys \ list(A)" by (simp add: app_list_def) (** rev **) lemma rev_type [TC]: "xs: list(A) ==> rev(xs) \ list(A)" by (simp add: rev_list_def) (** flat **) lemma flat_type [TC]: "ls: list(list(A)) ==> flat(ls) \ list(A)" by (simp add: flat_list_def) (** set_of_list **) lemma set_of_list_type [TC]: "l \ list(A) ==> set_of_list(l) \ Pow(A)" apply (unfold set_of_list_list_def) apply (erule list_rec_type, auto) done lemma set_of_list_append: "xs: list(A) ==> set_of_list (xs@ys) = set_of_list(xs) \ set_of_list(ys)" apply (erule list.induct) apply (simp_all (no_asm_simp) add: Un_cons) done (** list_add **) lemma list_add_type [TC]: "xs: list(nat) ==> list_add(xs) \ nat" by (simp add: list_add_list_def) (*** theorems about map ***) lemma map_ident [simp]: "l \ list(A) ==> map(%u. u, l) = l" apply (induct_tac "l") apply (simp_all (no_asm_simp)) done lemma map_compose: "l \ list(A) ==> map(h, map(j,l)) = map(%u. h(j(u)), l)" apply (induct_tac "l") apply (simp_all (no_asm_simp)) done lemma map_app_distrib: "xs: list(A) ==> map(h, xs@ys) = map(h,xs) @ map(h,ys)" apply (induct_tac "xs") apply (simp_all (no_asm_simp)) done lemma map_flat: "ls: list(list(A)) ==> map(h, flat(ls)) = flat(map(map(h),ls))" apply (induct_tac "ls") apply (simp_all (no_asm_simp) add: map_app_distrib) done lemma list_rec_map: "l \ list(A) ==> list_rec(c, d, map(h,l)) = list_rec(c, %x xs r. d(h(x), map(h,xs), r), l)" apply (induct_tac "l") apply (simp_all (no_asm_simp)) done (** theorems about list(Collect(A,P)) -- used in Induct/Term.thy **) (* @{term"c \ list(Collect(B,P)) ==> c \ list"} *) lemmas list_CollectD = Collect_subset [THEN list_mono, THEN subsetD] lemma map_list_Collect: "l \ list({x \ A. h(x)=j(x)}) ==> map(h,l) = map(j,l)" apply (induct_tac "l") apply (simp_all (no_asm_simp)) done (*** theorems about length ***) lemma length_map [simp]: "xs: list(A) ==> length(map(h,xs)) = length(xs)" by (induct_tac "xs", simp_all) lemma length_app [simp]: "[| xs: list(A); ys: list(A) |] ==> length(xs@ys) = length(xs) #+ length(ys)" by (induct_tac "xs", simp_all) lemma length_rev [simp]: "xs: list(A) ==> length(rev(xs)) = length(xs)" apply (induct_tac "xs") apply (simp_all (no_asm_simp) add: length_app) done lemma length_flat: "ls: list(list(A)) ==> length(flat(ls)) = list_add(map(length,ls))" apply (induct_tac "ls") apply (simp_all (no_asm_simp) add: length_app) done (** Length and drop **) (*Lemma for the inductive step of drop_length*) lemma drop_length_Cons [rule_format]: "xs: list(A) ==> \x. \z zs. drop(length(xs), Cons(x,xs)) = Cons(z,zs)" by (erule list.induct, simp_all) lemma drop_length [rule_format]: "l \ list(A) ==> \i \ length(l). (\z zs. drop(i,l) = Cons(z,zs))" apply (erule list.induct, simp_all, safe) apply (erule drop_length_Cons) apply (rule natE) apply (erule Ord_trans [OF asm_rl length_type Ord_nat], assumption, simp_all) apply (blast intro: succ_in_naturalD length_type) done (*** theorems about app ***) lemma app_right_Nil [simp]: "xs: list(A) ==> xs@Nil=xs" by (erule list.induct, simp_all) lemma app_assoc: "xs: list(A) ==> (xs@ys)@zs = xs@(ys@zs)" by (induct_tac "xs", simp_all) lemma flat_app_distrib: "ls: list(list(A)) ==> flat(ls@ms) = flat(ls)@flat(ms)" apply (induct_tac "ls") apply (simp_all (no_asm_simp) add: app_assoc) done (*** theorems about rev ***) lemma rev_map_distrib: "l \ list(A) ==> rev(map(h,l)) = map(h,rev(l))" apply (induct_tac "l") apply (simp_all (no_asm_simp) add: map_app_distrib) done (*Simplifier needs the premises as assumptions because rewriting will not instantiate the variable ?A in the rules' typing conditions; note that rev_type does not instantiate ?A. Only the premises do. *) lemma rev_app_distrib: "[| xs: list(A); ys: list(A) |] ==> rev(xs@ys) = rev(ys)@rev(xs)" apply (erule list.induct) apply (simp_all add: app_assoc) done lemma rev_rev_ident [simp]: "l \ list(A) ==> rev(rev(l))=l" apply (induct_tac "l") apply (simp_all (no_asm_simp) add: rev_app_distrib) done lemma rev_flat: "ls: list(list(A)) ==> rev(flat(ls)) = flat(map(rev,rev(ls)))" apply (induct_tac "ls") apply (simp_all add: map_app_distrib flat_app_distrib rev_app_distrib) done (*** theorems about list_add ***) lemma list_add_app: "[| xs: list(nat); ys: list(nat) |] ==> list_add(xs@ys) = list_add(ys) #+ list_add(xs)" apply (induct_tac "xs", simp_all) done lemma list_add_rev: "l \ list(nat) ==> list_add(rev(l)) = list_add(l)" apply (induct_tac "l") apply (simp_all (no_asm_simp) add: list_add_app) done lemma list_add_flat: "ls: list(list(nat)) ==> list_add(flat(ls)) = list_add(map(list_add,ls))" apply (induct_tac "ls") apply (simp_all (no_asm_simp) add: list_add_app) done (** New induction rules **) lemma list_append_induct [case_names Nil snoc, consumes 1]: "[| l \ list(A); P(Nil); !!x y. [| x \ A; y \ list(A); P(y) |] ==> P(y @ [x]) |] ==> P(l)" apply (subgoal_tac "P(rev(rev(l)))", simp) apply (erule rev_type [THEN list.induct], simp_all) done lemma list_complete_induct_lemma [rule_format]: assumes ih: "\l. [| l \ list(A); \l' \ list(A). length(l') < length(l) \ P(l')|] ==> P(l)" shows "n \ nat ==> \l \ list(A). length(l) < n \ P(l)" apply (induct_tac n, simp) apply (blast intro: ih elim!: leE) done theorem list_complete_induct: "[| l \ list(A); \l. [| l \ list(A); \l' \ list(A). length(l') < length(l) \ P(l')|] ==> P(l) |] ==> P(l)" apply (rule list_complete_induct_lemma [of A]) prefer 4 apply (rule le_refl, simp) apply blast apply simp apply assumption done (*** Thanks to Sidi Ehmety for these results about min, take, etc. ***) (** min FIXME: replace by Int! **) (* Min theorems are also true for i, j ordinals *) lemma min_sym: "[| i \ nat; j \ nat |] ==> min(i,j)=min(j,i)" apply (unfold min_def) apply (auto dest!: not_lt_imp_le dest: lt_not_sym intro: le_anti_sym) done lemma min_type [simp,TC]: "[| i \ nat; j \ nat |] ==> min(i,j):nat" by (unfold min_def, auto) lemma min_0 [simp]: "i \ nat ==> min(0,i) = 0" apply (unfold min_def) apply (auto dest: not_lt_imp_le) done lemma min_02 [simp]: "i \ nat ==> min(i, 0) = 0" apply (unfold min_def) apply (auto dest: not_lt_imp_le) done lemma lt_min_iff: "[| i \ nat; j \ nat; k \ nat |] ==> i i nat; j \ nat |] ==> min(succ(i), succ(j))= succ(min(i, j))" apply (unfold min_def, auto) done (*** more theorems about lists ***) (** filter **) lemma filter_append [simp]: "xs:list(A) ==> filter(P, xs@ys) = filter(P, xs) @ filter(P, ys)" by (induct_tac "xs", auto) lemma filter_type [simp,TC]: "xs:list(A) ==> filter(P, xs):list(A)" by (induct_tac "xs", auto) lemma length_filter: "xs:list(A) ==> length(filter(P, xs)) \ length(xs)" apply (induct_tac "xs", auto) apply (rule_tac j = "length (l) " in le_trans) apply (auto simp add: le_iff) done lemma filter_is_subset: "xs:list(A) ==> set_of_list(filter(P,xs)) \ set_of_list(xs)" by (induct_tac "xs", auto) lemma filter_False [simp]: "xs:list(A) ==> filter(%p. False, xs) = Nil" by (induct_tac "xs", auto) lemma filter_True [simp]: "xs:list(A) ==> filter(%p. True, xs) = xs" by (induct_tac "xs", auto) (** length **) lemma length_is_0_iff [simp]: "xs:list(A) ==> length(xs)=0 \ xs=Nil" by (erule list.induct, auto) lemma length_is_0_iff2 [simp]: "xs:list(A) ==> 0 = length(xs) \ xs=Nil" by (erule list.induct, auto) lemma length_tl [simp]: "xs:list(A) ==> length(tl(xs)) = length(xs) #- 1" by (erule list.induct, auto) lemma length_greater_0_iff: "xs:list(A) ==> 0 xs \ Nil" by (erule list.induct, auto) lemma length_succ_iff: "xs:list(A) ==> length(xs)=succ(n) \ (\y ys. xs=Cons(y, ys) & length(ys)=n)" by (erule list.induct, auto) (** more theorems about append **) lemma append_is_Nil_iff [simp]: "xs:list(A) ==> (xs@ys = Nil) \ (xs=Nil & ys = Nil)" by (erule list.induct, auto) lemma append_is_Nil_iff2 [simp]: "xs:list(A) ==> (Nil = xs@ys) \ (xs=Nil & ys = Nil)" by (erule list.induct, auto) lemma append_left_is_self_iff [simp]: "xs:list(A) ==> (xs@ys = xs) \ (ys = Nil)" by (erule list.induct, auto) lemma append_left_is_self_iff2 [simp]: "xs:list(A) ==> (xs = xs@ys) \ (ys = Nil)" by (erule list.induct, auto) (*TOO SLOW as a default simprule!*) lemma append_left_is_Nil_iff [rule_format]: "[| xs:list(A); ys:list(A); zs:list(A) |] ==> length(ys)=length(zs) \ (xs@ys=zs \ (xs=Nil & ys=zs))" apply (erule list.induct) apply (auto simp add: length_app) done (*TOO SLOW as a default simprule!*) lemma append_left_is_Nil_iff2 [rule_format]: "[| xs:list(A); ys:list(A); zs:list(A) |] ==> length(ys)=length(zs) \ (zs=ys@xs \ (xs=Nil & ys=zs))" apply (erule list.induct) apply (auto simp add: length_app) done lemma append_eq_append_iff [rule_format]: "xs:list(A) ==> \ys \ list(A). length(xs)=length(ys) \ (xs@us = ys@vs) \ (xs=ys & us=vs)" apply (erule list.induct) apply (simp (no_asm_simp)) apply clarify apply (erule_tac a = ys in list.cases, auto) done declare append_eq_append_iff [simp] lemma append_eq_append [rule_format]: "xs:list(A) ==> \ys \ list(A). \us \ list(A). \vs \ list(A). length(us) = length(vs) \ (xs@us = ys@vs) \ (xs=ys & us=vs)" apply (induct_tac "xs") apply (force simp add: length_app, clarify) apply (erule_tac a = ys in list.cases, simp) apply (subgoal_tac "Cons (a, l) @ us =vs") apply (drule rev_iffD1 [OF _ append_left_is_Nil_iff], simp_all, blast) done lemma append_eq_append_iff2 [simp]: "[| xs:list(A); ys:list(A); us:list(A); vs:list(A); length(us)=length(vs) |] ==> xs@us = ys@vs \ (xs=ys & us=vs)" apply (rule iffI) apply (rule append_eq_append, auto) done lemma append_self_iff [simp]: "[| xs:list(A); ys:list(A); zs:list(A) |] ==> xs@ys=xs@zs \ ys=zs" by simp lemma append_self_iff2 [simp]: "[| xs:list(A); ys:list(A); zs:list(A) |] ==> ys@xs=zs@xs \ ys=zs" by simp (* Can also be proved from append_eq_append_iff2, but the proof requires two more hypotheses: x \ A and y \ A *) lemma append1_eq_iff [rule_format]: "xs:list(A) ==> \ys \ list(A). xs@[x] = ys@[y] \ (xs = ys & x=y)" apply (erule list.induct) apply clarify apply (erule list.cases) apply simp_all txt\Inductive step\ apply clarify apply (erule_tac a=ys in list.cases, simp_all) done declare append1_eq_iff [simp] lemma append_right_is_self_iff [simp]: "[| xs:list(A); ys:list(A) |] ==> (xs@ys = ys) \ (xs=Nil)" by (simp (no_asm_simp) add: append_left_is_Nil_iff) lemma append_right_is_self_iff2 [simp]: "[| xs:list(A); ys:list(A) |] ==> (ys = xs@ys) \ (xs=Nil)" apply (rule iffI) apply (drule sym, auto) done lemma hd_append [rule_format]: "xs:list(A) ==> xs \ Nil \ hd(xs @ ys) = hd(xs)" by (induct_tac "xs", auto) declare hd_append [simp] lemma tl_append [rule_format]: "xs:list(A) ==> xs\Nil \ tl(xs @ ys) = tl(xs)@ys" by (induct_tac "xs", auto) declare tl_append [simp] (** rev **) lemma rev_is_Nil_iff [simp]: "xs:list(A) ==> (rev(xs) = Nil \ xs = Nil)" by (erule list.induct, auto) lemma Nil_is_rev_iff [simp]: "xs:list(A) ==> (Nil = rev(xs) \ xs = Nil)" by (erule list.induct, auto) lemma rev_is_rev_iff [rule_format]: "xs:list(A) ==> \ys \ list(A). rev(xs)=rev(ys) \ xs=ys" apply (erule list.induct, force, clarify) apply (erule_tac a = ys in list.cases, auto) done declare rev_is_rev_iff [simp] lemma rev_list_elim [rule_format]: "xs:list(A) ==> (xs=Nil \ P) \ (\ys \ list(A). \y \ A. xs =ys@[y] \P)\P" by (erule list_append_induct, auto) (** more theorems about drop **) lemma length_drop [rule_format]: "n \ nat ==> \xs \ list(A). length(drop(n, xs)) = length(xs) #- n" apply (erule nat_induct) apply (auto elim: list.cases) done declare length_drop [simp] lemma drop_all [rule_format]: "n \ nat ==> \xs \ list(A). length(xs) \ n \ drop(n, xs)=Nil" apply (erule nat_induct) apply (auto elim: list.cases) done declare drop_all [simp] lemma drop_append [rule_format]: "n \ nat ==> \xs \ list(A). drop(n, xs@ys) = drop(n,xs) @ drop(n #- length(xs), ys)" apply (induct_tac "n") apply (auto elim: list.cases) done lemma drop_drop: "m \ nat ==> \xs \ list(A). \n \ nat. drop(n, drop(m, xs))=drop(n #+ m, xs)" apply (induct_tac "m") apply (auto elim: list.cases) done (** take **) lemma take_0 [simp]: "xs:list(A) ==> take(0, xs) = Nil" apply (unfold take_def) apply (erule list.induct, auto) done lemma take_succ_Cons [simp]: "n \ nat ==> take(succ(n), Cons(a, xs)) = Cons(a, take(n, xs))" by (simp add: take_def) (* Needed for proving take_all *) lemma take_Nil [simp]: "n \ nat ==> take(n, Nil) = Nil" by (unfold take_def, auto) lemma take_all [rule_format]: "n \ nat ==> \xs \ list(A). length(xs) \ n \ take(n, xs) = xs" apply (erule nat_induct) apply (auto elim: list.cases) done declare take_all [simp] lemma take_type [rule_format]: "xs:list(A) ==> \n \ nat. take(n, xs):list(A)" apply (erule list.induct, simp, clarify) apply (erule natE, auto) done declare take_type [simp,TC] lemma take_append [rule_format]: "xs:list(A) ==> \ys \ list(A). \n \ nat. take(n, xs @ ys) = take(n, xs) @ take(n #- length(xs), ys)" apply (erule list.induct, simp, clarify) apply (erule natE, auto) done declare take_append [simp] lemma take_take [rule_format]: "m \ nat ==> \xs \ list(A). \n \ nat. take(n, take(m,xs))= take(min(n, m), xs)" apply (induct_tac "m", auto) apply (erule_tac a = xs in list.cases) apply (auto simp add: take_Nil) apply (erule_tac n=n in natE) apply (auto intro: take_0 take_type) done (** nth **) lemma nth_0 [simp]: "nth(0, Cons(a, l)) = a" by (simp add: nth_def) lemma nth_Cons [simp]: "n \ nat ==> nth(succ(n), Cons(a,l)) = nth(n,l)" by (simp add: nth_def) lemma nth_empty [simp]: "nth(n, Nil) = 0" by (simp add: nth_def) lemma nth_type [rule_format]: "xs:list(A) ==> \n. n < length(xs) \ nth(n,xs) \ A" apply (erule list.induct, simp, clarify) apply (subgoal_tac "n \ nat") apply (erule natE, auto dest!: le_in_nat) done declare nth_type [simp,TC] lemma nth_eq_0 [rule_format]: "xs:list(A) ==> \n \ nat. length(xs) \ n \ nth(n,xs) = 0" apply (erule list.induct, simp, clarify) apply (erule natE, auto) done lemma nth_append [rule_format]: "xs:list(A) ==> \n \ nat. nth(n, xs @ ys) = (if n < length(xs) then nth(n,xs) else nth(n #- length(xs), ys))" apply (induct_tac "xs", simp, clarify) apply (erule natE, auto) done lemma set_of_list_conv_nth: "xs:list(A) ==> set_of_list(xs) = {x \ A. \i\nat. i nat ==> \xs \ list(A). (\ys \ list(A). k \ length(xs) \ k \ length(ys) \ (\i \ nat. i nth(i,xs) = nth(i,ys))\ take(k,xs) = take(k,ys))" apply (induct_tac "k") apply (simp_all (no_asm_simp) add: lt_succ_eq_0_disj all_conj_distrib) apply clarify (*Both lists are non-empty*) apply (erule_tac a=xs in list.cases, simp) apply (erule_tac a=ys in list.cases, clarify) apply (simp (no_asm_use) ) apply clarify apply (simp (no_asm_simp)) apply (rule conjI, force) apply (rename_tac y ys z zs) apply (drule_tac x = zs and x1 = ys in bspec [THEN bspec], auto) done lemma nth_equalityI [rule_format]: "[| xs:list(A); ys:list(A); length(xs) = length(ys); \i \ nat. i < length(xs) \ nth(i,xs) = nth(i,ys) |] ==> xs = ys" apply (subgoal_tac "length (xs) \ length (ys) ") apply (cut_tac k="length(xs)" and xs=xs and ys=ys in nth_take_lemma) apply (simp_all add: take_all) done (*The famous take-lemma*) lemma take_equalityI [rule_format]: "[| xs:list(A); ys:list(A); (\i \ nat. take(i, xs) = take(i,ys)) |] ==> xs = ys" apply (case_tac "length (xs) \ length (ys) ") apply (drule_tac x = "length (ys) " in bspec) apply (drule_tac [3] not_lt_imp_le) apply (subgoal_tac [5] "length (ys) \ length (xs) ") apply (rule_tac [6] j = "succ (length (ys))" in le_trans) apply (rule_tac [6] leI) apply (drule_tac [5] x = "length (xs) " in bspec) apply (simp_all add: take_all) done lemma nth_drop [rule_format]: "n \ nat ==> \i \ nat. \xs \ list(A). nth(i, drop(n, xs)) = nth(n #+ i, xs)" apply (induct_tac "n", simp_all, clarify) apply (erule list.cases, auto) done lemma take_succ [rule_format]: "xs\list(A) ==> \i. i < length(xs) \ take(succ(i), xs) = take(i,xs) @ [nth(i, xs)]" apply (induct_tac "xs", auto) apply (subgoal_tac "i\nat") apply (erule natE) apply (auto simp add: le_in_nat) done lemma take_add [rule_format]: "[|xs\list(A); j\nat|] ==> \i\nat. take(i #+ j, xs) = take(i,xs) @ take(j, drop(i,xs))" apply (induct_tac "xs", simp_all, clarify) apply (erule_tac n = i in natE, simp_all) done lemma length_take: "l\list(A) ==> \n\nat. length(take(n,l)) = min(n, length(l))" apply (induct_tac "l", safe, simp_all) apply (erule natE, simp_all) done subsection\The function zip\ text\Crafty definition to eliminate a type argument\ consts zip_aux :: "[i,i]=>i" primrec (*explicit lambda is required because both arguments of "un" vary*) "zip_aux(B,[]) = (\ys \ list(B). list_case([], %y l. [], ys))" "zip_aux(B,Cons(x,l)) = (\ys \ list(B). list_case(Nil, %y zs. Cons(, zip_aux(B,l)`zs), ys))" definition zip :: "[i, i]=>i" where "zip(xs, ys) == zip_aux(set_of_list(ys),xs)`ys" (* zip equations *) lemma list_on_set_of_list: "xs \ list(A) ==> xs \ list(set_of_list(xs))" apply (induct_tac xs, simp_all) apply (blast intro: list_mono [THEN subsetD]) done lemma zip_Nil [simp]: "ys:list(A) ==> zip(Nil, ys)=Nil" apply (simp add: zip_def list_on_set_of_list [of _ A]) apply (erule list.cases, simp_all) done lemma zip_Nil2 [simp]: "xs:list(A) ==> zip(xs, Nil)=Nil" apply (simp add: zip_def list_on_set_of_list [of _ A]) apply (erule list.cases, simp_all) done lemma zip_aux_unique [rule_format]: "[|B<=C; xs \ list(A)|] ==> \ys \ list(B). zip_aux(C,xs) ` ys = zip_aux(B,xs) ` ys" apply (induct_tac xs) apply simp_all apply (blast intro: list_mono [THEN subsetD], clarify) apply (erule_tac a=ys in list.cases, auto) apply (blast intro: list_mono [THEN subsetD]) done lemma zip_Cons_Cons [simp]: "[| xs:list(A); ys:list(B); x \ A; y \ B |] ==> zip(Cons(x,xs), Cons(y, ys)) = Cons(, zip(xs, ys))" apply (simp add: zip_def, auto) apply (rule zip_aux_unique, auto) apply (simp add: list_on_set_of_list [of _ B]) apply (blast intro: list_on_set_of_list list_mono [THEN subsetD]) done lemma zip_type [rule_format]: "xs:list(A) ==> \ys \ list(B). zip(xs, ys):list(A*B)" apply (induct_tac "xs") apply (simp (no_asm)) apply clarify apply (erule_tac a = ys in list.cases, auto) done declare zip_type [simp,TC] (* zip length *) lemma length_zip [rule_format]: "xs:list(A) ==> \ys \ list(B). length(zip(xs,ys)) = min(length(xs), length(ys))" apply (unfold min_def) apply (induct_tac "xs", simp_all, clarify) apply (erule_tac a = ys in list.cases, auto) done declare length_zip [simp] lemma zip_append1 [rule_format]: "[| ys:list(A); zs:list(B) |] ==> \xs \ list(A). zip(xs @ ys, zs) = zip(xs, take(length(xs), zs)) @ zip(ys, drop(length(xs),zs))" apply (induct_tac "zs", force, clarify) apply (erule_tac a = xs in list.cases, simp_all) done lemma zip_append2 [rule_format]: "[| xs:list(A); zs:list(B) |] ==> \ys \ list(B). zip(xs, ys@zs) = zip(take(length(ys), xs), ys) @ zip(drop(length(ys), xs), zs)" apply (induct_tac "xs", force, clarify) apply (erule_tac a = ys in list.cases, auto) done lemma zip_append [simp]: "[| length(xs) = length(us); length(ys) = length(vs); xs:list(A); us:list(B); ys:list(A); vs:list(B) |] ==> zip(xs@ys,us@vs) = zip(xs, us) @ zip(ys, vs)" by (simp (no_asm_simp) add: zip_append1 drop_append diff_self_eq_0) lemma zip_rev [rule_format]: "ys:list(B) ==> \xs \ list(A). length(xs) = length(ys) \ zip(rev(xs), rev(ys)) = rev(zip(xs, ys))" apply (induct_tac "ys", force, clarify) apply (erule_tac a = xs in list.cases) apply (auto simp add: length_rev) done declare zip_rev [simp] lemma nth_zip [rule_format]: "ys:list(B) ==> \i \ nat. \xs \ list(A). i < length(xs) \ i < length(ys) \ nth(i,zip(xs, ys)) = " apply (induct_tac "ys", force, clarify) apply (erule_tac a = xs in list.cases, simp) apply (auto elim: natE) done declare nth_zip [simp] lemma set_of_list_zip [rule_format]: "[| xs:list(A); ys:list(B); i \ nat |] ==> set_of_list(zip(xs, ys)) = {:A*B. \i\nat. i < min(length(xs), length(ys)) & x = nth(i, xs) & y = nth(i, ys)}" by (force intro!: Collect_cong simp add: lt_min_iff set_of_list_conv_nth) (** list_update **) lemma list_update_Nil [simp]: "i \ nat ==>list_update(Nil, i, v) = Nil" by (unfold list_update_def, auto) lemma list_update_Cons_0 [simp]: "list_update(Cons(x, xs), 0, v)= Cons(v, xs)" by (unfold list_update_def, auto) lemma list_update_Cons_succ [simp]: "n \ nat ==> list_update(Cons(x, xs), succ(n), v)= Cons(x, list_update(xs, n, v))" apply (unfold list_update_def, auto) done lemma list_update_type [rule_format]: "[| xs:list(A); v \ A |] ==> \n \ nat. list_update(xs, n, v):list(A)" apply (induct_tac "xs") apply (simp (no_asm)) apply clarify apply (erule natE, auto) done declare list_update_type [simp,TC] lemma length_list_update [rule_format]: "xs:list(A) ==> \i \ nat. length(list_update(xs, i, v))=length(xs)" apply (induct_tac "xs") apply (simp (no_asm)) apply clarify apply (erule natE, auto) done declare length_list_update [simp] lemma nth_list_update [rule_format]: "[| xs:list(A) |] ==> \i \ nat. \j \ nat. i < length(xs) \ nth(j, list_update(xs, i, x)) = (if i=j then x else nth(j, xs))" apply (induct_tac "xs") apply simp_all apply clarify apply (rename_tac i j) apply (erule_tac n=i in natE) apply (erule_tac [2] n=j in natE) apply (erule_tac n=j in natE, simp_all, force) done lemma nth_list_update_eq [simp]: "[| i < length(xs); xs:list(A) |] ==> nth(i, list_update(xs, i,x)) = x" by (simp (no_asm_simp) add: lt_length_in_nat nth_list_update) lemma nth_list_update_neq [rule_format]: "xs:list(A) ==> \i \ nat. \j \ nat. i \ j \ nth(j, list_update(xs,i,x)) = nth(j,xs)" apply (induct_tac "xs") apply (simp (no_asm)) apply clarify apply (erule natE) apply (erule_tac [2] natE, simp_all) apply (erule natE, simp_all) done declare nth_list_update_neq [simp] lemma list_update_overwrite [rule_format]: "xs:list(A) ==> \i \ nat. i < length(xs) \ list_update(list_update(xs, i, x), i, y) = list_update(xs, i,y)" apply (induct_tac "xs") apply (simp (no_asm)) apply clarify apply (erule natE, auto) done declare list_update_overwrite [simp] lemma list_update_same_conv [rule_format]: "xs:list(A) ==> \i \ nat. i < length(xs) \ (list_update(xs, i, x) = xs) \ (nth(i, xs) = x)" apply (induct_tac "xs") apply (simp (no_asm)) apply clarify apply (erule natE, auto) done lemma update_zip [rule_format]: "ys:list(B) ==> \i \ nat. \xy \ A*B. \xs \ list(A). length(xs) = length(ys) \ list_update(zip(xs, ys), i, xy) = zip(list_update(xs, i, fst(xy)), list_update(ys, i, snd(xy)))" apply (induct_tac "ys") apply auto apply (erule_tac a = xs in list.cases) apply (auto elim: natE) done lemma set_update_subset_cons [rule_format]: "xs:list(A) ==> \i \ nat. set_of_list(list_update(xs, i, x)) \ cons(x, set_of_list(xs))" apply (induct_tac "xs") apply simp apply (rule ballI) apply (erule natE, simp_all, auto) done lemma set_of_list_update_subsetI: "[| set_of_list(xs) \ A; xs:list(A); x \ A; i \ nat|] ==> set_of_list(list_update(xs, i,x)) \ A" apply (rule subset_trans) apply (rule set_update_subset_cons, auto) done (** upt **) lemma upt_rec: "j \ nat ==> upt(i,j) = (if i i; j \ nat |] ==> upt(i,j) = Nil" apply (subst upt_rec, auto) apply (auto simp add: le_iff) apply (drule lt_asym [THEN notE], auto) done (*Only needed if upt_Suc is deleted from the simpset*) lemma upt_succ_append: "[| i \ j; j \ nat |] ==> upt(i,succ(j)) = upt(i, j)@[j]" by simp lemma upt_conv_Cons: "[| i nat |] ==> upt(i,j) = Cons(i,upt(succ(i),j))" apply (rule trans) apply (rule upt_rec, auto) done lemma upt_type [simp,TC]: "j \ nat ==> upt(i,j):list(nat)" by (induct_tac "j", auto) (*LOOPS as a simprule, since j<=j*) lemma upt_add_eq_append: "[| i \ j; j \ nat; k \ nat |] ==> upt(i, j #+k) = upt(i,j)@upt(j,j#+k)" apply (induct_tac "k") apply (auto simp add: app_assoc app_type) apply (rule_tac j = j in le_trans, auto) done lemma length_upt [simp]: "[| i \ nat; j \ nat |] ==>length(upt(i,j)) = j #- i" apply (induct_tac "j") apply (rule_tac [2] sym) apply (auto dest!: not_lt_imp_le simp add: diff_succ diff_is_0_iff) done -lemma nth_upt [rule_format]: - "[| i \ nat; j \ nat; k \ nat |] ==> i #+ k < j \ nth(k, upt(i,j)) = i #+ k" +lemma nth_upt [simp]: + "[| i \ nat; j \ nat; k \ nat; i #+ k < j |] ==> nth(k, upt(i,j)) = i #+ k" +apply (rotate_tac -1, erule rev_mp) apply (induct_tac "j", simp) -apply (simp add: nth_append le_iff) apply (auto dest!: not_lt_imp_le - simp add: nth_append less_diff_conv add_commute) + simp add: nth_append le_iff less_diff_conv add_commute) done -declare nth_upt [simp] lemma take_upt [rule_format]: "[| m \ nat; n \ nat |] ==> \i \ nat. i #+ m \ n \ take(m, upt(i,n)) = upt(i,i#+m)" apply (induct_tac "m") apply (simp (no_asm_simp) add: take_0) apply clarify apply (subst upt_rec, simp) apply (rule sym) apply (subst upt_rec, simp) apply (simp_all del: upt.simps) apply (rule_tac j = "succ (i #+ x) " in lt_trans2) apply auto done declare take_upt [simp] lemma map_succ_upt: "[| m \ nat; n \ nat |] ==> map(succ, upt(m,n))= upt(succ(m), succ(n))" apply (induct_tac "n") apply (auto simp add: map_app_distrib) done lemma nth_map [rule_format]: "xs:list(A) ==> \n \ nat. n < length(xs) \ nth(n, map(f, xs)) = f(nth(n, xs))" apply (induct_tac "xs", simp) apply (rule ballI) apply (induct_tac "n", auto) done declare nth_map [simp] lemma nth_map_upt [rule_format]: "[| m \ nat; n \ nat |] ==> \i \ nat. i < n #- m \ nth(i, map(f, upt(m,n))) = f(m #+ i)" apply (rule_tac n = m and m = n in diff_induct, typecheck, simp, simp) apply (subst map_succ_upt [symmetric], simp_all, clarify) apply (subgoal_tac "i < length (upt (0, x))") prefer 2 apply (simp add: less_diff_conv) apply (rule_tac j = "succ (i #+ y) " in lt_trans2) apply simp apply simp apply (subgoal_tac "i < length (upt (y, x))") apply (simp_all add: add_commute less_diff_conv) done (** sublist (a generalization of nth to sets) **) definition sublist :: "[i, i] => i" where "sublist(xs, A)== map(fst, (filter(%p. snd(p): A, zip(xs, upt(0,length(xs))))))" lemma sublist_0 [simp]: "xs:list(A) ==>sublist(xs, 0) =Nil" by (unfold sublist_def, auto) lemma sublist_Nil [simp]: "sublist(Nil, A) = Nil" by (unfold sublist_def, auto) lemma sublist_shift_lemma: "[| xs:list(B); i \ nat |] ==> map(fst, filter(%p. snd(p):A, zip(xs, upt(i,i #+ length(xs))))) = map(fst, filter(%p. snd(p):nat & snd(p) #+ i \ A, zip(xs,upt(0,length(xs)))))" apply (erule list_append_induct) apply (simp (no_asm_simp)) apply (auto simp add: add_commute length_app filter_append map_app_distrib) done lemma sublist_type [simp,TC]: "xs:list(B) ==> sublist(xs, A):list(B)" apply (unfold sublist_def) apply (induct_tac "xs") apply (auto simp add: filter_append map_app_distrib) done lemma upt_add_eq_append2: "[| i \ nat; j \ nat |] ==> upt(0, i #+ j) = upt(0, i) @ upt(i, i #+ j)" by (simp add: upt_add_eq_append [of 0] nat_0_le) lemma sublist_append: "[| xs:list(B); ys:list(B) |] ==> sublist(xs@ys, A) = sublist(xs, A) @ sublist(ys, {j \ nat. j #+ length(xs): A})" apply (unfold sublist_def) apply (erule_tac l = ys in list_append_induct, simp) apply (simp (no_asm_simp) add: upt_add_eq_append2 app_assoc [symmetric]) apply (auto simp add: sublist_shift_lemma length_type map_app_distrib app_assoc) apply (simp_all add: add_commute) done lemma sublist_Cons: "[| xs:list(B); x \ B |] ==> sublist(Cons(x, xs), A) = (if 0 \ A then [x] else []) @ sublist(xs, {j \ nat. succ(j) \ A})" apply (erule_tac l = xs in list_append_induct) apply (simp (no_asm_simp) add: sublist_def) apply (simp del: app_Cons add: app_Cons [symmetric] sublist_append, simp) done lemma sublist_singleton [simp]: "sublist([x], A) = (if 0 \ A then [x] else [])" by (simp add: sublist_Cons) lemma sublist_upt_eq_take [rule_format]: "xs:list(A) ==> \n\nat. sublist(xs,n) = take(n,xs)" apply (erule list.induct, simp) apply (clarify ) apply (erule natE) apply (simp_all add: nat_eq_Collect_lt Ord_mem_iff_lt sublist_Cons) done declare sublist_upt_eq_take [simp] lemma sublist_Int_eq: "xs \ list(B) ==> sublist(xs, A \ nat) = sublist(xs, A)" apply (erule list.induct) apply (simp_all add: sublist_Cons) done text\Repetition of a List Element\ consts repeat :: "[i,i]=>i" primrec "repeat(a,0) = []" "repeat(a,succ(n)) = Cons(a,repeat(a,n))" lemma length_repeat: "n \ nat ==> length(repeat(a,n)) = n" by (induct_tac n, auto) lemma repeat_succ_app: "n \ nat ==> repeat(a,succ(n)) = repeat(a,n) @ [a]" apply (induct_tac n) apply (simp_all del: app_Cons add: app_Cons [symmetric]) done lemma repeat_type [TC]: "[|a \ A; n \ nat|] ==> repeat(a,n) \ list(A)" by (induct_tac n, auto) end diff --git a/src/ZF/Nat.thy b/src/ZF/Nat.thy --- a/src/ZF/Nat.thy +++ b/src/ZF/Nat.thy @@ -1,302 +1,297 @@ (* Title: ZF/Nat.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge *) section\The Natural numbers As a Least Fixed Point\ theory Nat imports OrdQuant Bool begin definition nat :: i where "nat == lfp(Inf, %X. {0} \ {succ(i). i \ X})" definition quasinat :: "i => o" where "quasinat(n) == n=0 | (\m. n = succ(m))" definition (*Has an unconditional succ case, which is used in "recursor" below.*) nat_case :: "[i, i=>i, i]=>i" where "nat_case(a,b,k) == THE y. k=0 & y=a | (\x. k=succ(x) & y=b(x))" definition nat_rec :: "[i, i, [i,i]=>i]=>i" where "nat_rec(k,a,b) == wfrec(Memrel(nat), k, %n f. nat_case(a, %m. b(m, f`m), n))" (*Internalized relations on the naturals*) definition Le :: i where "Le == {:nat*nat. x \ y}" definition Lt :: i where "Lt == {:nat*nat. x < y}" definition Ge :: i where "Ge == {:nat*nat. y \ x}" definition Gt :: i where "Gt == {:nat*nat. y < x}" definition greater_than :: "i=>i" where "greater_than(n) == {i \ nat. n < i}" text\No need for a less-than operator: a natural number is its list of predecessors!\ lemma nat_bnd_mono: "bnd_mono(Inf, %X. {0} \ {succ(i). i \ X})" apply (rule bnd_monoI) apply (cut_tac infinity, blast, blast) done (* @{term"nat = {0} \ {succ(x). x \ nat}"} *) lemmas nat_unfold = nat_bnd_mono [THEN nat_def [THEN def_lfp_unfold]] (** Type checking of 0 and successor **) lemma nat_0I [iff,TC]: "0 \ nat" apply (subst nat_unfold) apply (rule singletonI [THEN UnI1]) done lemma nat_succI [intro!,TC]: "n \ nat ==> succ(n) \ nat" apply (subst nat_unfold) apply (erule RepFunI [THEN UnI2]) done lemma nat_1I [iff,TC]: "1 \ nat" by (rule nat_0I [THEN nat_succI]) lemma nat_2I [iff,TC]: "2 \ nat" by (rule nat_1I [THEN nat_succI]) lemma bool_subset_nat: "bool \ nat" by (blast elim!: boolE) lemmas bool_into_nat = bool_subset_nat [THEN subsetD] subsection\Injectivity Properties and Induction\ (*Mathematical induction*) lemma nat_induct [case_names 0 succ, induct set: nat]: "[| n \ nat; P(0); !!x. [| x \ nat; P(x) |] ==> P(succ(x)) |] ==> P(n)" by (erule def_induct [OF nat_def nat_bnd_mono], blast) lemma natE: assumes "n \ nat" obtains ("0") "n=0" | (succ) x where "x \ nat" "n=succ(x)" using assms by (rule nat_unfold [THEN equalityD1, THEN subsetD, THEN UnE]) auto lemma nat_into_Ord [simp]: "n \ nat ==> Ord(n)" by (erule nat_induct, auto) (* @{term"i \ nat ==> 0 \ i"}; same thing as @{term"0 nat ==> i \ i"}; same thing as @{term"i nat ==> ~ Limit(a)" by (induct a rule: nat_induct, auto) lemma succ_natD: "succ(i): nat ==> i \ nat" by (rule Ord_trans [OF succI1], auto) lemma nat_succ_iff [iff]: "succ(n): nat \ n \ nat" by (blast dest!: succ_natD) lemma nat_le_Limit: "Limit(i) ==> nat \ i" apply (rule subset_imp_le) apply (simp_all add: Limit_is_Ord) apply (rule subsetI) apply (erule nat_induct) apply (erule Limit_has_0 [THEN ltD]) apply (blast intro: Limit_has_succ [THEN ltD] ltI Limit_is_Ord) done (* [| succ(i): k; k \ nat |] ==> i \ k *) lemmas succ_in_naturalD = Ord_trans [OF succI1 _ nat_into_Ord] lemma lt_nat_in_nat: "[| m nat |] ==> m \ nat" apply (erule ltE) apply (erule Ord_trans, assumption, simp) done lemma le_in_nat: "[| m \ n; n \ nat |] ==> m \ nat" by (blast dest!: lt_nat_in_nat) subsection\Variations on Mathematical Induction\ (*complete induction*) lemmas complete_induct = Ord_induct [OF _ Ord_nat, case_names less, consumes 1] -lemmas complete_induct_rule = - complete_induct [rule_format, case_names less, consumes 1] - - -lemma nat_induct_from_lemma [rule_format]: - "[| n \ nat; m \ nat; - !!x. [| x \ nat; m \ x; P(x) |] ==> P(succ(x)) |] - ==> m \ n \ P(m) \ P(n)" -apply (erule nat_induct) -apply (simp_all add: distrib_simps le0_iff le_succ_iff) -done +lemma complete_induct_rule [case_names less, consumes 1]: + "i \ nat \ (\x. x \ nat \ (\y. y \ x \ P(y)) \ P(x)) \ P(i)" + using complete_induct [of i P] by simp (*Induction starting from m rather than 0*) lemma nat_induct_from: - "[| m \ n; m \ nat; n \ nat; - P(m); - !!x. [| x \ nat; m \ x; P(x) |] ==> P(succ(x)) |] - ==> P(n)" -apply (blast intro: nat_induct_from_lemma) -done + assumes "m \ n" "m \ nat" "n \ nat" + and "P(m)" + and "!!x. [| x \ nat; m \ x; P(x) |] ==> P(succ(x))" + shows "P(n)" +proof - + from assms(3) have "m \ n \ P(m) \ P(n)" + by (rule nat_induct) (use assms(5) in \simp_all add: distrib_simps le_succ_iff\) + with assms(1,2,4) show ?thesis by blast +qed (*Induction suitable for subtraction and less-than*) lemma diff_induct [case_names 0 0_succ succ_succ, consumes 2]: "[| m \ nat; n \ nat; !!x. x \ nat ==> P(x,0); !!y. y \ nat ==> P(0,succ(y)); !!x y. [| x \ nat; y \ nat; P(x,y) |] ==> P(succ(x),succ(y)) |] ==> P(m,n)" apply (erule_tac x = m in rev_bspec) apply (erule nat_induct, simp) apply (rule ballI) apply (rename_tac i j) apply (erule_tac n=j in nat_induct, auto) done (** Induction principle analogous to trancl_induct **) lemma succ_lt_induct_lemma [rule_format]: "m \ nat ==> P(m,succ(m)) \ (\x\nat. P(m,x) \ P(m,succ(x))) \ (\n\nat. m P(m,n))" apply (erule nat_induct) apply (intro impI, rule nat_induct [THEN ballI]) prefer 4 apply (intro impI, rule nat_induct [THEN ballI]) apply (auto simp add: le_iff) done lemma succ_lt_induct: "[| m nat; P(m,succ(m)); !!x. [| x \ nat; P(m,x) |] ==> P(m,succ(x)) |] ==> P(m,n)" by (blast intro: succ_lt_induct_lemma lt_nat_in_nat) subsection\quasinat: to allow a case-split rule for \<^term>\nat_case\\ text\True if the argument is zero or any successor\ lemma [iff]: "quasinat(0)" by (simp add: quasinat_def) lemma [iff]: "quasinat(succ(x))" by (simp add: quasinat_def) lemma nat_imp_quasinat: "n \ nat ==> quasinat(n)" by (erule natE, simp_all) lemma non_nat_case: "~ quasinat(x) ==> nat_case(a,b,x) = 0" by (simp add: quasinat_def nat_case_def) lemma nat_cases_disj: "k=0 | (\y. k = succ(y)) | ~ quasinat(k)" apply (case_tac "k=0", simp) apply (case_tac "\m. k = succ(m)") apply (simp_all add: quasinat_def) done lemma nat_cases: "[|k=0 ==> P; !!y. k = succ(y) ==> P; ~ quasinat(k) ==> P|] ==> P" by (insert nat_cases_disj [of k], blast) (** nat_case **) lemma nat_case_0 [simp]: "nat_case(a,b,0) = a" by (simp add: nat_case_def) lemma nat_case_succ [simp]: "nat_case(a,b,succ(n)) = b(n)" by (simp add: nat_case_def) lemma nat_case_type [TC]: "[| n \ nat; a \ C(0); !!m. m \ nat ==> b(m): C(succ(m)) |] ==> nat_case(a,b,n) \ C(n)" by (erule nat_induct, auto) lemma split_nat_case: "P(nat_case(a,b,k)) \ ((k=0 \ P(a)) & (\x. k=succ(x) \ P(b(x))) & (~ quasinat(k) \ P(0)))" apply (rule nat_cases [of k]) apply (auto simp add: non_nat_case) done subsection\Recursion on the Natural Numbers\ (** nat_rec is used to define eclose and transrec, then becomes obsolete. The operator rec, from arith.thy, has fewer typing conditions **) lemma nat_rec_0: "nat_rec(0,a,b) = a" apply (rule nat_rec_def [THEN def_wfrec, THEN trans]) apply (rule wf_Memrel) apply (rule nat_case_0) done lemma nat_rec_succ: "m \ nat ==> nat_rec(succ(m),a,b) = b(m, nat_rec(m,a,b))" apply (rule nat_rec_def [THEN def_wfrec, THEN trans]) apply (rule wf_Memrel) apply (simp add: vimage_singleton_iff) done (** The union of two natural numbers is a natural number -- their maximum **) lemma Un_nat_type [TC]: "[| i \ nat; j \ nat |] ==> i \ j \ nat" apply (rule Un_least_lt [THEN ltD]) apply (simp_all add: lt_def) done lemma Int_nat_type [TC]: "[| i \ nat; j \ nat |] ==> i \ j \ nat" apply (rule Int_greatest_lt [THEN ltD]) apply (simp_all add: lt_def) done (*needed to simplify unions over nat*) lemma nat_nonempty [simp]: "nat \ 0" by blast text\A natural number is the set of its predecessors\ lemma nat_eq_Collect_lt: "i \ nat ==> {j\nat. j \ Le \ x \ y & x \ nat & y \ nat" by (force simp add: Le_def) end diff --git a/src/ZF/Ordinal.thy b/src/ZF/Ordinal.thy --- a/src/ZF/Ordinal.thy +++ b/src/ZF/Ordinal.thy @@ -1,764 +1,765 @@ (* Title: ZF/Ordinal.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge *) section\Transitive Sets and Ordinals\ theory Ordinal imports WF Bool equalities begin definition Memrel :: "i=>i" where "Memrel(A) == {z\A*A . \x y. z= & x\y }" definition Transset :: "i=>o" where "Transset(i) == \x\i. x<=i" definition Ord :: "i=>o" where "Ord(i) == Transset(i) & (\x\i. Transset(x))" definition lt :: "[i,i] => o" (infixl \<\ 50) (*less-than on ordinals*) where "ij & Ord(j)" definition Limit :: "i=>o" where "Limit(i) == Ord(i) & 0y. y succ(y)\\ 50) where "x \ y == x < succ(y)" subsection\Rules for Transset\ subsubsection\Three Neat Characterisations of Transset\ lemma Transset_iff_Pow: "Transset(A) <-> A<=Pow(A)" by (unfold Transset_def, blast) lemma Transset_iff_Union_succ: "Transset(A) <-> \(succ(A)) = A" apply (unfold Transset_def) apply (blast elim!: equalityE) done lemma Transset_iff_Union_subset: "Transset(A) <-> \(A) \ A" by (unfold Transset_def, blast) subsubsection\Consequences of Downwards Closure\ lemma Transset_doubleton_D: "[| Transset(C); {a,b}: C |] ==> a\C & b\C" by (unfold Transset_def, blast) lemma Transset_Pair_D: "[| Transset(C); \C |] ==> a\C & b\C" apply (simp add: Pair_def) apply (blast dest: Transset_doubleton_D) done lemma Transset_includes_domain: "[| Transset(C); A*B \ C; b \ B |] ==> A \ C" by (blast dest: Transset_Pair_D) lemma Transset_includes_range: "[| Transset(C); A*B \ C; a \ A |] ==> B \ C" by (blast dest: Transset_Pair_D) subsubsection\Closure Properties\ lemma Transset_0: "Transset(0)" by (unfold Transset_def, blast) lemma Transset_Un: "[| Transset(i); Transset(j) |] ==> Transset(i \ j)" by (unfold Transset_def, blast) lemma Transset_Int: "[| Transset(i); Transset(j) |] ==> Transset(i \ j)" by (unfold Transset_def, blast) lemma Transset_succ: "Transset(i) ==> Transset(succ(i))" by (unfold Transset_def, blast) lemma Transset_Pow: "Transset(i) ==> Transset(Pow(i))" by (unfold Transset_def, blast) lemma Transset_Union: "Transset(A) ==> Transset(\(A))" by (unfold Transset_def, blast) lemma Transset_Union_family: "[| !!i. i\A ==> Transset(i) |] ==> Transset(\(A))" by (unfold Transset_def, blast) lemma Transset_Inter_family: "[| !!i. i\A ==> Transset(i) |] ==> Transset(\(A))" by (unfold Inter_def Transset_def, blast) lemma Transset_UN: "(!!x. x \ A ==> Transset(B(x))) ==> Transset (\x\A. B(x))" by (rule Transset_Union_family, auto) lemma Transset_INT: "(!!x. x \ A ==> Transset(B(x))) ==> Transset (\x\A. B(x))" by (rule Transset_Inter_family, auto) subsection\Lemmas for Ordinals\ lemma OrdI: "[| Transset(i); !!x. x\i ==> Transset(x) |] ==> Ord(i)" by (simp add: Ord_def) lemma Ord_is_Transset: "Ord(i) ==> Transset(i)" by (simp add: Ord_def) lemma Ord_contains_Transset: "[| Ord(i); j\i |] ==> Transset(j) " by (unfold Ord_def, blast) lemma Ord_in_Ord: "[| Ord(i); j\i |] ==> Ord(j)" by (unfold Ord_def Transset_def, blast) (*suitable for rewriting PROVIDED i has been fixed*) lemma Ord_in_Ord': "[| j\i; Ord(i) |] ==> Ord(j)" by (blast intro: Ord_in_Ord) (* Ord(succ(j)) ==> Ord(j) *) lemmas Ord_succD = Ord_in_Ord [OF _ succI1] lemma Ord_subset_Ord: "[| Ord(i); Transset(j); j<=i |] ==> Ord(j)" by (simp add: Ord_def Transset_def, blast) lemma OrdmemD: "[| j\i; Ord(i) |] ==> j<=i" by (unfold Ord_def Transset_def, blast) lemma Ord_trans: "[| i\j; j\k; Ord(k) |] ==> i\k" by (blast dest: OrdmemD) lemma Ord_succ_subsetI: "[| i\j; Ord(j) |] ==> succ(i) \ j" by (blast dest: OrdmemD) subsection\The Construction of Ordinals: 0, succ, Union\ lemma Ord_0 [iff,TC]: "Ord(0)" by (blast intro: OrdI Transset_0) lemma Ord_succ [TC]: "Ord(i) ==> Ord(succ(i))" by (blast intro: OrdI Transset_succ Ord_is_Transset Ord_contains_Transset) lemmas Ord_1 = Ord_0 [THEN Ord_succ] lemma Ord_succ_iff [iff]: "Ord(succ(i)) <-> Ord(i)" by (blast intro: Ord_succ dest!: Ord_succD) lemma Ord_Un [intro,simp,TC]: "[| Ord(i); Ord(j) |] ==> Ord(i \ j)" apply (unfold Ord_def) apply (blast intro!: Transset_Un) done lemma Ord_Int [TC]: "[| Ord(i); Ord(j) |] ==> Ord(i \ j)" apply (unfold Ord_def) apply (blast intro!: Transset_Int) done text\There is no set of all ordinals, for then it would contain itself\ lemma ON_class: "~ (\i. i\X <-> Ord(i))" proof (rule notI) assume X: "\i. i \ X \ Ord(i)" have "\x y. x\X \ y\x \ y\X" by (simp add: X, blast intro: Ord_in_Ord) hence "Transset(X)" by (auto simp add: Transset_def) moreover have "\x. x \ X \ Transset(x)" by (simp add: X Ord_def) ultimately have "Ord(X)" by (rule OrdI) hence "X \ X" by (simp add: X) thus "False" by (rule mem_irrefl) qed subsection\< is 'less Than' for Ordinals\ lemma ltI: "[| i\j; Ord(j) |] ==> ij; Ord(i); Ord(j) |] ==> P |] ==> P" apply (unfold lt_def) apply (blast intro: Ord_in_Ord) done lemma ltD: "i i\j" by (erule ltE, assumption) lemma not_lt0 [simp]: "~ i<0" by (unfold lt_def, blast) lemma lt_Ord: "j Ord(j)" by (erule ltE, assumption) lemma lt_Ord2: "j Ord(i)" by (erule ltE, assumption) (* @{term"ja \ j ==> Ord(j)"} *) lemmas le_Ord2 = lt_Ord2 [THEN Ord_succD] (* i<0 ==> R *) lemmas lt0E = not_lt0 [THEN notE, elim!] lemma lt_trans [trans]: "[| i i ~ (j j P *) lemmas lt_asym = lt_not_sym [THEN swap] lemma lt_irrefl [elim!]: "i P" by (blast intro: lt_asym) lemma lt_not_refl: "~ iRecall that \<^term>\i \ j\ abbreviates \<^term>\i !!\ lemma le_iff: "i \ j <-> i i < succ(j)*) lemma leI: "i i \ j" by (simp add: le_iff) lemma le_eqI: "[| i=j; Ord(j) |] ==> i \ j" by (simp add: le_iff) lemmas le_refl = refl [THEN le_eqI] lemma le_refl_iff [iff]: "i \ i <-> Ord(i)" by (simp (no_asm_simp) add: lt_not_refl le_iff) lemma leCI: "(~ (i=j & Ord(j)) ==> i i \ j" by (simp add: le_iff, blast) lemma leE: "[| i \ j; i P; [| i=j; Ord(j) |] ==> P |] ==> P" by (simp add: le_iff, blast) lemma le_anti_sym: "[| i \ j; j \ i |] ==> i=j" apply (simp add: le_iff) apply (blast elim: lt_asym) done lemma le0_iff [simp]: "i \ 0 <-> i=0" by (blast elim!: leE) lemmas le0D = le0_iff [THEN iffD1, dest!] subsection\Natural Deduction Rules for Memrel\ (*The lemmas MemrelI/E give better speed than [iff] here*) lemma Memrel_iff [simp]: " \ Memrel(A) <-> a\b & a\A & b\A" by (unfold Memrel_def, blast) lemma MemrelI [intro!]: "[| a \ b; a \ A; b \ A |] ==> \ Memrel(A)" by auto lemma MemrelE [elim!]: "[| \ Memrel(A); [| a \ A; b \ A; a\b |] ==> P |] ==> P" by auto lemma Memrel_type: "Memrel(A) \ A*A" by (unfold Memrel_def, blast) lemma Memrel_mono: "A<=B ==> Memrel(A) \ Memrel(B)" by (unfold Memrel_def, blast) lemma Memrel_0 [simp]: "Memrel(0) = 0" by (unfold Memrel_def, blast) lemma Memrel_1 [simp]: "Memrel(1) = 0" by (unfold Memrel_def, blast) lemma relation_Memrel: "relation(Memrel(A))" by (simp add: relation_def Memrel_def) (*The membership relation (as a set) is well-founded. Proof idea: show A<=B by applying the foundation axiom to A-B *) lemma wf_Memrel: "wf(Memrel(A))" apply (unfold wf_def) apply (rule foundation [THEN disjE, THEN allI], erule disjI1, blast) done text\The premise \<^term>\Ord(i)\ does not suffice.\ lemma trans_Memrel: "Ord(i) ==> trans(Memrel(i))" by (unfold Ord_def Transset_def trans_def, blast) text\However, the following premise is strong enough.\ lemma Transset_trans_Memrel: "\j\i. Transset(j) ==> trans(Memrel(i))" by (unfold Transset_def trans_def, blast) (*If Transset(A) then Memrel(A) internalizes the membership relation below A*) lemma Transset_Memrel_iff: "Transset(A) ==> \ Memrel(A) <-> a\b & b\A" by (unfold Transset_def, blast) subsection\Transfinite Induction\ (*Epsilon induction over a transitive set*) lemma Transset_induct: "[| i \ k; Transset(k); !!x.[| x \ k; \y\x. P(y) |] ==> P(x) |] ==> P(i)" apply (simp add: Transset_def) apply (erule wf_Memrel [THEN wf_induct2], blast+) done (*Induction over an ordinal*) -lemmas Ord_induct [consumes 2] = Transset_induct [rule_format, OF _ Ord_is_Transset] +lemma Ord_induct [consumes 2]: + "i \ k \ Ord(k) \ (\x. x \ k \ (\y. y \ x \ P(y)) \ P(x)) \ P(i)" + using Transset_induct [OF _ Ord_is_Transset, of i k P] by simp (*Induction over the class of ordinals -- a useful corollary of Ord_induct*) - -lemma trans_induct [rule_format, consumes 1, case_names step]: - "[| Ord(i); - !!x.[| Ord(x); \y\x. P(y) |] ==> P(x) |] - ==> P(i)" -apply (rule Ord_succ [THEN succI1 [THEN Ord_induct]], assumption) -apply (blast intro: Ord_succ [THEN Ord_in_Ord]) -done +lemma trans_induct [consumes 1, case_names step]: + "Ord(i) \ (\x. Ord(x) \ (\y. y \ x \ P(y)) \ P(x)) \ P(i)" + apply (rule Ord_succ [THEN succI1 [THEN Ord_induct]], assumption) + apply (blast intro: Ord_succ [THEN Ord_in_Ord]) + done section\Fundamental properties of the epsilon ordering (< on ordinals)\ subsubsection\Proving That < is a Linear Ordering on the Ordinals\ lemma Ord_linear: "Ord(i) \ Ord(j) \ i\j | i=j | j\i" proof (induct i arbitrary: j rule: trans_induct) case (step i) note step_i = step show ?case using \Ord(j)\ proof (induct j rule: trans_induct) case (step j) thus ?case using step_i by (blast dest: Ord_trans) qed qed text\The trichotomy law for ordinals\ lemma Ord_linear_lt: assumes o: "Ord(i)" "Ord(j)" obtains (lt) "i i" apply (rule_tac i = i and j = j in Ord_linear_lt) apply (blast intro: leI le_eqI sym o) + done lemma Ord_linear_le: assumes o: "Ord(i)" "Ord(j)" obtains (le) "i \ j" | (ge) "j \ i" apply (rule_tac i = i and j = j in Ord_linear_lt) apply (blast intro: leI le_eqI o) + done lemma le_imp_not_lt: "j \ i ==> ~ i j \ i" by (rule_tac i = i and j = j in Ord_linear2, auto) subsubsection \Some Rewrite Rules for \<\, \\\\ lemma Ord_mem_iff_lt: "Ord(j) ==> i\j <-> i ~ i j \ i" by (blast dest: le_imp_not_lt not_lt_imp_le) lemma not_le_iff_lt: "[| Ord(i); Ord(j) |] ==> ~ i \ j <-> j 0 \ i" by (erule not_lt_iff_le [THEN iffD1], auto) lemma Ord_0_lt: "[| Ord(i); i\0 |] ==> 0 i\0 <-> 0Results about Less-Than or Equals\ (** For ordinals, @{term"j\i"} implies @{term"j \ i"} (less-than or equals) **) lemma zero_le_succ_iff [iff]: "0 \ succ(x) <-> Ord(x)" by (blast intro: Ord_0_le elim: ltE) lemma subset_imp_le: "[| j<=i; Ord(i); Ord(j) |] ==> j \ i" apply (rule not_lt_iff_le [THEN iffD1], assumption+) apply (blast elim: ltE mem_irrefl) done lemma le_imp_subset: "i \ j ==> i<=j" by (blast dest: OrdmemD elim: ltE leE) lemma le_subset_iff: "j \ i <-> j<=i & Ord(i) & Ord(j)" by (blast dest: subset_imp_le le_imp_subset elim: ltE) lemma le_succ_iff: "i \ succ(j) <-> i \ j | i=succ(j) & Ord(i)" apply (simp (no_asm) add: le_iff) apply blast done (*Just a variant of subset_imp_le*) lemma all_lt_imp_le: "[| Ord(i); Ord(j); !!x. x x j \ i" by (blast intro: not_lt_imp_le dest: lt_irrefl) subsubsection\Transitivity Laws\ lemma lt_trans1: "[| i \ j; j i k |] ==> i j; j \ k |] ==> i \ k" by (blast intro: lt_trans1) lemma succ_leI: "i succ(i) \ j" apply (rule not_lt_iff_le [THEN iffD1]) apply (blast elim: ltE leE lt_asym)+ done (*Identical to succ(i) < succ(j) ==> i j ==> i j <-> i succ(j) ==> i \ j" by (blast dest!: succ_leE) lemma lt_subset_trans: "[| i \ j; j i 0 i j" apply auto apply (blast intro: lt_trans le_refl dest: lt_Ord) apply (frule lt_Ord) apply (rule not_le_iff_lt [THEN iffD1]) apply (blast intro: lt_Ord2) apply blast apply (simp add: lt_Ord lt_Ord2 le_iff) apply (blast dest: lt_asym) done lemma Ord_succ_mem_iff: "Ord(j) ==> succ(i) \ succ(j) <-> i\j" apply (insert succ_le_iff [of i j]) apply (simp add: lt_def) done subsubsection\Union and Intersection\ lemma Un_upper1_le: "[| Ord(i); Ord(j) |] ==> i \ i \ j" by (rule Un_upper1 [THEN subset_imp_le], auto) lemma Un_upper2_le: "[| Ord(i); Ord(j) |] ==> j \ i \ j" by (rule Un_upper2 [THEN subset_imp_le], auto) (*Replacing k by succ(k') yields the similar rule for le!*) lemma Un_least_lt: "[| i i \ j < k" apply (rule_tac i = i and j = j in Ord_linear_le) apply (auto simp add: Un_commute le_subset_iff subset_Un_iff lt_Ord) done lemma Un_least_lt_iff: "[| Ord(i); Ord(j) |] ==> i \ j < k <-> i i \ j \ k <-> i\k & j\k" apply (insert Un_least_lt_iff [of i j k]) apply (simp add: lt_def) done (*Replacing k by succ(k') yields the similar rule for le!*) lemma Int_greatest_lt: "[| i i \ j < k" apply (rule_tac i = i and j = j in Ord_linear_le) apply (auto simp add: Int_commute le_subset_iff subset_Int_iff lt_Ord) done lemma Ord_Un_if: "[| Ord(i); Ord(j) |] ==> i \ j = (if j succ(i \ j) = succ(i) \ succ(j)" by (simp add: Ord_Un_if lt_Ord le_Ord2) lemma lt_Un_iff: "[| Ord(i); Ord(j) |] ==> k < i \ j <-> k < i | k < j" apply (simp add: Ord_Un_if not_lt_iff_le) apply (blast intro: leI lt_trans2)+ done lemma le_Un_iff: "[| Ord(i); Ord(j) |] ==> k \ i \ j <-> k \ i | k \ j" by (simp add: succ_Un_distrib lt_Un_iff [symmetric]) lemma Un_upper1_lt: "[|k < i; Ord(j)|] ==> k < i \ j" by (simp add: lt_Un_iff lt_Ord2) lemma Un_upper2_lt: "[|k < j; Ord(i)|] ==> k < i \ j" by (simp add: lt_Un_iff lt_Ord2) (*See also Transset_iff_Union_succ*) lemma Ord_Union_succ_eq: "Ord(i) ==> \(succ(i)) = i" by (blast intro: Ord_trans) subsection\Results about Limits\ lemma Ord_Union [intro,simp,TC]: "[| !!i. i\A ==> Ord(i) |] ==> Ord(\(A))" apply (rule Ord_is_Transset [THEN Transset_Union_family, THEN OrdI]) apply (blast intro: Ord_contains_Transset)+ done lemma Ord_UN [intro,simp,TC]: "[| !!x. x\A ==> Ord(B(x)) |] ==> Ord(\x\A. B(x))" by (rule Ord_Union, blast) lemma Ord_Inter [intro,simp,TC]: "[| !!i. i\A ==> Ord(i) |] ==> Ord(\(A))" apply (rule Transset_Inter_family [THEN OrdI]) apply (blast intro: Ord_is_Transset) apply (simp add: Inter_def) apply (blast intro: Ord_contains_Transset) done lemma Ord_INT [intro,simp,TC]: "[| !!x. x\A ==> Ord(B(x)) |] ==> Ord(\x\A. B(x))" by (rule Ord_Inter, blast) (* No < version of this theorem: consider that @{term"(\i\nat.i)=nat"}! *) lemma UN_least_le: "[| Ord(i); !!x. x\A ==> b(x) \ i |] ==> (\x\A. b(x)) \ i" apply (rule le_imp_subset [THEN UN_least, THEN subset_imp_le]) apply (blast intro: Ord_UN elim: ltE)+ done lemma UN_succ_least_lt: "[| jA ==> b(x) (\x\A. succ(b(x))) < i" apply (rule ltE, assumption) apply (rule UN_least_le [THEN lt_trans2]) apply (blast intro: succ_leI)+ done lemma UN_upper_lt: "[| a\A; i < b(a); Ord(\x\A. b(x)) |] ==> i < (\x\A. b(x))" by (unfold lt_def, blast) lemma UN_upper_le: "[| a \ A; i \ b(a); Ord(\x\A. b(x)) |] ==> i \ (\x\A. b(x))" apply (frule ltD) apply (rule le_imp_subset [THEN subset_trans, THEN subset_imp_le]) apply (blast intro: lt_Ord UN_upper)+ done lemma lt_Union_iff: "\i\A. Ord(i) ==> (j < \(A)) <-> (\i\A. j J; i\j; Ord(\(J)) |] ==> i \ \J" apply (subst Union_eq_UN) apply (rule UN_upper_le, auto) done lemma le_implies_UN_le_UN: "[| !!x. x\A ==> c(x) \ d(x) |] ==> (\x\A. c(x)) \ (\x\A. d(x))" apply (rule UN_least_le) apply (rule_tac [2] UN_upper_le) apply (blast intro: Ord_UN le_Ord2)+ done lemma Ord_equality: "Ord(i) ==> (\y\i. succ(y)) = i" by (blast intro: Ord_trans) (*Holds for all transitive sets, not just ordinals*) lemma Ord_Union_subset: "Ord(i) ==> \(i) \ i" by (blast intro: Ord_trans) subsection\Limit Ordinals -- General Properties\ lemma Limit_Union_eq: "Limit(i) ==> \(i) = i" apply (unfold Limit_def) apply (fast intro!: ltI elim!: ltE elim: Ord_trans) done lemma Limit_is_Ord: "Limit(i) ==> Ord(i)" apply (unfold Limit_def) apply (erule conjunct1) done lemma Limit_has_0: "Limit(i) ==> 0 < i" apply (unfold Limit_def) apply (erule conjunct2 [THEN conjunct1]) done lemma Limit_nonzero: "Limit(i) ==> i \ 0" by (drule Limit_has_0, blast) lemma Limit_has_succ: "[| Limit(i); j succ(j) < i" by (unfold Limit_def, blast) lemma Limit_succ_lt_iff [simp]: "Limit(i) ==> succ(j) < i <-> (j 1 < i" by (blast intro: Limit_has_0 Limit_has_succ) lemma increasing_LimitI: "[| 0x\l. \y\l. x Limit(l)" apply (unfold Limit_def, simp add: lt_Ord2, clarify) apply (drule_tac i=y in ltD) apply (blast intro: lt_trans1 [OF _ ltI] lt_Ord2) done lemma non_succ_LimitI: assumes i: "0y. succ(y) \ i" shows "Limit(i)" proof - have Oi: "Ord(i)" using i by (simp add: lt_def) { fix y assume yi: "y y" using yi by (blast dest: le_imp_not_lt) hence "succ(y) < i" using nsucc [of y] by (blast intro: Ord_linear_lt [OF Osy Oi]) } thus ?thesis using i Oi by (auto simp add: Limit_def) qed lemma succ_LimitE [elim!]: "Limit(succ(i)) ==> P" apply (rule lt_irrefl) apply (rule Limit_has_succ, assumption) apply (erule Limit_is_Ord [THEN Ord_succD, THEN le_refl]) done lemma not_succ_Limit [simp]: "~ Limit(succ(i))" by blast lemma Limit_le_succD: "[| Limit(i); i \ succ(j) |] ==> i \ j" by (blast elim!: leE) subsubsection\Traditional 3-Way Case Analysis on Ordinals\ lemma Ord_cases_disj: "Ord(i) ==> i=0 | (\j. Ord(j) & i=succ(j)) | Limit(i)" by (blast intro!: non_succ_LimitI Ord_0_lt) lemma Ord_cases: assumes i: "Ord(i)" obtains ("0") "i=0" | (succ) j where "Ord(j)" "i=succ(j)" | (limit) "Limit(i)" by (insert Ord_cases_disj [OF i], auto) lemma trans_induct3_raw: "[| Ord(i); P(0); !!x. [| Ord(x); P(x) |] ==> P(succ(x)); !!x. [| Limit(x); \y\x. P(y) |] ==> P(x) |] ==> P(i)" apply (erule trans_induct) apply (erule Ord_cases, blast+) done -lemmas trans_induct3 = trans_induct3_raw [rule_format, case_names 0 succ limit, consumes 1] +lemma trans_induct3 [case_names 0 succ limit, consumes 1]: + "Ord(i) \ P(0) \ (\x. Ord(x) \ P(x) \ P(succ(x))) \ (\x. Limit(x) \ (\y. y \ x \ P(y)) \ P(x)) \ P(i)" + using trans_induct3_raw [of i P] by simp text\A set of ordinals is either empty, contains its own union, or its union is a limit ordinal.\ lemma Union_le: "[| !!x. x\I ==> x\j; Ord(j) |] ==> \(I) \ j" by (auto simp add: le_subset_iff Union_least) lemma Ord_set_cases: assumes I: "\i\I. Ord(i)" shows "I=0 \ \(I) \ I \ (\(I) \ I \ Limit(\(I)))" proof (cases "\(I)" rule: Ord_cases) show "Ord(\I)" using I by (blast intro: Ord_Union) next assume "\I = 0" thus ?thesis by (simp, blast intro: subst_elem) next fix j assume j: "Ord(j)" and UIj:"\(I) = succ(j)" { assume "\i\I. i\j" hence "\(I) \ j" by (simp add: Union_le j) hence False by (simp add: UIj lt_not_refl) } then obtain i where i: "i \ I" "succ(j) \ i" using I j by (atomize, auto simp add: not_le_iff_lt) have "\(I) \ succ(j)" using UIj j by auto hence "i \ succ(j)" using i by (simp add: le_subset_iff Union_subset_iff) hence "succ(j) = i" using i by (blast intro: le_anti_sym) hence "succ(j) \ I" by (simp add: i) thus ?thesis by (simp add: UIj) next assume "Limit(\I)" thus ?thesis by auto qed text\If the union of a set of ordinals is a successor, then it is an element of that set.\ lemma Ord_Union_eq_succD: "[|\x\X. Ord(x); \X = succ(j)|] ==> succ(j) \ X" by (drule Ord_set_cases, auto) -lemma Limit_Union [rule_format]: "[| I \ 0; \i\I. Limit(i) |] ==> Limit(\I)" +lemma Limit_Union [rule_format]: "[| I \ 0; (\i. i\I \ Limit(i)) |] ==> Limit(\I)" apply (simp add: Limit_def lt_def) apply (blast intro!: equalityI) done end diff --git a/src/ZF/WF.thy b/src/ZF/WF.thy --- a/src/ZF/WF.thy +++ b/src/ZF/WF.thy @@ -1,371 +1,372 @@ (* Title: ZF/WF.thy Author: Tobias Nipkow and Lawrence C Paulson Copyright 1994 University of Cambridge Derived first for transitive relations, and finally for arbitrary WF relations via wf_trancl and trans_trancl. It is difficult to derive this general case directly, using r^+ instead of r. In is_recfun, the two occurrences of the relation must have the same form. Inserting r^+ in the_recfun or wftrec yields a recursion rule with r^+ -`` {a} instead of r-``{a}. This recursion rule is stronger in principle, but harder to use, especially to prove wfrec_eclose_eq in epsilon.ML. Expanding out the definition of wftrec in wfrec would yield a mess. *) section\Well-Founded Recursion\ theory WF imports Trancl begin definition wf :: "i=>o" where (*r is a well-founded relation*) "wf(r) == \Z. Z=0 | (\x\Z. \y. :r \ ~ y \ Z)" definition wf_on :: "[i,i]=>o" (\wf[_]'(_')\) where (*r is well-founded on A*) "wf_on(A,r) == wf(r \ A*A)" definition is_recfun :: "[i, i, [i,i]=>i, i] =>o" where "is_recfun(r,a,H,f) == (f = (\x\r-``{a}. H(x, restrict(f, r-``{x}))))" definition the_recfun :: "[i, i, [i,i]=>i] =>i" where "the_recfun(r,a,H) == (THE f. is_recfun(r,a,H,f))" definition wftrec :: "[i, i, [i,i]=>i] =>i" where "wftrec(r,a,H) == H(a, the_recfun(r,a,H))" definition wfrec :: "[i, i, [i,i]=>i] =>i" where (*public version. Does not require r to be transitive*) "wfrec(r,a,H) == wftrec(r^+, a, %x f. H(x, restrict(f,r-``{x})))" definition wfrec_on :: "[i, i, i, [i,i]=>i] =>i" (\wfrec[_]'(_,_,_')\) where "wfrec[A](r,a,H) == wfrec(r \ A*A, a, H)" subsection\Well-Founded Relations\ subsubsection\Equivalences between \<^term>\wf\ and \<^term>\wf_on\\ lemma wf_imp_wf_on: "wf(r) ==> wf[A](r)" by (unfold wf_def wf_on_def, force) lemma wf_on_imp_wf: "[|wf[A](r); r \ A*A|] ==> wf(r)" by (simp add: wf_on_def subset_Int_iff) lemma wf_on_field_imp_wf: "wf[field(r)](r) ==> wf(r)" by (unfold wf_def wf_on_def, fast) lemma wf_iff_wf_on_field: "wf(r) \ wf[field(r)](r)" by (blast intro: wf_imp_wf_on wf_on_field_imp_wf) lemma wf_on_subset_A: "[| wf[A](r); B<=A |] ==> wf[B](r)" by (unfold wf_on_def wf_def, fast) lemma wf_on_subset_r: "[| wf[A](r); s<=r |] ==> wf[A](s)" by (unfold wf_on_def wf_def, fast) lemma wf_subset: "[|wf(s); r<=s|] ==> wf(r)" by (simp add: wf_def, fast) subsubsection\Introduction Rules for \<^term>\wf_on\\ text\If every non-empty subset of \<^term>\A\ has an \<^term>\r\-minimal element then we have \<^term>\wf[A](r)\.\ lemma wf_onI: assumes prem: "!!Z u. [| Z<=A; u \ Z; \x\Z. \y\Z. :r |] ==> False" shows "wf[A](r)" apply (unfold wf_on_def wf_def) apply (rule equals0I [THEN disjCI, THEN allI]) apply (rule_tac Z = Z in prem, blast+) done text\If \<^term>\r\ allows well-founded induction over \<^term>\A\ then we have \<^term>\wf[A](r)\. Premise is equivalent to \<^prop>\!!B. \x\A. (\y. : r \ y \ B) \ x \ B ==> A<=B\\ lemma wf_onI2: assumes prem: "!!y B. [| \x\A. (\y\A. :r \ y \ B) \ x \ B; y \ A |] ==> y \ B" shows "wf[A](r)" apply (rule wf_onI) apply (rule_tac c=u in prem [THEN DiffE]) prefer 3 apply blast apply fast+ done subsubsection\Well-founded Induction\ text\Consider the least \<^term>\z\ in \<^term>\domain(r)\ such that \<^term>\P(z)\ does not hold...\ lemma wf_induct_raw: "[| wf(r); !!x.[| \y. : r \ P(y) |] ==> P(x) |] ==> P(a)" apply (unfold wf_def) apply (erule_tac x = "{z \ domain(r). ~ P(z)}" in allE) apply blast done lemmas wf_induct = wf_induct_raw [rule_format, consumes 1, case_names step, induct set: wf] text\The form of this rule is designed to match \wfI\\ lemma wf_induct2: "[| wf(r); a \ A; field(r)<=A; !!x.[| x \ A; \y. : r \ P(y) |] ==> P(x) |] ==> P(a)" apply (erule_tac P="a \ A" in rev_mp) apply (erule_tac a=a in wf_induct, blast) done lemma field_Int_square: "field(r \ A*A) \ A" by blast lemma wf_on_induct_raw [consumes 2, induct set: wf_on]: "[| wf[A](r); a \ A; !!x.[| x \ A; \y\A. : r \ P(y) |] ==> P(x) |] ==> P(a)" apply (unfold wf_on_def) apply (erule wf_induct2, assumption) apply (rule field_Int_square, blast) done -lemmas wf_on_induct = - wf_on_induct_raw [rule_format, consumes 2, case_names step, induct set: wf_on] - +lemma wf_on_induct [consumes 2, case_names step, induct set: wf_on]: + "wf[A](r) \ a \ A \ (\x. x \ A \ (\y. y \ A \ \y, x\ \ r \ P(y)) \ P(x)) \ P(a)" + using wf_on_induct_raw [of A r a P] by simp text\If \<^term>\r\ allows well-founded induction then we have \<^term>\wf(r)\.\ lemma wfI: "[| field(r)<=A; !!y B. [| \x\A. (\y\A. :r \ y \ B) \ x \ B; y \ A|] ==> y \ B |] ==> wf(r)" apply (rule wf_on_subset_A [THEN wf_on_field_imp_wf]) apply (rule wf_onI2) prefer 2 apply blast apply blast done subsection\Basic Properties of Well-Founded Relations\ lemma wf_not_refl: "wf(r) ==> \ r" by (erule_tac a=a in wf_induct, blast) lemma wf_not_sym [rule_format]: "wf(r) ==> \x. :r \ \ r" by (erule_tac a=a in wf_induct, blast) (* @{term"[| wf(r); \ r; ~P ==> \ r |] ==> P"} *) lemmas wf_asym = wf_not_sym [THEN swap] lemma wf_on_not_refl: "[| wf[A](r); a \ A |] ==> \ r" by (erule_tac a=a in wf_on_induct, assumption, blast) -lemma wf_on_not_sym [rule_format]: - "[| wf[A](r); a \ A |] ==> \b\A. :r \ \r" +lemma wf_on_not_sym: + "[| wf[A](r); a \ A |] ==> (\b. b\A \ :r \ \r)" +apply (atomize (full), intro impI) apply (erule_tac a=a in wf_on_induct, assumption, blast) done lemma wf_on_asym: "[| wf[A](r); ~Z ==> \ r; \ r ==> Z; ~Z ==> a \ A; ~Z ==> b \ A |] ==> Z" by (blast dest: wf_on_not_sym) (*Needed to prove well_ordI. Could also reason that wf[A](r) means wf(r \ A*A); thus wf( (r \ A*A)^+ ) and use wf_not_refl *) lemma wf_on_chain3: "[| wf[A](r); :r; :r; :r; a \ A; b \ A; c \ A |] ==> P" apply (subgoal_tac "\y\A. \z\A. :r \ :r \ :r \ P", blast) apply (erule_tac a=a in wf_on_induct, assumption, blast) done text\transitive closure of a WF relation is WF provided \<^term>\A\ is downward closed\ lemma wf_on_trancl: "[| wf[A](r); r-``A \ A |] ==> wf[A](r^+)" apply (rule wf_onI2) apply (frule bspec [THEN mp], assumption+) apply (erule_tac a = y in wf_on_induct, assumption) apply (blast elim: tranclE, blast) done lemma wf_trancl: "wf(r) ==> wf(r^+)" apply (simp add: wf_iff_wf_on_field) apply (rule wf_on_subset_A) apply (erule wf_on_trancl) apply blast apply (rule trancl_type [THEN field_rel_subset]) done text\\<^term>\r-``{a}\ is the set of everything under \<^term>\a\ in \<^term>\r\\ lemmas underI = vimage_singleton_iff [THEN iffD2] lemmas underD = vimage_singleton_iff [THEN iffD1] subsection\The Predicate \<^term>\is_recfun\\ lemma is_recfun_type: "is_recfun(r,a,H,f) ==> f \ r-``{a} -> range(f)" apply (unfold is_recfun_def) apply (erule ssubst) apply (rule lamI [THEN rangeI, THEN lam_type], assumption) done lemmas is_recfun_imp_function = is_recfun_type [THEN fun_is_function] lemma apply_recfun: "[| is_recfun(r,a,H,f); :r |] ==> f`x = H(x, restrict(f,r-``{x}))" apply (unfold is_recfun_def) txt\replace f only on the left-hand side\ apply (erule_tac P = "%x. t(x) = u" for t u in ssubst) apply (simp add: underI) done lemma is_recfun_equal [rule_format]: "[| wf(r); trans(r); is_recfun(r,a,H,f); is_recfun(r,b,H,g) |] ==> :r \ :r \ f`x=g`x" apply (frule_tac f = f in is_recfun_type) apply (frule_tac f = g in is_recfun_type) apply (simp add: is_recfun_def) apply (erule_tac a=x in wf_induct) apply (intro impI) apply (elim ssubst) apply (simp (no_asm_simp) add: vimage_singleton_iff restrict_def) apply (rule_tac t = "%z. H (x, z)" for x in subst_context) apply (subgoal_tac "\y\r-``{x}. \z. :f \ :g") apply (blast dest: transD) apply (simp add: apply_iff) apply (blast dest: transD intro: sym) done lemma is_recfun_cut: "[| wf(r); trans(r); is_recfun(r,a,H,f); is_recfun(r,b,H,g); :r |] ==> restrict(f, r-``{b}) = g" apply (frule_tac f = f in is_recfun_type) apply (rule fun_extension) apply (blast dest: transD intro: restrict_type2) apply (erule is_recfun_type, simp) apply (blast dest: transD intro: is_recfun_equal) done subsection\Recursion: Main Existence Lemma\ lemma is_recfun_functional: "[| wf(r); trans(r); is_recfun(r,a,H,f); is_recfun(r,a,H,g) |] ==> f=g" by (blast intro: fun_extension is_recfun_type is_recfun_equal) lemma the_recfun_eq: "[| is_recfun(r,a,H,f); wf(r); trans(r) |] ==> the_recfun(r,a,H) = f" apply (unfold the_recfun_def) apply (blast intro: is_recfun_functional) done (*If some f satisfies is_recfun(r,a,H,-) then so does the_recfun(r,a,H) *) lemma is_the_recfun: "[| is_recfun(r,a,H,f); wf(r); trans(r) |] ==> is_recfun(r, a, H, the_recfun(r,a,H))" by (simp add: the_recfun_eq) lemma unfold_the_recfun: "[| wf(r); trans(r) |] ==> is_recfun(r, a, H, the_recfun(r,a,H))" apply (rule_tac a=a in wf_induct, assumption) apply (rename_tac a1) apply (rule_tac f = "\y\r-``{a1}. wftrec (r,y,H)" in is_the_recfun) apply typecheck apply (unfold is_recfun_def wftrec_def) \ \Applying the substitution: must keep the quantified assumption!\ apply (rule lam_cong [OF refl]) apply (drule underD) apply (fold is_recfun_def) apply (rule_tac t = "%z. H(x, z)" for x in subst_context) apply (rule fun_extension) apply (blast intro: is_recfun_type) apply (rule lam_type [THEN restrict_type2]) apply blast apply (blast dest: transD) apply atomize apply (frule spec [THEN mp], assumption) apply (subgoal_tac " \ r") apply (drule_tac x1 = xa in spec [THEN mp], assumption) apply (simp add: vimage_singleton_iff apply_recfun is_recfun_cut) apply (blast dest: transD) done subsection\Unfolding \<^term>\wftrec(r,a,H)\\ lemma the_recfun_cut: "[| wf(r); trans(r); :r |] ==> restrict(the_recfun(r,a,H), r-``{b}) = the_recfun(r,b,H)" by (blast intro: is_recfun_cut unfold_the_recfun) (*NOT SUITABLE FOR REWRITING: it is recursive!*) lemma wftrec: "[| wf(r); trans(r) |] ==> wftrec(r,a,H) = H(a, \x\r-``{a}. wftrec(r,x,H))" apply (unfold wftrec_def) apply (subst unfold_the_recfun [unfolded is_recfun_def]) apply (simp_all add: vimage_singleton_iff [THEN iff_sym] the_recfun_cut) done subsubsection\Removal of the Premise \<^term>\trans(r)\\ (*NOT SUITABLE FOR REWRITING: it is recursive!*) lemma wfrec: "wf(r) ==> wfrec(r,a,H) = H(a, \x\r-``{a}. wfrec(r,x,H))" apply (unfold wfrec_def) apply (erule wf_trancl [THEN wftrec, THEN ssubst]) apply (rule trans_trancl) apply (rule vimage_pair_mono [THEN restrict_lam_eq, THEN subst_context]) apply (erule r_into_trancl) apply (rule subset_refl) done (*This form avoids giant explosions in proofs. NOTE USE OF == *) lemma def_wfrec: "[| !!x. h(x)==wfrec(r,x,H); wf(r) |] ==> h(a) = H(a, \x\r-``{a}. h(x))" apply simp apply (elim wfrec) done lemma wfrec_type: "[| wf(r); a \ A; field(r)<=A; !!x u. [| x \ A; u \ Pi(r-``{x}, B) |] ==> H(x,u) \ B(x) |] ==> wfrec(r,a,H) \ B(a)" apply (rule_tac a = a in wf_induct2, assumption+) apply (subst wfrec, assumption) apply (simp add: lam_type underD) done lemma wfrec_on: "[| wf[A](r); a \ A |] ==> wfrec[A](r,a,H) = H(a, \x\(r-``{a}) \ A. wfrec[A](r,x,H))" apply (unfold wf_on_def wfrec_on_def) apply (erule wfrec [THEN trans]) apply (simp add: vimage_Int_square cons_subset_iff) done text\Minimal-element characterization of well-foundedness\ lemma wf_eq_minimal: "wf(r) \ (\Q x. x \ Q \ (\z\Q. \y. :r \ y\Q))" by (unfold wf_def, blast) end diff --git a/src/ZF/equalities.thy b/src/ZF/equalities.thy --- a/src/ZF/equalities.thy +++ b/src/ZF/equalities.thy @@ -1,984 +1,984 @@ (* Title: ZF/equalities.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1992 University of Cambridge *) section\Basic Equalities and Inclusions\ theory equalities imports pair begin text\These cover union, intersection, converse, domain, range, etc. Philippe de Groote proved many of the inclusions.\ lemma in_mono: "A\B ==> x\A \ x\B" by blast lemma the_eq_0 [simp]: "(THE x. False) = 0" by (blast intro: the_0) subsection\Bounded Quantifiers\ text \\medskip The following are not added to the default simpset because (a) they duplicate the body and (b) there are no similar rules for \Int\.\ lemma ball_Un: "(\x \ A\B. P(x)) \ (\x \ A. P(x)) & (\x \ B. P(x))" by blast lemma bex_Un: "(\x \ A\B. P(x)) \ (\x \ A. P(x)) | (\x \ B. P(x))" by blast lemma ball_UN: "(\z \ (\x\A. B(x)). P(z)) \ (\x\A. \z \ B(x). P(z))" by blast lemma bex_UN: "(\z \ (\x\A. B(x)). P(z)) \ (\x\A. \z\B(x). P(z))" by blast subsection\Converse of a Relation\ lemma converse_iff [simp]: "\ converse(r) \ \r" by (unfold converse_def, blast) lemma converseI [intro!]: "\r ==> \converse(r)" by (unfold converse_def, blast) lemma converseD: " \ converse(r) ==> \ r" by (unfold converse_def, blast) lemma converseE [elim!]: "[| yx \ converse(r); !!x y. [| yx=; \r |] ==> P |] ==> P" by (unfold converse_def, blast) lemma converse_converse: "r\Sigma(A,B) ==> converse(converse(r)) = r" by blast lemma converse_type: "r\A*B ==> converse(r)\B*A" by blast lemma converse_prod [simp]: "converse(A*B) = B*A" by blast lemma converse_empty [simp]: "converse(0) = 0" by blast lemma converse_subset_iff: "A \ Sigma(X,Y) ==> converse(A) \ converse(B) \ A \ B" by blast subsection\Finite Set Constructions Using \<^term>\cons\\ lemma cons_subsetI: "[| a\C; B\C |] ==> cons(a,B) \ C" by blast lemma subset_consI: "B \ cons(a,B)" by blast lemma cons_subset_iff [iff]: "cons(a,B)\C \ a\C & B\C" by blast (*A safe special case of subset elimination, adding no new variables [| cons(a,B) \ C; [| a \ C; B \ C |] ==> R |] ==> R *) lemmas cons_subsetE = cons_subset_iff [THEN iffD1, THEN conjE] lemma subset_empty_iff: "A\0 \ A=0" by blast lemma subset_cons_iff: "C\cons(a,B) \ C\B | (a\C & C-{a} \ B)" by blast (* cons_def refers to Upair; reversing the equality LOOPS in rewriting!*) lemma cons_eq: "{a} \ B = cons(a,B)" by blast lemma cons_commute: "cons(a, cons(b, C)) = cons(b, cons(a, C))" by blast lemma cons_absorb: "a: B ==> cons(a,B) = B" by blast lemma cons_Diff: "a: B ==> cons(a, B-{a}) = B" by blast lemma Diff_cons_eq: "cons(a,B) - C = (if a\C then B-C else cons(a,B-C))" by auto -lemma equal_singleton [rule_format]: "[| a: C; \y\C. y=b |] ==> C = {b}" +lemma equal_singleton: "[| a: C; \y. y \C \ y=b |] ==> C = {b}" by blast lemma [simp]: "cons(a,cons(a,B)) = cons(a,B)" by blast (** singletons **) lemma singleton_subsetI: "a\C ==> {a} \ C" by blast lemma singleton_subsetD: "{a} \ C ==> a\C" by blast (** succ **) lemma subset_succI: "i \ succ(i)" by blast (*But if j is an ordinal or is transitive, then @{term"i\j"} implies @{term"i\j"}! See @{text"Ord_succ_subsetI}*) lemma succ_subsetI: "[| i\j; i\j |] ==> succ(i)\j" by (unfold succ_def, blast) lemma succ_subsetE: "[| succ(i) \ j; [| i\j; i\j |] ==> P |] ==> P" by (unfold succ_def, blast) lemma succ_subset_iff: "succ(a) \ B \ (a \ B & a \ B)" by (unfold succ_def, blast) subsection\Binary Intersection\ (** Intersection is the greatest lower bound of two sets **) lemma Int_subset_iff: "C \ A \ B \ C \ A & C \ B" by blast lemma Int_lower1: "A \ B \ A" by blast lemma Int_lower2: "A \ B \ B" by blast lemma Int_greatest: "[| C\A; C\B |] ==> C \ A \ B" by blast lemma Int_cons: "cons(a,B) \ C \ cons(a, B \ C)" by blast lemma Int_absorb [simp]: "A \ A = A" by blast lemma Int_left_absorb: "A \ (A \ B) = A \ B" by blast lemma Int_commute: "A \ B = B \ A" by blast lemma Int_left_commute: "A \ (B \ C) = B \ (A \ C)" by blast lemma Int_assoc: "(A \ B) \ C = A \ (B \ C)" by blast (*Intersection is an AC-operator*) lemmas Int_ac= Int_assoc Int_left_absorb Int_commute Int_left_commute lemma Int_absorb1: "B \ A ==> A \ B = B" by blast lemma Int_absorb2: "A \ B ==> A \ B = A" by blast lemma Int_Un_distrib: "A \ (B \ C) = (A \ B) \ (A \ C)" by blast lemma Int_Un_distrib2: "(B \ C) \ A = (B \ A) \ (C \ A)" by blast lemma subset_Int_iff: "A\B \ A \ B = A" by (blast elim!: equalityE) lemma subset_Int_iff2: "A\B \ B \ A = A" by (blast elim!: equalityE) lemma Int_Diff_eq: "C\A ==> (A-B) \ C = C-B" by blast lemma Int_cons_left: "cons(a,A) \ B = (if a \ B then cons(a, A \ B) else A \ B)" by auto lemma Int_cons_right: "A \ cons(a, B) = (if a \ A then cons(a, A \ B) else A \ B)" by auto lemma cons_Int_distrib: "cons(x, A \ B) = cons(x, A) \ cons(x, B)" by auto subsection\Binary Union\ (** Union is the least upper bound of two sets *) lemma Un_subset_iff: "A \ B \ C \ A \ C & B \ C" by blast lemma Un_upper1: "A \ A \ B" by blast lemma Un_upper2: "B \ A \ B" by blast lemma Un_least: "[| A\C; B\C |] ==> A \ B \ C" by blast lemma Un_cons: "cons(a,B) \ C = cons(a, B \ C)" by blast lemma Un_absorb [simp]: "A \ A = A" by blast lemma Un_left_absorb: "A \ (A \ B) = A \ B" by blast lemma Un_commute: "A \ B = B \ A" by blast lemma Un_left_commute: "A \ (B \ C) = B \ (A \ C)" by blast lemma Un_assoc: "(A \ B) \ C = A \ (B \ C)" by blast (*Union is an AC-operator*) lemmas Un_ac = Un_assoc Un_left_absorb Un_commute Un_left_commute lemma Un_absorb1: "A \ B ==> A \ B = B" by blast lemma Un_absorb2: "B \ A ==> A \ B = A" by blast lemma Un_Int_distrib: "(A \ B) \ C = (A \ C) \ (B \ C)" by blast lemma subset_Un_iff: "A\B \ A \ B = B" by (blast elim!: equalityE) lemma subset_Un_iff2: "A\B \ B \ A = B" by (blast elim!: equalityE) lemma Un_empty [iff]: "(A \ B = 0) \ (A = 0 & B = 0)" by blast lemma Un_eq_Union: "A \ B = \({A, B})" by blast subsection\Set Difference\ lemma Diff_subset: "A-B \ A" by blast lemma Diff_contains: "[| C\A; C \ B = 0 |] ==> C \ A-B" by blast lemma subset_Diff_cons_iff: "B \ A - cons(c,C) \ B\A-C & c \ B" by blast lemma Diff_cancel: "A - A = 0" by blast lemma Diff_triv: "A \ B = 0 ==> A - B = A" by blast lemma empty_Diff [simp]: "0 - A = 0" by blast lemma Diff_0 [simp]: "A - 0 = A" by blast lemma Diff_eq_0_iff: "A - B = 0 \ A \ B" by (blast elim: equalityE) (*NOT SUITABLE FOR REWRITING since {a} == cons(a,0)*) lemma Diff_cons: "A - cons(a,B) = A - B - {a}" by blast (*NOT SUITABLE FOR REWRITING since {a} == cons(a,0)*) lemma Diff_cons2: "A - cons(a,B) = A - {a} - B" by blast lemma Diff_disjoint: "A \ (B-A) = 0" by blast lemma Diff_partition: "A\B ==> A \ (B-A) = B" by blast lemma subset_Un_Diff: "A \ B \ (A - B)" by blast lemma double_complement: "[| A\B; B\C |] ==> B-(C-A) = A" by blast lemma double_complement_Un: "(A \ B) - (B-A) = A" by blast lemma Un_Int_crazy: "(A \ B) \ (B \ C) \ (C \ A) = (A \ B) \ (B \ C) \ (C \ A)" apply blast done lemma Diff_Un: "A - (B \ C) = (A-B) \ (A-C)" by blast lemma Diff_Int: "A - (B \ C) = (A-B) \ (A-C)" by blast lemma Un_Diff: "(A \ B) - C = (A - C) \ (B - C)" by blast lemma Int_Diff: "(A \ B) - C = A \ (B - C)" by blast lemma Diff_Int_distrib: "C \ (A-B) = (C \ A) - (C \ B)" by blast lemma Diff_Int_distrib2: "(A-B) \ C = (A \ C) - (B \ C)" by blast (*Halmos, Naive Set Theory, page 16.*) lemma Un_Int_assoc_iff: "(A \ B) \ C = A \ (B \ C) \ C\A" by (blast elim!: equalityE) subsection\Big Union and Intersection\ (** Big Union is the least upper bound of a set **) lemma Union_subset_iff: "\(A) \ C \ (\x\A. x \ C)" by blast lemma Union_upper: "B\A ==> B \ \(A)" by blast lemma Union_least: "[| !!x. x\A ==> x\C |] ==> \(A) \ C" by blast lemma Union_cons [simp]: "\(cons(a,B)) = a \ \(B)" by blast lemma Union_Un_distrib: "\(A \ B) = \(A) \ \(B)" by blast lemma Union_Int_subset: "\(A \ B) \ \(A) \ \(B)" by blast lemma Union_disjoint: "\(C) \ A = 0 \ (\B\C. B \ A = 0)" by (blast elim!: equalityE) lemma Union_empty_iff: "\(A) = 0 \ (\B\A. B=0)" by blast lemma Int_Union2: "\(B) \ A = (\C\B. C \ A)" by blast (** Big Intersection is the greatest lower bound of a nonempty set **) lemma Inter_subset_iff: "A\0 ==> C \ \(A) \ (\x\A. C \ x)" by blast lemma Inter_lower: "B\A ==> \(A) \ B" by blast lemma Inter_greatest: "[| A\0; !!x. x\A ==> C\x |] ==> C \ \(A)" by blast (** Intersection of a family of sets **) lemma INT_lower: "x\A ==> (\x\A. B(x)) \ B(x)" by blast lemma INT_greatest: "[| A\0; !!x. x\A ==> C\B(x) |] ==> C \ (\x\A. B(x))" by force lemma Inter_0 [simp]: "\(0) = 0" by (unfold Inter_def, blast) lemma Inter_Un_subset: "[| z\A; z\B |] ==> \(A) \ \(B) \ \(A \ B)" by blast (* A good challenge: Inter is ill-behaved on the empty set *) lemma Inter_Un_distrib: "[| A\0; B\0 |] ==> \(A \ B) = \(A) \ \(B)" by blast lemma Union_singleton: "\({b}) = b" by blast lemma Inter_singleton: "\({b}) = b" by blast lemma Inter_cons [simp]: "\(cons(a,B)) = (if B=0 then a else a \ \(B))" by force subsection\Unions and Intersections of Families\ lemma subset_UN_iff_eq: "A \ (\i\I. B(i)) \ A = (\i\I. A \ B(i))" by (blast elim!: equalityE) lemma UN_subset_iff: "(\x\A. B(x)) \ C \ (\x\A. B(x) \ C)" by blast lemma UN_upper: "x\A ==> B(x) \ (\x\A. B(x))" by (erule RepFunI [THEN Union_upper]) lemma UN_least: "[| !!x. x\A ==> B(x)\C |] ==> (\x\A. B(x)) \ C" by blast lemma Union_eq_UN: "\(A) = (\x\A. x)" by blast lemma Inter_eq_INT: "\(A) = (\x\A. x)" by (unfold Inter_def, blast) lemma UN_0 [simp]: "(\i\0. A(i)) = 0" by blast lemma UN_singleton: "(\x\A. {x}) = A" by blast lemma UN_Un: "(\i\ A \ B. C(i)) = (\i\ A. C(i)) \ (\i\B. C(i))" by blast lemma INT_Un: "(\i\I \ J. A(i)) = (if I=0 then \j\J. A(j) else if J=0 then \i\I. A(i) else ((\i\I. A(i)) \ (\j\J. A(j))))" by (simp, blast intro!: equalityI) lemma UN_UN_flatten: "(\x \ (\y\A. B(y)). C(x)) = (\y\A. \x\ B(y). C(x))" by blast (*Halmos, Naive Set Theory, page 35.*) lemma Int_UN_distrib: "B \ (\i\I. A(i)) = (\i\I. B \ A(i))" by blast lemma Un_INT_distrib: "I\0 ==> B \ (\i\I. A(i)) = (\i\I. B \ A(i))" by auto lemma Int_UN_distrib2: "(\i\I. A(i)) \ (\j\J. B(j)) = (\i\I. \j\J. A(i) \ B(j))" by blast lemma Un_INT_distrib2: "[| I\0; J\0 |] ==> (\i\I. A(i)) \ (\j\J. B(j)) = (\i\I. \j\J. A(i) \ B(j))" by auto lemma UN_constant [simp]: "(\y\A. c) = (if A=0 then 0 else c)" by force lemma INT_constant [simp]: "(\y\A. c) = (if A=0 then 0 else c)" by force lemma UN_RepFun [simp]: "(\y\ RepFun(A,f). B(y)) = (\x\A. B(f(x)))" by blast lemma INT_RepFun [simp]: "(\x\RepFun(A,f). B(x)) = (\a\A. B(f(a)))" by (auto simp add: Inter_def) lemma INT_Union_eq: "0 \ A ==> (\x\ \(A). B(x)) = (\y\A. \x\y. B(x))" apply (subgoal_tac "\x\A. x\0") prefer 2 apply blast apply (force simp add: Inter_def ball_conj_distrib) done lemma INT_UN_eq: "(\x\A. B(x) \ 0) ==> (\z\ (\x\A. B(x)). C(z)) = (\x\A. \z\ B(x). C(z))" apply (subst INT_Union_eq, blast) apply (simp add: Inter_def) done (** Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5: Union of a family of unions **) lemma UN_Un_distrib: "(\i\I. A(i) \ B(i)) = (\i\I. A(i)) \ (\i\I. B(i))" by blast lemma INT_Int_distrib: "I\0 ==> (\i\I. A(i) \ B(i)) = (\i\I. A(i)) \ (\i\I. B(i))" by (blast elim!: not_emptyE) lemma UN_Int_subset: "(\z\I \ J. A(z)) \ (\z\I. A(z)) \ (\z\J. A(z))" by blast (** Devlin, page 12, exercise 5: Complements **) lemma Diff_UN: "I\0 ==> B - (\i\I. A(i)) = (\i\I. B - A(i))" by (blast elim!: not_emptyE) lemma Diff_INT: "I\0 ==> B - (\i\I. A(i)) = (\i\I. B - A(i))" by (blast elim!: not_emptyE) (** Unions and Intersections with General Sum **) (*Not suitable for rewriting: LOOPS!*) lemma Sigma_cons1: "Sigma(cons(a,B), C) = ({a}*C(a)) \ Sigma(B,C)" by blast (*Not suitable for rewriting: LOOPS!*) lemma Sigma_cons2: "A * cons(b,B) = A*{b} \ A*B" by blast lemma Sigma_succ1: "Sigma(succ(A), B) = ({A}*B(A)) \ Sigma(A,B)" by blast lemma Sigma_succ2: "A * succ(B) = A*{B} \ A*B" by blast lemma SUM_UN_distrib1: "(\x \ (\y\A. C(y)). B(x)) = (\y\A. \x\C(y). B(x))" by blast lemma SUM_UN_distrib2: "(\i\I. \j\J. C(i,j)) = (\j\J. \i\I. C(i,j))" by blast lemma SUM_Un_distrib1: "(\i\I \ J. C(i)) = (\i\I. C(i)) \ (\j\J. C(j))" by blast lemma SUM_Un_distrib2: "(\i\I. A(i) \ B(i)) = (\i\I. A(i)) \ (\i\I. B(i))" by blast (*First-order version of the above, for rewriting*) lemma prod_Un_distrib2: "I * (A \ B) = I*A \ I*B" by (rule SUM_Un_distrib2) lemma SUM_Int_distrib1: "(\i\I \ J. C(i)) = (\i\I. C(i)) \ (\j\J. C(j))" by blast lemma SUM_Int_distrib2: "(\i\I. A(i) \ B(i)) = (\i\I. A(i)) \ (\i\I. B(i))" by blast (*First-order version of the above, for rewriting*) lemma prod_Int_distrib2: "I * (A \ B) = I*A \ I*B" by (rule SUM_Int_distrib2) (*Cf Aczel, Non-Well-Founded Sets, page 115*) lemma SUM_eq_UN: "(\i\I. A(i)) = (\i\I. {i} * A(i))" by blast lemma times_subset_iff: "(A'*B' \ A*B) \ (A' = 0 | B' = 0 | (A'\A) & (B'\B))" by blast lemma Int_Sigma_eq: "(\x \ A'. B'(x)) \ (\x \ A. B(x)) = (\x \ A' \ A. B'(x) \ B(x))" by blast (** Domain **) lemma domain_iff: "a: domain(r) \ (\y. \ r)" by (unfold domain_def, blast) lemma domainI [intro]: "\ r ==> a: domain(r)" by (unfold domain_def, blast) lemma domainE [elim!]: "[| a \ domain(r); !!y. \ r ==> P |] ==> P" by (unfold domain_def, blast) lemma domain_subset: "domain(Sigma(A,B)) \ A" by blast lemma domain_of_prod: "b\B ==> domain(A*B) = A" by blast lemma domain_0 [simp]: "domain(0) = 0" by blast lemma domain_cons [simp]: "domain(cons(,r)) = cons(a, domain(r))" by blast lemma domain_Un_eq [simp]: "domain(A \ B) = domain(A) \ domain(B)" by blast lemma domain_Int_subset: "domain(A \ B) \ domain(A) \ domain(B)" by blast lemma domain_Diff_subset: "domain(A) - domain(B) \ domain(A - B)" by blast lemma domain_UN: "domain(\x\A. B(x)) = (\x\A. domain(B(x)))" by blast lemma domain_Union: "domain(\(A)) = (\x\A. domain(x))" by blast (** Range **) lemma rangeI [intro]: "\ r ==> b \ range(r)" apply (unfold range_def) apply (erule converseI [THEN domainI]) done lemma rangeE [elim!]: "[| b \ range(r); !!x. \ r ==> P |] ==> P" by (unfold range_def, blast) lemma range_subset: "range(A*B) \ B" apply (unfold range_def) apply (subst converse_prod) apply (rule domain_subset) done lemma range_of_prod: "a\A ==> range(A*B) = B" by blast lemma range_0 [simp]: "range(0) = 0" by blast lemma range_cons [simp]: "range(cons(,r)) = cons(b, range(r))" by blast lemma range_Un_eq [simp]: "range(A \ B) = range(A) \ range(B)" by blast lemma range_Int_subset: "range(A \ B) \ range(A) \ range(B)" by blast lemma range_Diff_subset: "range(A) - range(B) \ range(A - B)" by blast lemma domain_converse [simp]: "domain(converse(r)) = range(r)" by blast lemma range_converse [simp]: "range(converse(r)) = domain(r)" by blast (** Field **) lemma fieldI1: "\ r ==> a \ field(r)" by (unfold field_def, blast) lemma fieldI2: "\ r ==> b \ field(r)" by (unfold field_def, blast) lemma fieldCI [intro]: "(~ \r ==> \ r) ==> a \ field(r)" apply (unfold field_def, blast) done lemma fieldE [elim!]: "[| a \ field(r); !!x. \ r ==> P; !!x. \ r ==> P |] ==> P" by (unfold field_def, blast) lemma field_subset: "field(A*B) \ A \ B" by blast lemma domain_subset_field: "domain(r) \ field(r)" apply (unfold field_def) apply (rule Un_upper1) done lemma range_subset_field: "range(r) \ field(r)" apply (unfold field_def) apply (rule Un_upper2) done lemma domain_times_range: "r \ Sigma(A,B) ==> r \ domain(r)*range(r)" by blast lemma field_times_field: "r \ Sigma(A,B) ==> r \ field(r)*field(r)" by blast lemma relation_field_times_field: "relation(r) ==> r \ field(r)*field(r)" by (simp add: relation_def, blast) lemma field_of_prod: "field(A*A) = A" by blast lemma field_0 [simp]: "field(0) = 0" by blast lemma field_cons [simp]: "field(cons(,r)) = cons(a, cons(b, field(r)))" by blast lemma field_Un_eq [simp]: "field(A \ B) = field(A) \ field(B)" by blast lemma field_Int_subset: "field(A \ B) \ field(A) \ field(B)" by blast lemma field_Diff_subset: "field(A) - field(B) \ field(A - B)" by blast lemma field_converse [simp]: "field(converse(r)) = field(r)" by blast (** The Union of a set of relations is a relation -- Lemma for fun_Union **) lemma rel_Union: "(\x\S. \A B. x \ A*B) ==> \(S) \ domain(\(S)) * range(\(S))" by blast (** The Union of 2 relations is a relation (Lemma for fun_Un) **) lemma rel_Un: "[| r \ A*B; s \ C*D |] ==> (r \ s) \ (A \ C) * (B \ D)" by blast lemma domain_Diff_eq: "[| \ r; c\b |] ==> domain(r-{}) = domain(r)" by blast lemma range_Diff_eq: "[| \ r; c\a |] ==> range(r-{}) = range(r)" by blast subsection\Image of a Set under a Function or Relation\ lemma image_iff: "b \ r``A \ (\x\A. \r)" by (unfold image_def, blast) lemma image_singleton_iff: "b \ r``{a} \ \r" by (rule image_iff [THEN iff_trans], blast) lemma imageI [intro]: "[| \ r; a\A |] ==> b \ r``A" by (unfold image_def, blast) lemma imageE [elim!]: "[| b: r``A; !!x.[| \ r; x\A |] ==> P |] ==> P" by (unfold image_def, blast) lemma image_subset: "r \ A*B ==> r``C \ B" by blast lemma image_0 [simp]: "r``0 = 0" by blast lemma image_Un [simp]: "r``(A \ B) = (r``A) \ (r``B)" by blast lemma image_UN: "r `` (\x\A. B(x)) = (\x\A. r `` B(x))" by blast lemma Collect_image_eq: "{z \ Sigma(A,B). P(z)} `` C = (\x \ A. {y \ B(x). x \ C & P()})" by blast lemma image_Int_subset: "r``(A \ B) \ (r``A) \ (r``B)" by blast lemma image_Int_square_subset: "(r \ A*A)``B \ (r``B) \ A" by blast lemma image_Int_square: "B\A ==> (r \ A*A)``B = (r``B) \ A" by blast (*Image laws for special relations*) lemma image_0_left [simp]: "0``A = 0" by blast lemma image_Un_left: "(r \ s)``A = (r``A) \ (s``A)" by blast lemma image_Int_subset_left: "(r \ s)``A \ (r``A) \ (s``A)" by blast subsection\Inverse Image of a Set under a Function or Relation\ lemma vimage_iff: "a \ r-``B \ (\y\B. \r)" by (unfold vimage_def image_def converse_def, blast) lemma vimage_singleton_iff: "a \ r-``{b} \ \r" by (rule vimage_iff [THEN iff_trans], blast) lemma vimageI [intro]: "[| \ r; b\B |] ==> a \ r-``B" by (unfold vimage_def, blast) lemma vimageE [elim!]: "[| a: r-``B; !!x.[| \ r; x\B |] ==> P |] ==> P" apply (unfold vimage_def, blast) done lemma vimage_subset: "r \ A*B ==> r-``C \ A" apply (unfold vimage_def) apply (erule converse_type [THEN image_subset]) done lemma vimage_0 [simp]: "r-``0 = 0" by blast lemma vimage_Un [simp]: "r-``(A \ B) = (r-``A) \ (r-``B)" by blast lemma vimage_Int_subset: "r-``(A \ B) \ (r-``A) \ (r-``B)" by blast (*NOT suitable for rewriting*) lemma vimage_eq_UN: "f -``B = (\y\B. f-``{y})" by blast lemma function_vimage_Int: "function(f) ==> f-``(A \ B) = (f-``A) \ (f-``B)" by (unfold function_def, blast) lemma function_vimage_Diff: "function(f) ==> f-``(A-B) = (f-``A) - (f-``B)" by (unfold function_def, blast) lemma function_image_vimage: "function(f) ==> f `` (f-`` A) \ A" by (unfold function_def, blast) lemma vimage_Int_square_subset: "(r \ A*A)-``B \ (r-``B) \ A" by blast lemma vimage_Int_square: "B\A ==> (r \ A*A)-``B = (r-``B) \ A" by blast (*Invese image laws for special relations*) lemma vimage_0_left [simp]: "0-``A = 0" by blast lemma vimage_Un_left: "(r \ s)-``A = (r-``A) \ (s-``A)" by blast lemma vimage_Int_subset_left: "(r \ s)-``A \ (r-``A) \ (s-``A)" by blast (** Converse **) lemma converse_Un [simp]: "converse(A \ B) = converse(A) \ converse(B)" by blast lemma converse_Int [simp]: "converse(A \ B) = converse(A) \ converse(B)" by blast lemma converse_Diff [simp]: "converse(A - B) = converse(A) - converse(B)" by blast lemma converse_UN [simp]: "converse(\x\A. B(x)) = (\x\A. converse(B(x)))" by blast (*Unfolding Inter avoids using excluded middle on A=0*) lemma converse_INT [simp]: "converse(\x\A. B(x)) = (\x\A. converse(B(x)))" apply (unfold Inter_def, blast) done subsection\Powerset Operator\ lemma Pow_0 [simp]: "Pow(0) = {0}" by blast lemma Pow_insert: "Pow (cons(a,A)) = Pow(A) \ {cons(a,X) . X: Pow(A)}" apply (rule equalityI, safe) apply (erule swap) apply (rule_tac a = "x-{a}" in RepFun_eqI, auto) done lemma Un_Pow_subset: "Pow(A) \ Pow(B) \ Pow(A \ B)" by blast lemma UN_Pow_subset: "(\x\A. Pow(B(x))) \ Pow(\x\A. B(x))" by blast lemma subset_Pow_Union: "A \ Pow(\(A))" by blast lemma Union_Pow_eq [simp]: "\(Pow(A)) = A" by blast lemma Union_Pow_iff: "\(A) \ Pow(B) \ A \ Pow(Pow(B))" by blast lemma Pow_Int_eq [simp]: "Pow(A \ B) = Pow(A) \ Pow(B)" by blast lemma Pow_INT_eq: "A\0 ==> Pow(\x\A. B(x)) = (\x\A. Pow(B(x)))" by (blast elim!: not_emptyE) subsection\RepFun\ lemma RepFun_subset: "[| !!x. x\A ==> f(x) \ B |] ==> {f(x). x\A} \ B" by blast lemma RepFun_eq_0_iff [simp]: "{f(x).x\A}=0 \ A=0" by blast lemma RepFun_constant [simp]: "{c. x\A} = (if A=0 then 0 else {c})" by force subsection\Collect\ lemma Collect_subset: "Collect(A,P) \ A" by blast lemma Collect_Un: "Collect(A \ B, P) = Collect(A,P) \ Collect(B,P)" by blast lemma Collect_Int: "Collect(A \ B, P) = Collect(A,P) \ Collect(B,P)" by blast lemma Collect_Diff: "Collect(A - B, P) = Collect(A,P) - Collect(B,P)" by blast lemma Collect_cons: "{x\cons(a,B). P(x)} = (if P(a) then cons(a, {x\B. P(x)}) else {x\B. P(x)})" by (simp, blast) lemma Int_Collect_self_eq: "A \ Collect(A,P) = Collect(A,P)" by blast lemma Collect_Collect_eq [simp]: "Collect(Collect(A,P), Q) = Collect(A, %x. P(x) & Q(x))" by blast lemma Collect_Int_Collect_eq: "Collect(A,P) \ Collect(A,Q) = Collect(A, %x. P(x) & Q(x))" by blast lemma Collect_Union_eq [simp]: "Collect(\x\A. B(x), P) = (\x\A. Collect(B(x), P))" by blast lemma Collect_Int_left: "{x\A. P(x)} \ B = {x \ A \ B. P(x)}" by blast lemma Collect_Int_right: "A \ {x\B. P(x)} = {x \ A \ B. P(x)}" by blast lemma Collect_disj_eq: "{x\A. P(x) | Q(x)} = Collect(A, P) \ Collect(A, Q)" by blast lemma Collect_conj_eq: "{x\A. P(x) & Q(x)} = Collect(A, P) \ Collect(A, Q)" by blast lemmas subset_SIs = subset_refl cons_subsetI subset_consI Union_least UN_least Un_least Inter_greatest Int_greatest RepFun_subset Un_upper1 Un_upper2 Int_lower1 Int_lower2 ML \ val subset_cs = claset_of (\<^context> delrules [@{thm subsetI}, @{thm subsetCE}] addSIs @{thms subset_SIs} addIs [@{thm Union_upper}, @{thm Inter_lower}] addSEs [@{thm cons_subsetE}]); val ZF_cs = claset_of (\<^context> delrules [@{thm equalityI}]); \ end diff --git a/src/ZF/func.thy b/src/ZF/func.thy --- a/src/ZF/func.thy +++ b/src/ZF/func.thy @@ -1,611 +1,611 @@ (* Title: ZF/func.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1991 University of Cambridge *) section\Functions, Function Spaces, Lambda-Abstraction\ theory func imports equalities Sum begin subsection\The Pi Operator: Dependent Function Space\ lemma subset_Sigma_imp_relation: "r \ Sigma(A,B) ==> relation(r)" by (simp add: relation_def, blast) lemma relation_converse_converse [simp]: "relation(r) ==> converse(converse(r)) = r" by (simp add: relation_def, blast) lemma relation_restrict [simp]: "relation(restrict(r,A))" by (simp add: restrict_def relation_def, blast) lemma Pi_iff: "f \ Pi(A,B) \ function(f) & f<=Sigma(A,B) & A<=domain(f)" by (unfold Pi_def, blast) (*For upward compatibility with the former definition*) lemma Pi_iff_old: "f \ Pi(A,B) \ f<=Sigma(A,B) & (\x\A. \!y. : f)" by (unfold Pi_def function_def, blast) lemma fun_is_function: "f \ Pi(A,B) ==> function(f)" by (simp only: Pi_iff) lemma function_imp_Pi: "[|function(f); relation(f)|] ==> f \ domain(f) -> range(f)" by (simp add: Pi_iff relation_def, blast) lemma functionI: "[| !!x y y'. [| :r; :r |] ==> y=y' |] ==> function(r)" by (simp add: function_def, blast) (*Functions are relations*) lemma fun_is_rel: "f \ Pi(A,B) ==> f \ Sigma(A,B)" by (unfold Pi_def, blast) lemma Pi_cong: "[| A=A'; !!x. x \ A' ==> B(x)=B'(x) |] ==> Pi(A,B) = Pi(A',B')" by (simp add: Pi_def cong add: Sigma_cong) (*Sigma_cong, Pi_cong NOT given to Addcongs: they cause flex-flex pairs and the "Check your prover" error. Most Sigmas and Pis are abbreviated as * or -> *) (*Weakening one function type to another; see also Pi_type*) lemma fun_weaken_type: "[| f \ A->B; B<=D |] ==> f \ A->D" by (unfold Pi_def, best) subsection\Function Application\ lemma apply_equality2: "[| : f; : f; f \ Pi(A,B) |] ==> b=c" by (unfold Pi_def function_def, blast) lemma function_apply_equality: "[| : f; function(f) |] ==> f`a = b" by (unfold apply_def function_def, blast) lemma apply_equality: "[| : f; f \ Pi(A,B) |] ==> f`a = b" apply (unfold Pi_def) apply (blast intro: function_apply_equality) done (*Applying a function outside its domain yields 0*) lemma apply_0: "a \ domain(f) ==> f`a = 0" by (unfold apply_def, blast) lemma Pi_memberD: "[| f \ Pi(A,B); c \ f |] ==> \x\A. c = " apply (frule fun_is_rel) apply (blast dest: apply_equality) done lemma function_apply_Pair: "[| function(f); a \ domain(f)|] ==> : f" apply (simp add: function_def, clarify) apply (subgoal_tac "f`a = y", blast) apply (simp add: apply_def, blast) done lemma apply_Pair: "[| f \ Pi(A,B); a \ A |] ==> : f" apply (simp add: Pi_iff) apply (blast intro: function_apply_Pair) done (*Conclusion is flexible -- use rule_tac or else apply_funtype below!*) lemma apply_type [TC]: "[| f \ Pi(A,B); a \ A |] ==> f`a \ B(a)" by (blast intro: apply_Pair dest: fun_is_rel) (*This version is acceptable to the simplifier*) lemma apply_funtype: "[| f \ A->B; a \ A |] ==> f`a \ B" by (blast dest: apply_type) lemma apply_iff: "f \ Pi(A,B) ==> : f \ a \ A & f`a = b" apply (frule fun_is_rel) apply (blast intro!: apply_Pair apply_equality) done (*Refining one Pi type to another*) lemma Pi_type: "[| f \ Pi(A,C); !!x. x \ A ==> f`x \ B(x) |] ==> f \ Pi(A,B)" apply (simp only: Pi_iff) apply (blast dest: function_apply_equality) done (*Such functions arise in non-standard datatypes, ZF/ex/Ntree for instance*) lemma Pi_Collect_iff: "(f \ Pi(A, %x. {y \ B(x). P(x,y)})) \ f \ Pi(A,B) & (\x\A. P(x, f`x))" by (blast intro: Pi_type dest: apply_type) lemma Pi_weaken_type: "[| f \ Pi(A,B); !!x. x \ A ==> B(x)<=C(x) |] ==> f \ Pi(A,C)" by (blast intro: Pi_type dest: apply_type) (** Elimination of membership in a function **) lemma domain_type: "[| \ f; f \ Pi(A,B) |] ==> a \ A" by (blast dest: fun_is_rel) lemma range_type: "[| \ f; f \ Pi(A,B) |] ==> b \ B(a)" by (blast dest: fun_is_rel) lemma Pair_mem_PiD: "[| : f; f \ Pi(A,B) |] ==> a \ A & b \ B(a) & f`a = b" by (blast intro: domain_type range_type apply_equality) subsection\Lambda Abstraction\ lemma lamI: "a \ A ==> \ (\x\A. b(x))" apply (unfold lam_def) apply (erule RepFunI) done lemma lamE: "[| p: (\x\A. b(x)); !!x.[| x \ A; p= |] ==> P |] ==> P" by (simp add: lam_def, blast) lemma lamD: "[| : (\x\A. b(x)) |] ==> c = b(a)" by (simp add: lam_def) lemma lam_type [TC]: "[| !!x. x \ A ==> b(x): B(x) |] ==> (\x\A. b(x)) \ Pi(A,B)" by (simp add: lam_def Pi_def function_def, blast) lemma lam_funtype: "(\x\A. b(x)) \ A -> {b(x). x \ A}" by (blast intro: lam_type) lemma function_lam: "function (\x\A. b(x))" by (simp add: function_def lam_def) lemma relation_lam: "relation (\x\A. b(x))" by (simp add: relation_def lam_def) lemma beta_if [simp]: "(\x\A. b(x)) ` a = (if a \ A then b(a) else 0)" by (simp add: apply_def lam_def, blast) lemma beta: "a \ A ==> (\x\A. b(x)) ` a = b(a)" by (simp add: apply_def lam_def, blast) lemma lam_empty [simp]: "(\x\0. b(x)) = 0" by (simp add: lam_def) lemma domain_lam [simp]: "domain(Lambda(A,b)) = A" by (simp add: lam_def, blast) (*congruence rule for lambda abstraction*) lemma lam_cong [cong]: "[| A=A'; !!x. x \ A' ==> b(x)=b'(x) |] ==> Lambda(A,b) = Lambda(A',b')" by (simp only: lam_def cong add: RepFun_cong) lemma lam_theI: "(!!x. x \ A ==> \!y. Q(x,y)) ==> \f. \x\A. Q(x, f`x)" apply (rule_tac x = "\x\A. THE y. Q (x,y)" in exI) apply simp apply (blast intro: theI) done lemma lam_eqE: "[| (\x\A. f(x)) = (\x\A. g(x)); a \ A |] ==> f(a)=g(a)" by (fast intro!: lamI elim: equalityE lamE) (*Empty function spaces*) lemma Pi_empty1 [simp]: "Pi(0,A) = {0}" by (unfold Pi_def function_def, blast) (*The singleton function*) lemma singleton_fun [simp]: "{} \ {a} -> {b}" by (unfold Pi_def function_def, blast) lemma Pi_empty2 [simp]: "(A->0) = (if A=0 then {0} else 0)" by (unfold Pi_def function_def, force) lemma fun_space_empty_iff [iff]: "(A->X)=0 \ X=0 & (A \ 0)" apply auto apply (fast intro!: equals0I intro: lam_type) done subsection\Extensionality\ (*Semi-extensionality!*) lemma fun_subset: "[| f \ Pi(A,B); g \ Pi(C,D); A<=C; !!x. x \ A ==> f`x = g`x |] ==> f<=g" by (force dest: Pi_memberD intro: apply_Pair) lemma fun_extension: "[| f \ Pi(A,B); g \ Pi(A,D); !!x. x \ A ==> f`x = g`x |] ==> f=g" by (blast del: subsetI intro: subset_refl sym fun_subset) lemma eta [simp]: "f \ Pi(A,B) ==> (\x\A. f`x) = f" apply (rule fun_extension) apply (auto simp add: lam_type apply_type beta) done lemma fun_extension_iff: "[| f \ Pi(A,B); g \ Pi(A,C) |] ==> (\a\A. f`a = g`a) \ f=g" by (blast intro: fun_extension) (*thm by Mark Staples, proof by lcp*) lemma fun_subset_eq: "[| f \ Pi(A,B); g \ Pi(A,C) |] ==> f \ g \ (f = g)" by (blast dest: apply_Pair intro: fun_extension apply_equality [symmetric]) (*Every element of Pi(A,B) may be expressed as a lambda abstraction!*) lemma Pi_lamE: assumes major: "f \ Pi(A,B)" and minor: "!!b. [| \x\A. b(x):B(x); f = (\x\A. b(x)) |] ==> P" shows "P" apply (rule minor) apply (rule_tac [2] eta [symmetric]) apply (blast intro: major apply_type)+ done subsection\Images of Functions\ lemma image_lam: "C \ A ==> (\x\A. b(x)) `` C = {b(x). x \ C}" by (unfold lam_def, blast) lemma Repfun_function_if: "function(f) ==> {f`x. x \ C} = (if C \ domain(f) then f``C else cons(0,f``C))" apply simp apply (intro conjI impI) apply (blast dest: function_apply_equality intro: function_apply_Pair) apply (rule equalityI) apply (blast intro!: function_apply_Pair apply_0) apply (blast dest: function_apply_equality intro: apply_0 [symmetric]) done (*For this lemma and the next, the right-hand side could equivalently be written \x\C. {f`x} *) lemma image_function: "[| function(f); C \ domain(f) |] ==> f``C = {f`x. x \ C}" by (simp add: Repfun_function_if) lemma image_fun: "[| f \ Pi(A,B); C \ A |] ==> f``C = {f`x. x \ C}" apply (simp add: Pi_iff) apply (blast intro: image_function) done lemma image_eq_UN: assumes f: "f \ Pi(A,B)" "C \ A" shows "f``C = (\x\C. {f ` x})" by (auto simp add: image_fun [OF f]) lemma Pi_image_cons: "[| f \ Pi(A,B); x \ A |] ==> f `` cons(x,y) = cons(f`x, f``y)" by (blast dest: apply_equality apply_Pair) subsection\Properties of \<^term>\restrict(f,A)\\ lemma restrict_subset: "restrict(f,A) \ f" by (unfold restrict_def, blast) lemma function_restrictI: "function(f) ==> function(restrict(f,A))" by (unfold restrict_def function_def, blast) lemma restrict_type2: "[| f \ Pi(C,B); A<=C |] ==> restrict(f,A) \ Pi(A,B)" by (simp add: Pi_iff function_def restrict_def, blast) lemma restrict: "restrict(f,A) ` a = (if a \ A then f`a else 0)" by (simp add: apply_def restrict_def, blast) lemma restrict_empty [simp]: "restrict(f,0) = 0" by (unfold restrict_def, simp) lemma restrict_iff: "z \ restrict(r,A) \ z \ r & (\x\A. \y. z = \x, y\)" by (simp add: restrict_def) lemma restrict_restrict [simp]: "restrict(restrict(r,A),B) = restrict(r, A \ B)" by (unfold restrict_def, blast) lemma domain_restrict [simp]: "domain(restrict(f,C)) = domain(f) \ C" apply (unfold restrict_def) apply (auto simp add: domain_def) done lemma restrict_idem: "f \ Sigma(A,B) ==> restrict(f,A) = f" by (simp add: restrict_def, blast) (*converse probably holds too*) lemma domain_restrict_idem: "[| domain(r) \ A; relation(r) |] ==> restrict(r,A) = r" by (simp add: restrict_def relation_def, blast) lemma domain_restrict_lam [simp]: "domain(restrict(Lambda(A,f),C)) = A \ C" apply (unfold restrict_def lam_def) apply (rule equalityI) apply (auto simp add: domain_iff) done lemma restrict_if [simp]: "restrict(f,A) ` a = (if a \ A then f`a else 0)" by (simp add: restrict apply_0) lemma restrict_lam_eq: "A<=C ==> restrict(\x\C. b(x), A) = (\x\A. b(x))" by (unfold restrict_def lam_def, auto) lemma fun_cons_restrict_eq: "f \ cons(a, b) -> B ==> f = cons(, restrict(f, b))" apply (rule equalityI) prefer 2 apply (blast intro: apply_Pair restrict_subset [THEN subsetD]) apply (auto dest!: Pi_memberD simp add: restrict_def lam_def) done subsection\Unions of Functions\ (** The Union of a set of COMPATIBLE functions is a function **) lemma function_Union: "[| \x\S. function(x); \x\S. \y\S. x<=y | y<=x |] ==> function(\(S))" by (unfold function_def, blast) lemma fun_Union: "[| \f\S. \C D. f \ C->D; \f\S. \y\S. f<=y | y<=f |] ==> \(S) \ domain(\(S)) -> range(\(S))" apply (unfold Pi_def) apply (blast intro!: rel_Union function_Union) done -lemma gen_relation_Union [rule_format]: - "\f\F. relation(f) \ relation(\(F))" +lemma gen_relation_Union: + "(\f. f\F \ relation(f)) \ relation(\(F))" by (simp add: relation_def) (** The Union of 2 disjoint functions is a function **) lemmas Un_rls = Un_subset_iff SUM_Un_distrib1 prod_Un_distrib2 subset_trans [OF _ Un_upper1] subset_trans [OF _ Un_upper2] lemma fun_disjoint_Un: "[| f \ A->B; g \ C->D; A \ C = 0 |] ==> (f \ g) \ (A \ C) -> (B \ D)" (*Prove the product and domain subgoals using distributive laws*) apply (simp add: Pi_iff extension Un_rls) apply (unfold function_def, blast) done lemma fun_disjoint_apply1: "a \ domain(g) ==> (f \ g)`a = f`a" by (simp add: apply_def, blast) lemma fun_disjoint_apply2: "c \ domain(f) ==> (f \ g)`c = g`c" by (simp add: apply_def, blast) subsection\Domain and Range of a Function or Relation\ lemma domain_of_fun: "f \ Pi(A,B) ==> domain(f)=A" by (unfold Pi_def, blast) lemma apply_rangeI: "[| f \ Pi(A,B); a \ A |] ==> f`a \ range(f)" by (erule apply_Pair [THEN rangeI], assumption) lemma range_of_fun: "f \ Pi(A,B) ==> f \ A->range(f)" by (blast intro: Pi_type apply_rangeI) subsection\Extensions of Functions\ lemma fun_extend: "[| f \ A->B; c\A |] ==> cons(,f) \ cons(c,A) -> cons(b,B)" apply (frule singleton_fun [THEN fun_disjoint_Un], blast) apply (simp add: cons_eq) done lemma fun_extend3: "[| f \ A->B; c\A; b \ B |] ==> cons(,f) \ cons(c,A) -> B" by (blast intro: fun_extend [THEN fun_weaken_type]) lemma extend_apply: "c \ domain(f) ==> cons(,f)`a = (if a=c then b else f`a)" by (auto simp add: apply_def) lemma fun_extend_apply [simp]: "[| f \ A->B; c\A |] ==> cons(,f)`a = (if a=c then b else f`a)" apply (rule extend_apply) apply (simp add: Pi_def, blast) done lemmas singleton_apply = apply_equality [OF singletonI singleton_fun, simp] (*For Finite.ML. Inclusion of right into left is easy*) lemma cons_fun_eq: "c \ A ==> cons(c,A) -> B = (\f \ A->B. \b\B. {cons(, f)})" apply (rule equalityI) apply (safe elim!: fun_extend3) (*Inclusion of left into right*) apply (subgoal_tac "restrict (x, A) \ A -> B") prefer 2 apply (blast intro: restrict_type2) apply (rule UN_I, assumption) apply (rule apply_funtype [THEN UN_I]) apply assumption apply (rule consI1) apply (simp (no_asm)) apply (rule fun_extension) apply assumption apply (blast intro: fun_extend) apply (erule consE, simp_all) done lemma succ_fun_eq: "succ(n) -> B = (\f \ n->B. \b\B. {cons(, f)})" by (simp add: succ_def mem_not_refl cons_fun_eq) subsection\Function Updates\ definition update :: "[i,i,i] => i" where "update(f,a,b) == \x\cons(a, domain(f)). if(x=a, b, f`x)" nonterminal updbinds and updbind syntax (* Let expressions *) "_updbind" :: "[i, i] => updbind" (\(2_ :=/ _)\) "" :: "updbind => updbinds" (\_\) "_updbinds" :: "[updbind, updbinds] => updbinds" (\_,/ _\) "_Update" :: "[i, updbinds] => i" (\_/'((_)')\ [900,0] 900) translations "_Update (f, _updbinds(b,bs))" == "_Update (_Update(f,b), bs)" "f(x:=y)" == "CONST update(f,x,y)" lemma update_apply [simp]: "f(x:=y) ` z = (if z=x then y else f`z)" apply (simp add: update_def) apply (case_tac "z \ domain(f)") apply (simp_all add: apply_0) done lemma update_idem: "[| f`x = y; f \ Pi(A,B); x \ A |] ==> f(x:=y) = f" apply (unfold update_def) apply (simp add: domain_of_fun cons_absorb) apply (rule fun_extension) apply (best intro: apply_type if_type lam_type, assumption, simp) done (* [| f \ Pi(A, B); x \ A |] ==> f(x := f`x) = f *) declare refl [THEN update_idem, simp] lemma domain_update [simp]: "domain(f(x:=y)) = cons(x, domain(f))" by (unfold update_def, simp) lemma update_type: "[| f \ Pi(A,B); x \ A; y \ B(x) |] ==> f(x:=y) \ Pi(A, B)" apply (unfold update_def) apply (simp add: domain_of_fun cons_absorb apply_funtype lam_type) done subsection\Monotonicity Theorems\ subsubsection\Replacement in its Various Forms\ (*Not easy to express monotonicity in P, since any "bigger" predicate would have to be single-valued*) lemma Replace_mono: "A<=B ==> Replace(A,P) \ Replace(B,P)" by (blast elim!: ReplaceE) lemma RepFun_mono: "A<=B ==> {f(x). x \ A} \ {f(x). x \ B}" by blast lemma Pow_mono: "A<=B ==> Pow(A) \ Pow(B)" by blast lemma Union_mono: "A<=B ==> \(A) \ \(B)" by blast lemma UN_mono: "[| A<=C; !!x. x \ A ==> B(x)<=D(x) |] ==> (\x\A. B(x)) \ (\x\C. D(x))" by blast (*Intersection is ANTI-monotonic. There are TWO premises! *) lemma Inter_anti_mono: "[| A<=B; A\0 |] ==> \(B) \ \(A)" by blast lemma cons_mono: "C<=D ==> cons(a,C) \ cons(a,D)" by blast lemma Un_mono: "[| A<=C; B<=D |] ==> A \ B \ C \ D" by blast lemma Int_mono: "[| A<=C; B<=D |] ==> A \ B \ C \ D" by blast lemma Diff_mono: "[| A<=C; D<=B |] ==> A-B \ C-D" by blast subsubsection\Standard Products, Sums and Function Spaces\ lemma Sigma_mono [rule_format]: "[| A<=C; !!x. x \ A \ B(x) \ D(x) |] ==> Sigma(A,B) \ Sigma(C,D)" by blast lemma sum_mono: "[| A<=C; B<=D |] ==> A+B \ C+D" by (unfold sum_def, blast) (*Note that B->A and C->A are typically disjoint!*) lemma Pi_mono: "B<=C ==> A->B \ A->C" by (blast intro: lam_type elim: Pi_lamE) lemma lam_mono: "A<=B ==> Lambda(A,c) \ Lambda(B,c)" apply (unfold lam_def) apply (erule RepFun_mono) done subsubsection\Converse, Domain, Range, Field\ lemma converse_mono: "r<=s ==> converse(r) \ converse(s)" by blast lemma domain_mono: "r<=s ==> domain(r)<=domain(s)" by blast lemmas domain_rel_subset = subset_trans [OF domain_mono domain_subset] lemma range_mono: "r<=s ==> range(r)<=range(s)" by blast lemmas range_rel_subset = subset_trans [OF range_mono range_subset] lemma field_mono: "r<=s ==> field(r)<=field(s)" by blast lemma field_rel_subset: "r \ A*A ==> field(r) \ A" by (erule field_mono [THEN subset_trans], blast) subsubsection\Images\ lemma image_pair_mono: "[| !! x y. :r ==> :s; A<=B |] ==> r``A \ s``B" by blast lemma vimage_pair_mono: "[| !! x y. :r ==> :s; A<=B |] ==> r-``A \ s-``B" by blast lemma image_mono: "[| r<=s; A<=B |] ==> r``A \ s``B" by blast lemma vimage_mono: "[| r<=s; A<=B |] ==> r-``A \ s-``B" by blast lemma Collect_mono: "[| A<=B; !!x. x \ A ==> P(x) \ Q(x) |] ==> Collect(A,P) \ Collect(B,Q)" by blast (*Used in intr_elim.ML and in individual datatype definitions*) lemmas basic_monos = subset_refl imp_refl disj_mono conj_mono ex_mono Collect_mono Part_mono in_mono (* Useful with simp; contributed by Clemens Ballarin. *) lemma bex_image_simp: "[| f \ Pi(X, Y); A \ X |] ==> (\x\f``A. P(x)) \ (\x\A. P(f`x))" apply safe apply rule prefer 2 apply assumption apply (simp add: apply_equality) apply (blast intro: apply_Pair) done lemma ball_image_simp: "[| f \ Pi(X, Y); A \ X |] ==> (\x\f``A. P(x)) \ (\x\A. P(f`x))" apply safe apply (blast intro: apply_Pair) apply (drule bspec) apply assumption apply (simp add: apply_equality) done end