--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL4/Importer.thy Sun Apr 01 14:50:47 2012 +0200
@@ -0,0 +1,236 @@
+(* Title: HOL/Import/Importer.thy
+ Author: Sebastian Skalberg, TU Muenchen
+*)
+
+theory Importer
+imports Main
+keywords
+ "import_segment" "import_theory" "end_import" "setup_theory" "end_setup"
+ "setup_dump" "append_dump" "flush_dump" "ignore_thms" "type_maps"
+ "def_maps" "thm_maps" "const_renames" "const_moves" "const_maps" :: thy_decl and ">"
+uses "shuffler.ML" "import_rews.ML" ("proof_kernel.ML") ("replay.ML") ("import.ML")
+begin
+
+setup {* Shuffler.setup #> importer_setup *}
+
+parse_ast_translation smarter_trueprop_parsing
+
+lemma conj_norm [shuffle_rule]: "(A & B ==> PROP C) == ([| A ; B |] ==> PROP C)"
+proof
+ assume "A & B ==> PROP C" A B
+ thus "PROP C"
+ by auto
+next
+ assume "[| A; B |] ==> PROP C" "A & B"
+ thus "PROP C"
+ by auto
+qed
+
+lemma imp_norm [shuffle_rule]: "(Trueprop (A --> B)) == (A ==> B)"
+proof
+ assume "A --> B" A
+ thus B ..
+next
+ assume "A ==> B"
+ thus "A --> B"
+ by auto
+qed
+
+lemma all_norm [shuffle_rule]: "(Trueprop (ALL x. P x)) == (!!x. P x)"
+proof
+ fix x
+ assume "ALL x. P x"
+ thus "P x" ..
+next
+ assume "!!x. P x"
+ thus "ALL x. P x"
+ ..
+qed
+
+lemma ex_norm [shuffle_rule]: "(EX x. P x ==> PROP Q) == (!!x. P x ==> PROP Q)"
+proof
+ fix x
+ assume ex: "EX x. P x ==> PROP Q"
+ assume "P x"
+ hence "EX x. P x" ..
+ with ex show "PROP Q" .
+next
+ assume allx: "!!x. P x ==> PROP Q"
+ assume "EX x. P x"
+ hence p: "P (SOME x. P x)"
+ ..
+ from allx
+ have "P (SOME x. P x) ==> PROP Q"
+ .
+ with p
+ show "PROP Q"
+ by auto
+qed
+
+lemma eq_norm [shuffle_rule]: "Trueprop (t = u) == (t == u)"
+proof
+ assume "t = u"
+ thus "t == u" by simp
+next
+ assume "t == u"
+ thus "t = u"
+ by simp
+qed
+
+section {* General Setup *}
+
+lemma eq_imp: "P = Q \<Longrightarrow> P \<longrightarrow> Q"
+ by auto
+
+lemma HOLallI: "(!! bogus. P bogus) \<Longrightarrow> (ALL bogus. P bogus)"
+proof -
+ assume "!! bogus. P bogus"
+ thus "ALL x. P x"
+ ..
+qed
+
+consts
+ ONE_ONE :: "('a => 'b) => bool"
+
+defs
+ ONE_ONE_DEF: "ONE_ONE f == ALL x y. f x = f y --> x = y"
+
+lemma ONE_ONE_rew: "ONE_ONE f = inj_on f UNIV"
+ by (simp add: ONE_ONE_DEF inj_on_def)
+
+lemma INFINITY_AX: "EX (f::ind \<Rightarrow> ind). (inj f & ~(surj f))"
+proof (rule exI,safe)
+ show "inj Suc_Rep"
+ by (rule injI) (rule Suc_Rep_inject)
+next
+ assume "surj Suc_Rep"
+ hence "ALL y. EX x. y = Suc_Rep x"
+ by (simp add: surj_def)
+ hence "EX x. Zero_Rep = Suc_Rep x"
+ by (rule spec)
+ thus False
+ proof (rule exE)
+ fix x
+ assume "Zero_Rep = Suc_Rep x"
+ hence "Suc_Rep x = Zero_Rep"
+ ..
+ with Suc_Rep_not_Zero_Rep
+ show False
+ ..
+ qed
+qed
+
+lemma EXISTS_DEF: "Ex P = P (Eps P)"
+proof (rule iffI)
+ assume "Ex P"
+ thus "P (Eps P)"
+ ..
+next
+ assume "P (Eps P)"
+ thus "Ex P"
+ ..
+qed
+
+consts
+ TYPE_DEFINITION :: "('a => bool) => ('b => 'a) => bool"
+
+defs
+ TYPE_DEFINITION: "TYPE_DEFINITION p rep == ((ALL x y. (rep x = rep y) --> (x = y)) & (ALL x. (p x = (EX y. x = rep y))))"
+
+lemma ex_imp_nonempty: "Ex P ==> EX x. x : (Collect P)"
+ by simp
+
+lemma light_ex_imp_nonempty: "P t ==> EX x. x : (Collect P)"
+proof -
+ assume "P t"
+ hence "EX x. P x"
+ ..
+ thus ?thesis
+ by (rule ex_imp_nonempty)
+qed
+
+lemma light_imp_as: "[| Q --> P; P --> Q |] ==> P = Q"
+ by blast
+
+lemma typedef_hol2hol4:
+ assumes a: "type_definition (Rep::'a=>'b) Abs (Collect P)"
+ shows "EX rep. TYPE_DEFINITION P (rep::'a=>'b)"
+proof -
+ from a
+ have td: "(ALL x. P (Rep x)) & (ALL x. Abs (Rep x) = x) & (ALL y. P y \<longrightarrow> Rep (Abs y) = y)"
+ by (simp add: type_definition_def)
+ have ed: "TYPE_DEFINITION P Rep"
+ proof (auto simp add: TYPE_DEFINITION)
+ fix x y
+ assume b: "Rep x = Rep y"
+ from td have "x = Abs (Rep x)"
+ by auto
+ also have "Abs (Rep x) = Abs (Rep y)"
+ by (simp add: b)
+ also from td have "Abs (Rep y) = y"
+ by auto
+ finally show "x = y" .
+ next
+ fix x
+ assume "P x"
+ with td
+ have "Rep (Abs x) = x"
+ by auto
+ hence "x = Rep (Abs x)"
+ ..
+ thus "EX y. x = Rep y"
+ ..
+ next
+ fix y
+ from td
+ show "P (Rep y)"
+ by auto
+ qed
+ show ?thesis
+ apply (rule exI [of _ Rep])
+ apply (rule ed)
+ .
+qed
+
+lemma typedef_hol2hollight:
+ assumes a: "type_definition (Rep::'a=>'b) Abs (Collect P)"
+ shows "(Abs (Rep a) = a) & (P r = (Rep (Abs r) = r))"
+proof
+ from a
+ show "Abs (Rep a) = a"
+ by (rule type_definition.Rep_inverse)
+next
+ show "P r = (Rep (Abs r) = r)"
+ proof
+ assume "P r"
+ hence "r \<in> (Collect P)"
+ by simp
+ with a
+ show "Rep (Abs r) = r"
+ by (rule type_definition.Abs_inverse)
+ next
+ assume ra: "Rep (Abs r) = r"
+ from a
+ have "Rep (Abs r) \<in> (Collect P)"
+ by (rule type_definition.Rep)
+ thus "P r"
+ by (simp add: ra)
+ qed
+qed
+
+lemma termspec_help: "[| Ex P ; c == Eps P |] ==> P c"
+ apply simp
+ apply (rule someI_ex)
+ .
+
+lemma typedef_helper: "EX x. P x \<Longrightarrow> EX x. x \<in> (Collect P)"
+ by simp
+
+use "proof_kernel.ML"
+use "replay.ML"
+use "import.ML"
+
+setup Import.setup
+
+end
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL4/README Sun Apr 01 14:50:47 2012 +0200
@@ -0,0 +1,19 @@
+
+ATTENTION! This is largely outdated. The hint to PROOF_DIRS might be
+everything which is still relevant.
+
+All the files in this directory (except this README, Importer.thy, and
+ROOT.ML) are automatically generated. Edit the files in
+../Generate-HOL and run "isabelle make HOL-Complex-Generate-HOL" in
+~~/src/HOL, if something needs to be changed.
+
+To build the logic in this directory, simply do a "isabelle make
+HOL-Import-HOL" in ~~/src/HOL.
+
+Note that the quick_and_dirty flag is on as default for this
+directory, which means that the original external proofs are not consulted
+at all. If a real replay of the external proofs is desired, get and
+unpack the external proof objects to somewhere on your harddisk, and set
+the variable PROOF_DIRS to the directory where the directory "hol4"
+is. Now edit the ROOT.ML file to unset the quick_and_dirty flag and
+do "isabelle make HOL-Import-HOL" in ~~/src/HOL.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL4/import.ML Sun Apr 01 14:50:47 2012 +0200
@@ -0,0 +1,280 @@
+(* Title: HOL/Import/import.ML
+ Author: Sebastian Skalberg (TU Muenchen)
+*)
+
+signature IMPORT =
+sig
+ val debug : bool Unsynchronized.ref
+ val import_tac : Proof.context -> string * string -> tactic
+ val setup : theory -> theory
+end
+
+structure ImportData = Theory_Data
+(
+ type T = ProofKernel.thm option array option (* FIXME mutable array !??!!! *)
+ val empty = NONE
+ val extend = I
+ fun merge _ = NONE
+)
+
+structure Import : IMPORT =
+struct
+
+val debug = Unsynchronized.ref false
+fun message s = if !debug then writeln s else ()
+
+fun import_tac ctxt (thyname, thmname) =
+ if ! quick_and_dirty
+ then Skip_Proof.cheat_tac (Proof_Context.theory_of ctxt)
+ else
+ fn th =>
+ let
+ val thy = Proof_Context.theory_of ctxt
+ val prem = hd (prems_of th)
+ val _ = message ("Import_tac: thyname=" ^ thyname ^ ", thmname=" ^ thmname)
+ val _ = message ("Import trying to prove " ^ Syntax.string_of_term ctxt prem)
+ val int_thms = case ImportData.get thy of
+ NONE => fst (Replay.setup_int_thms thyname thy)
+ | SOME a => a
+ val proof = snd (ProofKernel.import_proof thyname thmname thy) thy
+ val importerthm = snd (Replay.replay_proof int_thms thyname thmname proof thy)
+ val thm = snd (ProofKernel.to_isa_thm importerthm)
+ val rew = ProofKernel.rewrite_importer_term (concl_of thm) thy
+ val thm = Thm.equal_elim rew thm
+ val prew = ProofKernel.rewrite_importer_term prem thy
+ val prem' = #2 (Logic.dest_equals (prop_of prew))
+ val _ = message ("Import proved " ^ Display.string_of_thm ctxt thm)
+ val thm = ProofKernel.disambiguate_frees thm
+ val _ = message ("Disambiguate: " ^ Display.string_of_thm ctxt thm)
+ in
+ case Shuffler.set_prop thy prem' [("",thm)] of
+ SOME (_,thm) =>
+ let
+ val _ = if prem' aconv (prop_of thm)
+ then message "import: Terms match up"
+ else message "import: Terms DO NOT match up"
+ val thm' = Thm.equal_elim (Thm.symmetric prew) thm
+ val res = Thm.bicompose true (false,thm',0) 1 th
+ in
+ res
+ end
+ | NONE => (message "import: set_prop didn't succeed"; no_tac th)
+ end
+
+val setup = Method.setup @{binding import}
+ (Scan.lift (Args.name -- Args.name) >>
+ (fn arg => fn ctxt => SIMPLE_METHOD (import_tac ctxt arg)))
+ "import external theorem"
+
+(* Parsers *)
+
+val import_segment = Parse.name >> set_import_segment
+
+
+val import_theory = (Parse.name -- Parse.name) >> (fn (location, thyname) =>
+ fn thy =>
+ thy |> set_generating_thy thyname
+ |> Sign.add_path thyname
+ |> add_dump ("setup_theory " ^ quote location ^ " " ^ thyname))
+
+fun end_import toks =
+ Scan.succeed
+ (fn thy =>
+ let
+ val thyname = get_generating_thy thy
+ val segname = get_import_segment thy
+ val (int_thms,facts) = Replay.setup_int_thms thyname thy
+ val thmnames = filter_out (should_ignore thyname thy) facts
+ fun replay thy = Replay.import_thms thyname int_thms thmnames thy
+ in
+ thy |> clear_import_thy
+ |> set_segment thyname segname
+ |> set_used_names facts
+ |> replay
+ |> clear_used_names
+ |> export_importer_pending
+ |> Sign.parent_path
+ |> dump_import_thy thyname
+ |> add_dump ";end_setup"
+ end) toks
+
+val ignore_thms = Scan.repeat1 Parse.name
+ >> (fn ignored =>
+ fn thy =>
+ let
+ val thyname = get_import_thy thy
+ in
+ Library.foldl (fn (thy,thmname) => ignore_importer thyname thmname thy) (thy,ignored)
+ end)
+
+val thm_maps = Scan.repeat1 (Parse.name --| Parse.$$$ ">" -- Parse.xname)
+ >> (fn thmmaps =>
+ fn thy =>
+ let
+ val thyname = get_import_thy thy
+ in
+ Library.foldl (fn (thy,(hol,isa)) => add_importer_mapping thyname hol isa thy) (thy,thmmaps)
+ end)
+
+val type_maps = Scan.repeat1 (Parse.name --| Parse.$$$ ">" -- Parse.xname)
+ >> (fn typmaps =>
+ fn thy =>
+ let
+ val thyname = get_import_thy thy
+ in
+ Library.foldl (fn (thy,(hol,isa)) => add_importer_type_mapping thyname hol false isa thy) (thy,typmaps)
+ end)
+
+val def_maps = Scan.repeat1 (Parse.name --| Parse.$$$ ">" -- Parse.xname)
+ >> (fn defmaps =>
+ fn thy =>
+ let
+ val thyname = get_import_thy thy
+ in
+ Library.foldl (fn (thy,(hol,isa)) => add_defmap thyname hol isa thy) (thy,defmaps)
+ end)
+
+val const_renames = Scan.repeat1 (Parse.name --| Parse.$$$ ">" -- Parse.name)
+ >> (fn renames =>
+ fn thy =>
+ let
+ val thyname = get_import_thy thy
+ in
+ Library.foldl (fn (thy,(hol,isa)) => add_importer_const_renaming thyname hol isa thy) (thy,renames)
+ end)
+
+val const_maps = Scan.repeat1 (Parse.name --| Parse.$$$ ">" -- Parse.xname -- Scan.option (Parse.$$$ "::" |-- Parse.typ))
+ >> (fn constmaps =>
+ fn thy =>
+ let
+ val thyname = get_import_thy thy
+ in
+ Library.foldl (fn (thy,((hol,isa),NONE)) => add_importer_const_mapping thyname hol false isa thy
+ | (thy,((hol,isa),SOME ty)) => add_importer_const_wt_mapping thyname hol false isa (Syntax.read_typ_global thy ty) thy) (thy,constmaps)
+ end)
+
+val const_moves = Scan.repeat1 (Parse.name --| Parse.$$$ ">" -- Parse.xname -- Scan.option (Parse.$$$ "::" |-- Parse.typ))
+ >> (fn constmaps =>
+ fn thy =>
+ let
+ val thyname = get_import_thy thy
+ in
+ Library.foldl (fn (thy,((hol,isa),NONE)) => add_importer_const_mapping thyname hol true isa thy
+ | (thy,((hol,isa),SOME ty)) => add_importer_const_wt_mapping thyname hol true isa (Syntax.read_typ_global thy ty) thy) (thy,constmaps)
+ end)
+
+fun import_thy location s =
+ let
+ val src_location = Path.append (Path.explode location) (Path.basic (s ^ ".imp"))
+ val is = TextIO.openIn (File.platform_path src_location)
+ val inp = TextIO.inputAll is
+ val _ = TextIO.closeIn is
+ val orig_source = Source.of_string inp
+ val symb_source = Symbol.source orig_source
+ val lexes =
+ (Scan.make_lexicon
+ (map Symbol.explode ["import_segment","ignore_thms","import","end",">","::","const_maps","const_moves","thm_maps","const_renames","type_maps","def_maps"]),
+ Scan.empty_lexicon)
+ val token_source = Token.source {do_recover = NONE} (K lexes) Position.start symb_source
+ val token_list = filter_out (Token.is_kind Token.Space) (Source.exhaust token_source)
+ val import_segmentP = Parse.$$$ "import_segment" |-- import_segment
+ val type_mapsP = Parse.$$$ "type_maps" |-- type_maps
+ val def_mapsP = Parse.$$$ "def_maps" |-- def_maps
+ val const_mapsP = Parse.$$$ "const_maps" |-- const_maps
+ val const_movesP = Parse.$$$ "const_moves" |-- const_moves
+ val const_renamesP = Parse.$$$ "const_renames" |-- const_renames
+ val ignore_thmsP = Parse.$$$ "ignore_thms" |-- ignore_thms
+ val thm_mapsP = Parse.$$$ "thm_maps" |-- thm_maps
+ val parser = type_mapsP || def_mapsP || const_mapsP || const_movesP || const_renamesP || thm_mapsP || ignore_thmsP || import_segmentP
+ val importP = Parse.$$$ "import" |-- Scan.repeat parser --| Parse.$$$ "end"
+ fun apply [] thy = thy
+ | apply (f::fs) thy = apply fs (f thy)
+ in
+ apply (set_replaying_thy s :: Sign.add_path s :: fst (importP token_list))
+ end
+
+fun create_int_thms thyname thy =
+ if ! quick_and_dirty
+ then thy
+ else
+ case ImportData.get thy of
+ NONE => ImportData.put (SOME (fst (Replay.setup_int_thms thyname thy))) thy
+ | SOME _ => error "Import data not closed - forgotten an end_setup, mayhap?"
+
+fun clear_int_thms thy =
+ if ! quick_and_dirty
+ then thy
+ else
+ case ImportData.get thy of
+ NONE => error "Import data already cleared - forgotten a setup_theory?"
+ | SOME _ => ImportData.put NONE thy
+
+val setup_theory = (Parse.name -- Parse.name)
+ >>
+ (fn (location, thyname) =>
+ fn thy =>
+ (case Importer_DefThy.get thy of
+ NoImport => thy
+ | Generating _ => error "Currently generating a theory!"
+ | Replaying _ => thy |> clear_import_thy)
+ |> import_thy location thyname
+ |> create_int_thms thyname)
+
+fun end_setup toks =
+ Scan.succeed
+ (fn thy =>
+ let
+ val thyname = get_import_thy thy
+ val segname = get_import_segment thy
+ in
+ thy |> set_segment thyname segname
+ |> clear_import_thy
+ |> clear_int_thms
+ |> Sign.parent_path
+ end) toks
+
+val set_dump = Parse.string -- Parse.string >> setup_dump
+fun fl_dump toks = Scan.succeed flush_dump toks
+val append_dump = (Parse.verbatim || Parse.string) >> add_dump
+
+val _ =
+ (Outer_Syntax.command @{command_spec "import_segment"} "set import segment name"
+ (import_segment >> Toplevel.theory);
+ Outer_Syntax.command @{command_spec "import_theory"} "set default external theory name"
+ (import_theory >> Toplevel.theory);
+ Outer_Syntax.command @{command_spec "end_import"} "end external import"
+ (end_import >> Toplevel.theory);
+ Outer_Syntax.command @{command_spec "setup_theory"} "setup external theory replaying"
+ (setup_theory >> Toplevel.theory);
+ Outer_Syntax.command @{command_spec "end_setup"} "end external setup"
+ (end_setup >> Toplevel.theory);
+ Outer_Syntax.command @{command_spec "setup_dump"} "setup the dump file name"
+ (set_dump >> Toplevel.theory);
+ Outer_Syntax.command @{command_spec "append_dump"} "add text to dump file"
+ (append_dump >> Toplevel.theory);
+ Outer_Syntax.command @{command_spec "flush_dump"} "write the dump to file"
+ (fl_dump >> Toplevel.theory);
+ Outer_Syntax.command @{command_spec "ignore_thms"}
+ "theorems to ignore in next external theory import"
+ (ignore_thms >> Toplevel.theory);
+ Outer_Syntax.command @{command_spec "type_maps"}
+ "map external type names to existing Isabelle/HOL type names"
+ (type_maps >> Toplevel.theory);
+ Outer_Syntax.command @{command_spec "def_maps"}
+ "map external constant names to their primitive definitions"
+ (def_maps >> Toplevel.theory);
+ Outer_Syntax.command @{command_spec "thm_maps"}
+ "map external theorem names to existing Isabelle/HOL theorem names"
+ (thm_maps >> Toplevel.theory);
+ Outer_Syntax.command @{command_spec "const_renames"}
+ "rename external const names"
+ (const_renames >> Toplevel.theory);
+ Outer_Syntax.command @{command_spec "const_moves"}
+ "rename external const names to other external constants"
+ (const_moves >> Toplevel.theory);
+ Outer_Syntax.command @{command_spec "const_maps"}
+ "map external const names to existing Isabelle/HOL const names"
+ (const_maps >> Toplevel.theory));
+
+end
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL4/import_rews.ML Sun Apr 01 14:50:47 2012 +0200
@@ -0,0 +1,633 @@
+(* Title: HOL/Import/import_rews.ML
+ Author: Sebastian Skalberg (TU Muenchen)
+*)
+
+structure StringPair = Table(type key = string * string val ord = prod_ord string_ord string_ord);
+
+type holthm = (term * term) list * thm
+
+datatype ImportStatus =
+ NoImport
+ | Generating of string
+ | Replaying of string
+
+structure Importer_DefThy = Theory_Data
+(
+ type T = ImportStatus
+ val empty = NoImport
+ val extend = I
+ fun merge (NoImport, NoImport) = NoImport
+ | merge _ = (warning "Import status set during merge"; NoImport)
+)
+
+structure ImportSegment = Theory_Data
+(
+ type T = string
+ val empty = ""
+ val extend = I
+ fun merge ("",arg) = arg
+ | merge (arg,"") = arg
+ | merge (s1,s2) =
+ if s1 = s2
+ then s1
+ else error "Trying to merge two different import segments"
+)
+
+val get_import_segment = ImportSegment.get
+val set_import_segment = ImportSegment.put
+
+structure Importer_UNames = Theory_Data
+(
+ type T = string list
+ val empty = []
+ val extend = I
+ fun merge ([], []) = []
+ | merge _ = error "Used names not empty during merge"
+)
+
+structure Importer_Dump = Theory_Data
+(
+ type T = string * string * string list
+ val empty = ("","",[])
+ val extend = I
+ fun merge (("","",[]),("","",[])) = ("","",[])
+ | merge _ = error "Dump data not empty during merge"
+)
+
+structure Importer_Moves = Theory_Data
+(
+ type T = string Symtab.table
+ val empty = Symtab.empty
+ val extend = I
+ fun merge data = Symtab.merge (K true) data
+)
+
+structure Importer_Imports = Theory_Data
+(
+ type T = string Symtab.table
+ val empty = Symtab.empty
+ val extend = I
+ fun merge data = Symtab.merge (K true) data
+)
+
+fun get_segment2 thyname thy =
+ Symtab.lookup (Importer_Imports.get thy) thyname
+
+fun set_segment thyname segname thy =
+ let
+ val imps = Importer_Imports.get thy
+ val imps' = Symtab.update_new (thyname,segname) imps
+ in
+ Importer_Imports.put imps' thy
+ end
+
+structure Importer_CMoves = Theory_Data
+(
+ type T = string Symtab.table
+ val empty = Symtab.empty
+ val extend = I
+ fun merge data = Symtab.merge (K true) data
+)
+
+structure Importer_Maps = Theory_Data
+(
+ type T = string option StringPair.table
+ val empty = StringPair.empty
+ val extend = I
+ fun merge data = StringPair.merge (K true) data
+)
+
+structure Importer_Thms = Theory_Data
+(
+ type T = holthm StringPair.table
+ val empty = StringPair.empty
+ val extend = I
+ fun merge data = StringPair.merge (K true) data
+)
+
+structure Importer_ConstMaps = Theory_Data
+(
+ type T = (bool * string * typ option) StringPair.table
+ val empty = StringPair.empty
+ val extend = I
+ fun merge data = StringPair.merge (K true) data
+)
+
+structure Importer_Rename = Theory_Data
+(
+ type T = string StringPair.table
+ val empty = StringPair.empty
+ val extend = I
+ fun merge data = StringPair.merge (K true) data
+)
+
+structure Importer_DefMaps = Theory_Data
+(
+ type T = string StringPair.table
+ val empty = StringPair.empty
+ val extend = I
+ fun merge data = StringPair.merge (K true) data
+)
+
+structure Importer_TypeMaps = Theory_Data
+(
+ type T = (bool * string) StringPair.table
+ val empty = StringPair.empty
+ val extend = I
+ fun merge data = StringPair.merge (K true) data
+)
+
+structure Importer_Pending = Theory_Data
+(
+ type T = ((term * term) list * thm) StringPair.table
+ val empty = StringPair.empty
+ val extend = I
+ fun merge data = StringPair.merge (K true) data
+)
+
+structure Importer_Rewrites = Theory_Data
+(
+ type T = thm list
+ val empty = []
+ val extend = I
+ val merge = Thm.merge_thms
+)
+
+val importer_debug = Unsynchronized.ref false
+fun message s = if !importer_debug then writeln s else ()
+
+fun add_importer_rewrite (Context.Theory thy, th) =
+ let
+ val _ = message "Adding external rewrite"
+ val th1 = th RS @{thm eq_reflection}
+ handle THM _ => th
+ val current_rews = Importer_Rewrites.get thy
+ val new_rews = insert Thm.eq_thm th1 current_rews
+ val updated_thy = Importer_Rewrites.put new_rews thy
+ in
+ (Context.Theory updated_thy,th1)
+ end
+ | add_importer_rewrite (context, th) = (context,
+ (th RS @{thm eq_reflection} handle THM _ => th)
+ );
+
+fun ignore_importer bthy bthm thy =
+ let
+ val _ = message ("Ignoring " ^ bthy ^ "." ^ bthm)
+ val curmaps = Importer_Maps.get thy
+ val newmaps = StringPair.update_new ((bthy,bthm),NONE) curmaps
+ val upd_thy = Importer_Maps.put newmaps thy
+ in
+ upd_thy
+ end
+
+val opt_get_output_thy = #2 o Importer_Dump.get
+
+fun get_output_thy thy =
+ case #2 (Importer_Dump.get thy) of
+ "" => error "No theory file being output"
+ | thyname => thyname
+
+val get_output_dir = #1 o Importer_Dump.get
+
+fun add_importer_move bef aft thy =
+ let
+ val curmoves = Importer_Moves.get thy
+ val newmoves = Symtab.update_new (bef, aft) curmoves
+ in
+ Importer_Moves.put newmoves thy
+ end
+
+fun get_importer_move bef thy =
+ Symtab.lookup (Importer_Moves.get thy) bef
+
+fun follow_name thmname thy =
+ let
+ val moves = Importer_Moves.get thy
+ fun F thmname =
+ case Symtab.lookup moves thmname of
+ SOME name => F name
+ | NONE => thmname
+ in
+ F thmname
+ end
+
+fun add_importer_cmove bef aft thy =
+ let
+ val curmoves = Importer_CMoves.get thy
+ val newmoves = Symtab.update_new (bef, aft) curmoves
+ in
+ Importer_CMoves.put newmoves thy
+ end
+
+fun get_importer_cmove bef thy =
+ Symtab.lookup (Importer_CMoves.get thy) bef
+
+fun follow_cname thmname thy =
+ let
+ val moves = Importer_CMoves.get thy
+ fun F thmname =
+ case Symtab.lookup moves thmname of
+ SOME name => F name
+ | NONE => thmname
+ in
+ F thmname
+ end
+
+fun add_importer_mapping bthy bthm isathm thy =
+ let
+ (* val _ = writeln ("Before follow_name: "^isathm) *)
+ val isathm = follow_name isathm thy
+ (* val _ = writeln ("Adding theorem map: " ^ bthy ^ "." ^ bthm ^ " --> " ^ isathm)*)
+ val curmaps = Importer_Maps.get thy
+ val newmaps = StringPair.update_new ((bthy,bthm),SOME isathm) curmaps
+ val upd_thy = Importer_Maps.put newmaps thy
+ in
+ upd_thy
+ end;
+
+fun get_importer_type_mapping bthy tycon thy =
+ let
+ val maps = Importer_TypeMaps.get thy
+ in
+ StringPair.lookup maps (bthy,tycon)
+ end
+
+fun get_importer_mapping bthy bthm thy =
+ let
+ val curmaps = Importer_Maps.get thy
+ in
+ StringPair.lookup curmaps (bthy,bthm)
+ end;
+
+fun add_importer_const_mapping bthy bconst internal isaconst thy =
+ let
+ val thy = case opt_get_output_thy thy of
+ "" => thy
+ | output_thy => if internal
+ then add_importer_cmove (Sign.full_bname thy bconst) (output_thy ^ "." ^ bthy ^ "." ^ bconst) thy
+ else thy
+ val _ = message ("Adding cmap " ^ bthy ^ "." ^ bconst ^ " -> " ^ isaconst ^ (if internal then " (*)" else ""))
+ val curmaps = Importer_ConstMaps.get thy
+ val newmaps = StringPair.update_new ((bthy,bconst),(internal,isaconst,NONE)) curmaps
+ val upd_thy = Importer_ConstMaps.put newmaps thy
+ in
+ upd_thy
+ end;
+
+fun add_importer_const_renaming bthy bconst newname thy =
+ let
+ val currens = Importer_Rename.get thy
+ val _ = message ("Adding renaming " ^ bthy ^ "." ^ bconst ^ " -> " ^ newname)
+ val newrens = StringPair.update_new ((bthy,bconst),newname) currens
+ val upd_thy = Importer_Rename.put newrens thy
+ in
+ upd_thy
+ end;
+
+fun get_importer_const_renaming bthy bconst thy =
+ let
+ val currens = Importer_Rename.get thy
+ in
+ StringPair.lookup currens (bthy,bconst)
+ end;
+
+fun get_importer_const_mapping bthy bconst thy =
+ let
+ val bconst = case get_importer_const_renaming bthy bconst thy of
+ SOME name => name
+ | NONE => bconst
+ val maps = Importer_ConstMaps.get thy
+ in
+ StringPair.lookup maps (bthy,bconst)
+ end
+
+fun add_importer_const_wt_mapping bthy bconst internal isaconst typ thy =
+ let
+ val thy = case opt_get_output_thy thy of
+ "" => thy
+ | output_thy => if internal
+ then add_importer_cmove (Sign.full_bname thy bconst) (output_thy ^ "." ^ bthy ^ "." ^ bconst) thy
+ else thy
+ val _ = message ("Adding cmap " ^ bthy ^ "." ^ bconst ^ " -> " ^ isaconst ^ (if internal then " (*)" else ""))
+ val curmaps = Importer_ConstMaps.get thy
+ val newmaps = StringPair.update_new ((bthy,bconst),(internal,isaconst,SOME typ)) curmaps
+ val upd_thy = Importer_ConstMaps.put newmaps thy
+ in
+ upd_thy
+ end;
+
+fun add_importer_type_mapping bthy bconst internal isaconst thy =
+ let
+ val curmaps = Importer_TypeMaps.get thy
+ val _ = writeln ("Adding tmap " ^ bthy ^ "." ^ bconst ^ " -> " ^ isaconst ^ (if internal then " (*)" else ""))
+ val newmaps = StringPair.update_new ((bthy,bconst),(internal,isaconst)) curmaps
+ (* FIXME avoid handle x *)
+ handle x => let val (_, isaconst') = the (StringPair.lookup curmaps (bthy, bconst)) in
+ warning ("couldn't map type "^bthy^"."^bconst^" to "^isaconst^": already mapped to "^isaconst'); raise x end
+ val upd_thy = Importer_TypeMaps.put newmaps thy
+ in
+ upd_thy
+ end;
+
+fun add_importer_pending bthy bthm hth thy =
+ let
+ val thmname = Sign.full_bname thy bthm
+ val _ = message ("Add pending " ^ bthy ^ "." ^ bthm)
+ val curpend = Importer_Pending.get thy
+ val newpend = StringPair.update_new ((bthy,bthm),hth) curpend
+ val upd_thy = Importer_Pending.put newpend thy
+ val thy' = case opt_get_output_thy upd_thy of
+ "" => add_importer_mapping bthy bthm thmname upd_thy
+ | output_thy =>
+ let
+ val new_thmname = output_thy ^ "." ^ bthy ^ "." ^ bthm
+ in
+ upd_thy |> add_importer_move thmname new_thmname
+ |> add_importer_mapping bthy bthm new_thmname
+ end
+ in
+ thy'
+ end;
+
+fun get_importer_theorem thyname thmname thy =
+ let
+ val isathms = Importer_Thms.get thy
+ in
+ StringPair.lookup isathms (thyname,thmname)
+ end;
+
+fun add_importer_theorem thyname thmname hth =
+ let
+ val _ = message ("Adding external theorem " ^ thyname ^ "." ^ thmname)
+ in
+ Importer_Thms.map (StringPair.update_new ((thyname, thmname), hth))
+ end;
+
+fun export_importer_pending thy =
+ let
+ val rews = Importer_Rewrites.get thy;
+ val pending = Importer_Pending.get thy;
+ fun process ((bthy,bthm), hth as (_,thm)) thy =
+ let
+ val thm1 = rewrite_rule (map (Thm.transfer thy) rews) (Thm.transfer thy thm);
+ val thm2 = Drule.export_without_context thm1;
+ in
+ thy
+ |> Global_Theory.store_thm (Binding.name bthm, thm2)
+ |> snd
+ |> add_importer_theorem bthy bthm hth
+ end;
+ in
+ thy
+ |> StringPair.fold process pending
+ |> Importer_Pending.put StringPair.empty
+ end;
+
+fun setup_dump (dir,thyname) thy =
+ Importer_Dump.put (dir,thyname,["(* AUTOMATICALLY GENERATED, DO NOT EDIT! *)"]) thy
+
+fun add_dump str thy =
+ let
+ val (dir,thyname,curdump) = Importer_Dump.get thy
+ in
+ Importer_Dump.put (dir,thyname,str::curdump) thy
+ end
+
+fun flush_dump thy =
+ let
+ val (dir,thyname,dumpdata) = Importer_Dump.get thy
+ val os = TextIO.openOut (OS.Path.joinDirFile {dir=dir,
+ file=thyname ^ ".thy"})
+ val _ = app (fn s => TextIO.output(os,s ^ "\n\n")) (rev dumpdata)
+ val _ = TextIO.closeOut os
+ in
+ Importer_Dump.put ("","",[]) thy
+ end
+
+fun set_generating_thy thyname thy =
+ case Importer_DefThy.get thy of
+ NoImport => Importer_DefThy.put (Generating thyname) thy
+ | _ => error "Import already in progess"
+
+fun set_replaying_thy thyname thy =
+ case Importer_DefThy.get thy of
+ NoImport => Importer_DefThy.put (Replaying thyname) thy
+ | _ => error "Import already in progess"
+
+fun clear_import_thy thy =
+ case Importer_DefThy.get thy of
+ NoImport => error "No import in progress"
+ | _ => Importer_DefThy.put NoImport thy
+
+fun get_generating_thy thy =
+ case Importer_DefThy.get thy of
+ Generating thyname => thyname
+ | _ => error "No theory being generated"
+
+fun get_replaying_thy thy =
+ case Importer_DefThy.get thy of
+ Replaying thyname => thyname
+ | _ => error "No theory being replayed"
+
+fun get_import_thy thy =
+ case Importer_DefThy.get thy of
+ Replaying thyname => thyname
+ | Generating thyname => thyname
+ | _ => error "No theory being imported"
+
+fun should_ignore thyname thy thmname =
+ case get_importer_mapping thyname thmname thy of
+ SOME NONE => true
+ | _ => false
+
+val trans_string =
+ let
+ fun quote s = "\"" ^ s ^ "\""
+ fun F [] = []
+ | F (#"\\" :: cs) = patch #"\\" cs
+ | F (#"\"" :: cs) = patch #"\"" cs
+ | F (c :: cs) = c :: F cs
+ and patch c rest = #"\\" :: c :: F rest
+ in
+ quote o String.implode o F o String.explode
+ end
+
+fun dump_import_thy thyname thy =
+ let
+ val output_dir = get_output_dir thy
+ val output_thy = get_output_thy thy
+ val input_thy = Context.theory_name thy
+ val import_segment = get_import_segment thy
+ val os = TextIO.openOut (OS.Path.joinDirFile {dir=output_dir,
+ file=thyname ^ ".imp"})
+ fun out s = TextIO.output(os,s)
+ val (ignored, mapped) = StringPair.fold
+ (fn ((bthy, bthm), v) => fn (ign, map) =>
+ if bthy = thyname
+ then case v
+ of NONE => (bthm :: ign, map)
+ | SOME w => (ign, (bthm, w) :: map)
+ else (ign, map)) (Importer_Maps.get thy) ([],[]);
+ fun mk init = StringPair.fold
+ (fn ((bthy, bthm), v) => if bthy = thyname then cons (bthm, v) else I) init [];
+ val constmaps = mk (Importer_ConstMaps.get thy);
+ val constrenames = mk (Importer_Rename.get thy);
+ val typemaps = mk (Importer_TypeMaps.get thy);
+ val defmaps = mk (Importer_DefMaps.get thy);
+ fun new_name internal isa =
+ if internal
+ then
+ let
+ val paths = Long_Name.explode isa
+ val i = drop (length paths - 2) paths
+ in
+ case i of
+ [seg,con] => output_thy ^ "." ^ seg ^ "." ^ con
+ | _ => error "import_rews.dump internal error"
+ end
+ else
+ isa
+
+ val _ = out "import\n\n"
+
+ val _ = out ("import_segment " ^ trans_string import_segment ^ "\n\n")
+ val _ = if null defmaps
+ then ()
+ else out "def_maps"
+ val _ = app (fn (hol,isa) =>
+ out ("\n " ^ (trans_string hol) ^ " > " ^ (trans_string isa))) defmaps
+ val _ = if null defmaps
+ then ()
+ else out "\n\n"
+
+ val _ = if null typemaps
+ then ()
+ else out "type_maps"
+ val _ = app (fn (hol,(internal,isa)) =>
+ out ("\n " ^ (trans_string hol) ^ " > " ^ (trans_string (new_name internal isa)))) typemaps
+ val _ = if null typemaps
+ then ()
+ else out "\n\n"
+
+ val _ = if null constmaps
+ then ()
+ else out "const_maps"
+ val _ = app (fn (hol,(_,isa,opt_ty)) =>
+ (out ("\n " ^ (trans_string hol) ^ " > " ^ (trans_string (follow_cname isa thy)));
+ case opt_ty of
+ SOME ty => out (" :: \"" ^ Syntax.string_of_typ_global thy ty ^ "\"")
+ | NONE => ())) constmaps
+ val _ = if null constmaps
+ then ()
+ else out "\n\n"
+
+ val _ = if null constrenames
+ then ()
+ else out "const_renames"
+ val _ = app (fn (old,new) =>
+ out ("\n " ^ (trans_string old) ^ " > " ^ (trans_string new))) constrenames
+ val _ = if null constrenames
+ then ()
+ else out "\n\n"
+
+ fun gen2replay in_thy out_thy s =
+ let
+ val ss = Long_Name.explode s
+ in
+ if (hd ss = in_thy) then
+ Long_Name.implode (out_thy::(tl ss))
+ else
+ s
+ end
+
+ val _ = if null mapped
+ then ()
+ else out "thm_maps"
+ val _ = app (fn (hol,isa) => out ("\n " ^ (trans_string hol) ^ " > " ^ (trans_string (gen2replay input_thy output_thy isa)))) mapped
+ val _ = if null mapped
+ then ()
+ else out "\n\n"
+
+ val _ = if null ignored
+ then ()
+ else out "ignore_thms"
+ val _ = app (fn ign => out ("\n " ^ (trans_string ign))) ignored
+ val _ = if null ignored
+ then ()
+ else out "\n\n"
+
+ val _ = out "end\n"
+ val _ = TextIO.closeOut os
+ in
+ thy
+ end
+
+fun set_used_names names thy =
+ let
+ val unames = Importer_UNames.get thy
+ in
+ case unames of
+ [] => Importer_UNames.put names thy
+ | _ => error "import_rews.set_used_names called on initialized data!"
+ end
+
+val clear_used_names = Importer_UNames.put [];
+
+fun get_defmap thyname const thy =
+ let
+ val maps = Importer_DefMaps.get thy
+ in
+ StringPair.lookup maps (thyname,const)
+ end
+
+fun add_defmap thyname const defname thy =
+ let
+ val _ = message ("Adding defmap " ^ thyname ^ "." ^ const ^ " --> " ^ defname)
+ val maps = Importer_DefMaps.get thy
+ val maps' = StringPair.update_new ((thyname,const),defname) maps
+ val thy' = Importer_DefMaps.put maps' thy
+ in
+ thy'
+ end
+
+fun get_defname thyname name thy =
+ let
+ val maps = Importer_DefMaps.get thy
+ fun F dname = (dname,add_defmap thyname name dname thy)
+ in
+ case StringPair.lookup maps (thyname,name) of
+ SOME thmname => (thmname,thy)
+ | NONE =>
+ let
+ val used = Importer_UNames.get thy
+ val defname = Thm.def_name name
+ val pdefname = name ^ "_primdef"
+ in
+ if not (member (op =) used defname)
+ then F defname (* name_def *)
+ else if not (member (op =) used pdefname)
+ then F pdefname (* name_primdef *)
+ else F (singleton (Name.variant_list used) pdefname) (* last resort *)
+ end
+ end
+
+local
+ fun handle_meta [x as Ast.Appl [Ast.Appl [Ast.Constant "_constrain", Ast.Constant @{const_syntax "=="}, _],_,_]] = x
+ | handle_meta [x as Ast.Appl [Ast.Appl [Ast.Constant "_constrain", Ast.Constant @{const_syntax all}, _],_]] = x
+ | handle_meta [x as Ast.Appl [Ast.Appl [Ast.Constant "_constrain", Ast.Constant @{const_syntax "==>"}, _],_,_]] = x
+ | handle_meta [x] = Ast.Appl [Ast.Constant @{const_syntax Trueprop}, x]
+ | handle_meta _ = error "import_rews error: Trueprop not applied to single argument"
+in
+val smarter_trueprop_parsing = [(@{const_syntax Trueprop},handle_meta)]
+end
+
+val importer_setup =
+ add_importer_type_mapping "min" "bool" false @{type_name bool}
+ #> add_importer_type_mapping "min" "fun" false "fun"
+ #> add_importer_type_mapping "min" "ind" false @{type_name ind}
+ #> add_importer_const_mapping "min" "=" false @{const_name HOL.eq}
+ #> add_importer_const_mapping "min" "==>" false @{const_name HOL.implies}
+ #> add_importer_const_mapping "min" "@" false @{const_name "Eps"}
+ #> Attrib.setup @{binding import_rew}
+ (Scan.succeed (Thm.mixed_attribute add_importer_rewrite)) "external rewrite rule";
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL4/proof_kernel.ML Sun Apr 01 14:50:47 2012 +0200
@@ -0,0 +1,2084 @@
+(* Title: HOL/Import/proof_kernel.ML
+ Author: Sebastian Skalberg and Steven Obua, TU Muenchen
+*)
+
+signature ProofKernel =
+sig
+ type hol_type
+ type tag
+ type term
+ type thm
+ type ('a,'b) subst
+
+ type proof_info
+ datatype proof = Proof of proof_info * proof_content
+ and proof_content
+ = PRefl of term
+ | PInstT of proof * (hol_type,hol_type) subst
+ | PSubst of proof list * term * proof
+ | PAbs of proof * term
+ | PDisch of proof * term
+ | PMp of proof * proof
+ | PHyp of term
+ | PAxm of string * term
+ | PDef of string * string * term
+ | PTmSpec of string * string list * proof
+ | PTyDef of string * string * proof
+ | PTyIntro of string * string * string * string * term * term * proof
+ | POracle of tag * term list * term
+ | PDisk
+ | PSpec of proof * term
+ | PInst of proof * (term,term) subst
+ | PGen of proof * term
+ | PGenAbs of proof * term option * term list
+ | PImpAS of proof * proof
+ | PSym of proof
+ | PTrans of proof * proof
+ | PComb of proof * proof
+ | PEqMp of proof * proof
+ | PEqImp of proof
+ | PExists of proof * term * term
+ | PChoose of term * proof * proof
+ | PConj of proof * proof
+ | PConjunct1 of proof
+ | PConjunct2 of proof
+ | PDisj1 of proof * term
+ | PDisj2 of proof * term
+ | PDisjCases of proof * proof * proof
+ | PNotI of proof
+ | PNotE of proof
+ | PContr of proof * term
+
+ exception PK of string * string
+
+ val get_proof_dir: string -> theory -> string option
+ val disambiguate_frees : Thm.thm -> Thm.thm
+ val debug : bool Unsynchronized.ref
+ val disk_info_of : proof -> (string * string) option
+ val set_disk_info_of : proof -> string -> string -> unit
+ val mk_proof : proof_content -> proof
+ val content_of : proof -> proof_content
+ val import_proof : string -> string -> theory -> (theory -> term) option * (theory -> proof)
+
+ val rewrite_importer_term: Term.term -> theory -> Thm.thm
+
+ val type_of : term -> hol_type
+
+ val get_thm : string -> string -> theory -> (theory * thm option)
+ val get_def : string -> string -> term -> theory -> (theory * thm option)
+ val get_axiom: string -> string -> theory -> (theory * thm option)
+
+ val store_thm : string -> string -> thm -> theory -> theory * thm
+
+ val to_isa_thm : thm -> (term * term) list * Thm.thm
+ val to_isa_term: term -> Term.term
+ val to_hol_thm : Thm.thm -> thm
+
+ val REFL : term -> theory -> theory * thm
+ val ASSUME : term -> theory -> theory * thm
+ val INST_TYPE : (hol_type,hol_type) subst -> thm -> theory -> theory * thm
+ val INST : (term,term)subst -> thm -> theory -> theory * thm
+ val EQ_MP : thm -> thm -> theory -> theory * thm
+ val EQ_IMP_RULE : thm -> theory -> theory * thm
+ val SUBST : thm list -> term -> thm -> theory -> theory * thm
+ val DISJ_CASES : thm -> thm -> thm -> theory -> theory * thm
+ val DISJ1: thm -> term -> theory -> theory * thm
+ val DISJ2: term -> thm -> theory -> theory * thm
+ val IMP_ANTISYM: thm -> thm -> theory -> theory * thm
+ val SYM : thm -> theory -> theory * thm
+ val MP : thm -> thm -> theory -> theory * thm
+ val GEN : term -> thm -> theory -> theory * thm
+ val CHOOSE : term -> thm -> thm -> theory -> theory * thm
+ val EXISTS : term -> term -> thm -> theory -> theory * thm
+ val ABS : term -> thm -> theory -> theory * thm
+ val GEN_ABS : term option -> term list -> thm -> theory -> theory * thm
+ val TRANS : thm -> thm -> theory -> theory * thm
+ val CCONTR : term -> thm -> theory -> theory * thm
+ val CONJ : thm -> thm -> theory -> theory * thm
+ val CONJUNCT1: thm -> theory -> theory * thm
+ val CONJUNCT2: thm -> theory -> theory * thm
+ val NOT_INTRO: thm -> theory -> theory * thm
+ val NOT_ELIM : thm -> theory -> theory * thm
+ val SPEC : term -> thm -> theory -> theory * thm
+ val COMB : thm -> thm -> theory -> theory * thm
+ val DISCH: term -> thm -> theory -> theory * thm
+
+ val type_introduction: string -> string -> string -> string -> string -> term * term -> thm -> theory -> theory * thm
+
+ val new_definition : string -> string -> term -> theory -> theory * thm
+ val new_specification : string -> string -> string list -> thm -> theory -> theory * thm
+ val new_type_definition : string -> string -> string -> thm -> theory -> theory * thm
+ val new_axiom : string -> term -> theory -> theory * thm
+
+ val prin : term -> unit
+ val protect_factname : string -> string
+ val replay_protect_varname : string -> string -> unit
+ val replay_add_dump : string -> theory -> theory
+end
+
+structure ProofKernel : ProofKernel =
+struct
+type hol_type = Term.typ
+type term = Term.term
+datatype tag = Tag of string list
+type ('a,'b) subst = ('a * 'b) list
+datatype thm = HOLThm of (Term.term * Term.term) list * Thm.thm
+
+fun hthm2thm (HOLThm (_, th)) = th
+
+fun to_hol_thm th = HOLThm ([], th)
+
+val replay_add_dump = add_dump
+fun add_dump s thy = replay_add_dump s thy
+
+datatype proof_info
+ = Info of {disk_info: (string * string) option Unsynchronized.ref}
+
+datatype proof = Proof of proof_info * proof_content
+ and proof_content
+ = PRefl of term
+ | PInstT of proof * (hol_type,hol_type) subst
+ | PSubst of proof list * term * proof
+ | PAbs of proof * term
+ | PDisch of proof * term
+ | PMp of proof * proof
+ | PHyp of term
+ | PAxm of string * term
+ | PDef of string * string * term
+ | PTmSpec of string * string list * proof
+ | PTyDef of string * string * proof
+ | PTyIntro of string * string * string * string * term * term * proof
+ | POracle of tag * term list * term
+ | PDisk
+ | PSpec of proof * term
+ | PInst of proof * (term,term) subst
+ | PGen of proof * term
+ | PGenAbs of proof * term option * term list
+ | PImpAS of proof * proof
+ | PSym of proof
+ | PTrans of proof * proof
+ | PComb of proof * proof
+ | PEqMp of proof * proof
+ | PEqImp of proof
+ | PExists of proof * term * term
+ | PChoose of term * proof * proof
+ | PConj of proof * proof
+ | PConjunct1 of proof
+ | PConjunct2 of proof
+ | PDisj1 of proof * term
+ | PDisj2 of proof * term
+ | PDisjCases of proof * proof * proof
+ | PNotI of proof
+ | PNotE of proof
+ | PContr of proof * term
+
+exception PK of string * string
+fun ERR f mesg = PK (f,mesg)
+
+
+(* Compatibility. *)
+
+val string_of_mixfix = Pretty.string_of o Mixfix.pretty_mixfix;
+
+fun mk_syn thy c =
+ if Lexicon.is_identifier c andalso not (Syntax.is_keyword (Sign.syn_of thy) c) then NoSyn
+ else Delimfix (Syntax_Ext.escape c)
+
+fun quotename c =
+ if Lexicon.is_identifier c andalso not (Keyword.is_keyword c) then c else quote c
+
+exception SMART_STRING
+
+fun no_vars context tm =
+ let
+ val ctxt = Variable.set_body false context;
+ val ([tm'], _) = Variable.import_terms true [tm] ctxt;
+ in tm' end
+
+fun smart_string_of_cterm ctxt0 ct =
+ let
+ val ctxt = ctxt0
+ |> Config.put show_brackets false
+ |> Config.put show_all_types false
+ |> Config.put show_types false
+ |> Config.put show_sorts false;
+ val {t,T,...} = rep_cterm ct
+ (* Hack to avoid parse errors with Trueprop *)
+ val t' = HOLogic.dest_Trueprop t
+ handle TERM _ => t
+ val tn = no_vars ctxt t'
+ fun match u =
+ let val _ = Thm.match (ct, cterm_of (Proof_Context.theory_of ctxt) u) in true end
+ handle Pattern.MATCH => false
+ fun G 0 f ctxt x = f ctxt x
+ | G 1 f ctxt x = f (Config.put show_types true ctxt) x
+ | G 2 f ctxt x = G 1 f (Config.put show_sorts true ctxt) x
+ | G 3 f ctxt x = G 2 f (Config.put show_all_types true ctxt) x
+ | G 4 f ctxt x = G 3 f (Config.put show_brackets true ctxt) x
+ | G _ _ _ _ = raise SMART_STRING
+ fun F n =
+ let
+ val str = G n Syntax.string_of_term ctxt tn
+ val _ = warning (@{make_string} (n, str))
+ val u = Syntax.parse_term ctxt str
+ val u = if t = t' then u else HOLogic.mk_Trueprop u
+ val u = Syntax.check_term ctxt (Type.constraint T u)
+ in
+ if match u
+ then quote str
+ else F (n+1)
+ end
+ handle ERROR _ => F (n + 1)
+ | SMART_STRING =>
+ let val _ =
+ warning ("smart_string failed for: "^ G 0 Syntax.string_of_term ctxt (term_of ct))
+ in quote (G 2 Syntax.string_of_term ctxt tn) end
+ in
+ Print_Mode.setmp [] F 0
+ end
+
+fun smart_string_of_thm ctxt = smart_string_of_cterm ctxt o cprop_of
+
+fun prth th = writeln (Print_Mode.setmp [] Display.string_of_thm_without_context th);
+val topctxt = ML_Context.the_local_context ();
+fun prin t = writeln (Print_Mode.setmp []
+ (fn () => Syntax.string_of_term topctxt t) ());
+fun pth (HOLThm(_,thm)) =
+ let
+ (*val _ = writeln "Renaming:"
+ val _ = app (fn(v,w) => (prin v; writeln " -->"; prin w)) ren*)
+ val _ = prth thm
+ in
+ ()
+ end
+
+fun disk_info_of (Proof(Info{disk_info,...},_)) = !disk_info
+fun mk_proof p = Proof(Info{disk_info = Unsynchronized.ref NONE},p)
+fun content_of (Proof(_,p)) = p
+
+fun set_disk_info_of (Proof(Info{disk_info,...},_)) thyname thmname =
+ disk_info := SOME(thyname,thmname)
+
+structure Lib =
+struct
+
+fun assoc x =
+ let
+ fun F [] = raise PK("Lib.assoc","Not found")
+ | F ((x',y)::rest) = if x = x'
+ then y
+ else F rest
+ in
+ F
+ end
+infix mem;
+fun i mem L =
+ let fun itr [] = false
+ | itr (a::rst) = i=a orelse itr rst
+ in itr L end;
+
+infix union;
+fun [] union S = S
+ | S union [] = S
+ | (a::rst) union S2 = rst union (insert (op =) a S2);
+
+fun implode_subst [] = []
+ | implode_subst (x::r::rest) = ((x,r)::(implode_subst rest))
+ | implode_subst _ = raise ERR "implode_subst" "malformed substitution list"
+
+end
+open Lib
+
+structure Tag =
+struct
+val empty_tag = Tag []
+fun read name = Tag [name]
+fun merge (Tag tag1) (Tag tag2) = Tag (Lib.union(tag1,tag2))
+end
+
+(* Actual code. *)
+
+fun get_segment thyname l = (Lib.assoc "s" l
+ handle PK _ => thyname)
+val get_name : (string * string) list -> string = Lib.assoc "n"
+
+exception XML of string
+
+datatype xml = Elem of string * (string * string) list * xml list
+datatype XMLtype = XMLty of xml | FullType of hol_type
+datatype XMLterm = XMLtm of xml | FullTerm of term
+
+fun xml_to_import_xml (XML.Elem ((n, l), ts)) = Elem (n, l, map xml_to_import_xml ts)
+ | xml_to_import_xml (XML.Text _) = raise XML "Incorrect proof file: text";
+
+val type_of = Term.type_of
+
+val propT = Type("prop",[])
+
+fun mk_defeq name rhs thy =
+ let
+ val ty = type_of rhs
+ in
+ Logic.mk_equals (Const(Sign.intern_const thy name,ty),rhs)
+ end
+
+fun mk_teq name rhs thy =
+ let
+ val ty = type_of rhs
+ in
+ HOLogic.mk_eq (Const(Sign.intern_const thy name,ty),rhs)
+ end
+
+fun intern_const_name thyname const thy =
+ case get_importer_const_mapping thyname const thy of
+ SOME (_,cname,_) => cname
+ | NONE => (case get_importer_const_renaming thyname const thy of
+ SOME cname => Sign.intern_const thy (thyname ^ "." ^ cname)
+ | NONE => Sign.intern_const thy (thyname ^ "." ^ const))
+
+fun intern_type_name thyname const thy =
+ case get_importer_type_mapping thyname const thy of
+ SOME (_,cname) => cname
+ | NONE => Sign.intern_const thy (thyname ^ "." ^ const)
+
+fun mk_vartype name = TFree(name,["HOL.type"])
+fun mk_thy_type thy Thy Tyop Args = Type(intern_type_name Thy Tyop thy,Args)
+
+val mk_var = Free
+
+fun mk_thy_const thy Thy Nam Ty = Const(intern_const_name Thy Nam thy,Ty)
+
+local
+ fun get_const sg _ name =
+ (case Sign.const_type sg name of
+ SOME ty => Const (name, ty)
+ | NONE => raise ERR "get_type" (name ^ ": No such constant"))
+in
+fun prim_mk_const thy Thy Nam =
+ let
+ val name = intern_const_name Thy Nam thy
+ val cmaps = Importer_ConstMaps.get thy
+ in
+ case StringPair.lookup cmaps (Thy,Nam) of
+ SOME(_,_,SOME ty) => Const(name,ty)
+ | _ => get_const thy Thy name
+ end
+end
+
+fun mk_comb(f,a) = f $ a
+
+(* Needed for HOL Light *)
+fun protect_tyvarname s =
+ let
+ fun no_quest s =
+ if Char.contains s #"?"
+ then String.translate (fn #"?" => "q_" | c => Char.toString c) s
+ else s
+ fun beg_prime s =
+ if String.isPrefix "'" s
+ then s
+ else "'" ^ s
+ in
+ s |> no_quest |> beg_prime
+ end
+
+val protected_varnames = Unsynchronized.ref (Symtab.empty:string Symtab.table)
+val invented_isavar = Unsynchronized.ref 0
+
+fun innocent_varname s = Lexicon.is_identifier s andalso not (String.isPrefix "u_" s)
+
+fun valid_boundvarname s =
+ can (fn () => Syntax.read_term_global @{theory Main} ("SOME "^s^". True")) ();
+
+fun valid_varname s =
+ can (fn () => Syntax.read_term_global @{theory Main} s) ();
+
+fun protect_varname s =
+ if innocent_varname s andalso valid_varname s then s else
+ case Symtab.lookup (!protected_varnames) s of
+ SOME t => t
+ | NONE =>
+ let
+ val _ = Unsynchronized.inc invented_isavar
+ val t = "u_" ^ string_of_int (!invented_isavar)
+ val _ = protected_varnames := Symtab.update (s, t) (!protected_varnames)
+ in
+ t
+ end
+
+exception REPLAY_PROTECT_VARNAME of string*string*string
+
+fun replay_protect_varname s t =
+ case Symtab.lookup (!protected_varnames) s of
+ SOME t' => raise REPLAY_PROTECT_VARNAME (s, t, t')
+ | NONE =>
+ let
+ val _ = Unsynchronized.inc invented_isavar
+ val t = "u_" ^ string_of_int (!invented_isavar)
+ val _ = protected_varnames := Symtab.update (s, t) (!protected_varnames)
+ in
+ ()
+ end
+
+fun protect_boundvarname s = if innocent_varname s andalso valid_boundvarname s then s else "u"
+
+fun mk_lambda (v as Free (x, T)) t = Abs (protect_boundvarname x, T, abstract_over (v, t))
+ | mk_lambda (v as Var ((x, _), T)) t = Abs (protect_boundvarname x, T, abstract_over (v, t))
+ | mk_lambda v t = raise TERM ("lambda", [v, t]);
+
+fun replacestr x y s =
+let
+ val xl = raw_explode x
+ val yl = raw_explode y
+ fun isprefix [] _ = true
+ | isprefix (x::xs) (y::ys) = if x = y then isprefix xs ys else false
+ | isprefix _ _ = false
+ fun isp s = isprefix xl s
+ fun chg s = yl@(List.drop (s, List.length xl))
+ fun r [] = []
+ | r (S as (s::ss)) = if isp S then r (chg S) else s::(r ss)
+in
+ implode(r (raw_explode s))
+end
+
+fun protect_factname s = replacestr "." "_dot_" s
+fun unprotect_factname s = replacestr "_dot_" "." s
+
+val ty_num_prefix = "N_"
+
+fun startsWithDigit s = Char.isDigit (hd (String.explode s))
+
+fun protect_tyname tyn =
+ let
+ val tyn' =
+ if String.isPrefix ty_num_prefix tyn then raise (ERR "protect_ty_name" ("type name '"^tyn^"' is reserved")) else
+ (if startsWithDigit tyn then ty_num_prefix^tyn else tyn)
+ in
+ tyn'
+ end
+
+fun protect_constname tcn = tcn
+ (* if tcn = ".." then "dotdot"
+ else if tcn = "==" then "eqeq"
+ else tcn*)
+
+structure TypeNet =
+struct
+
+fun get_type_from_index thy thyname types is =
+ case Int.fromString is of
+ SOME i => (case Array.sub(types,i) of
+ FullType ty => ty
+ | XMLty xty =>
+ let
+ val ty = get_type_from_xml thy thyname types xty
+ val _ = Array.update(types,i,FullType ty)
+ in
+ ty
+ end)
+ | NONE => raise ERR "get_type_from_index" "Bad index"
+and get_type_from_xml thy thyname types =
+ let
+ fun gtfx (Elem("tyi",[("i",iS)],[])) =
+ get_type_from_index thy thyname types iS
+ | gtfx (Elem("tyc",atts,[])) =
+ mk_thy_type thy
+ (get_segment thyname atts)
+ (protect_tyname (get_name atts))
+ []
+ | gtfx (Elem("tyv",[("n",s)],[])) = mk_vartype (protect_tyvarname s)
+ | gtfx (Elem("tya",[],(Elem("tyc",atts,[]))::tys)) =
+ mk_thy_type thy
+ (get_segment thyname atts)
+ (protect_tyname (get_name atts))
+ (map gtfx tys)
+ | gtfx _ = raise ERR "get_type" "Bad type"
+ in
+ gtfx
+ end
+
+fun input_types _ (Elem("tylist",[("i",i)],xtys)) =
+ let
+ val types = Array.array(the (Int.fromString i),XMLty (Elem("",[],[])))
+ fun IT _ [] = ()
+ | IT n (xty::xtys) =
+ (Array.update(types,n,XMLty xty);
+ IT (n+1) xtys)
+ val _ = IT 0 xtys
+ in
+ types
+ end
+ | input_types _ _ = raise ERR "input_types" "Bad type list"
+end
+
+structure TermNet =
+struct
+
+fun get_term_from_index thy thyname types terms is =
+ case Int.fromString is of
+ SOME i => (case Array.sub(terms,i) of
+ FullTerm tm => tm
+ | XMLtm xtm =>
+ let
+ val tm = get_term_from_xml thy thyname types terms xtm
+ val _ = Array.update(terms,i,FullTerm tm)
+ in
+ tm
+ end)
+ | NONE => raise ERR "get_term_from_index" "Bad index"
+and get_term_from_xml thy thyname types terms =
+ let
+ fun gtfx (Elem("tmv",[("n",name),("t",tyi)],[])) =
+ mk_var(protect_varname name,TypeNet.get_type_from_index thy thyname types tyi)
+ | gtfx (Elem("tmc",atts,[])) =
+ let
+ val segment = get_segment thyname atts
+ val name = protect_constname(get_name atts)
+ in
+ mk_thy_const thy segment name (TypeNet.get_type_from_index thy thyname types (Lib.assoc "t" atts))
+ handle PK _ => prim_mk_const thy segment name
+ end
+ | gtfx (Elem("tma",[("f",tmf),("a",tma)],[])) =
+ let
+ val f = get_term_from_index thy thyname types terms tmf
+ val a = get_term_from_index thy thyname types terms tma
+ in
+ mk_comb(f,a)
+ end
+ | gtfx (Elem("tml",[("x",tmx),("a",tma)],[])) =
+ let
+ val x = get_term_from_index thy thyname types terms tmx
+ val a = get_term_from_index thy thyname types terms tma
+ in
+ mk_lambda x a
+ end
+ | gtfx (Elem("tmi",[("i",iS)],[])) =
+ get_term_from_index thy thyname types terms iS
+ | gtfx (Elem(tag,_,_)) =
+ raise ERR "get_term" ("Not a term: "^tag)
+ in
+ gtfx
+ end
+
+fun input_terms _ _ (Elem("tmlist",[("i",i)],xtms)) =
+ let
+ val terms = Array.array(the (Int.fromString i), XMLtm (Elem("",[],[])))
+
+ fun IT _ [] = ()
+ | IT n (xtm::xtms) =
+ (Array.update(terms,n,XMLtm xtm);
+ IT (n+1) xtms)
+ val _ = IT 0 xtms
+ in
+ terms
+ end
+ | input_terms _ _ _ = raise ERR "input_terms" "Bad term list"
+end
+
+fun get_proof_dir (thyname:string) thy =
+ let
+ val import_segment =
+ case get_segment2 thyname thy of
+ SOME seg => seg
+ | NONE => get_import_segment thy
+ val path = space_explode ":" (getenv "IMPORTER_PROOFS")
+ fun find [] = NONE
+ | find (p::ps) =
+ (let
+ val dir = OS.Path.joinDirFile {dir = p,file=import_segment}
+ in
+ if OS.FileSys.isDir dir
+ then SOME dir
+ else find ps
+ end) handle OS.SysErr _ => find ps
+ in
+ Option.map (fn p => OS.Path.joinDirFile {dir = p, file = thyname}) (find path)
+ end
+
+fun proof_file_name thyname thmname thy =
+ let
+ val path = case get_proof_dir thyname thy of
+ SOME p => p
+ | NONE => error "Cannot find proof files"
+ val _ = OS.FileSys.mkDir path handle OS.SysErr _ => ()
+ in
+ OS.Path.joinDirFile {dir = path, file = OS.Path.joinBaseExt {base = (unprotect_factname thmname), ext = SOME "prf"}}
+ end
+
+fun xml_to_proof thyname types terms prf thy =
+ let
+ val xml_to_hol_type = TypeNet.get_type_from_xml thy thyname types
+ val xml_to_term = TermNet.get_term_from_xml thy thyname types terms
+
+ fun index_to_term is =
+ TermNet.get_term_from_index thy thyname types terms is
+
+ fun x2p (Elem("prefl",[("i",is)],[])) = mk_proof (PRefl (index_to_term is))
+ | x2p (Elem("pinstt",[],p::lambda)) =
+ let
+ val p = x2p p
+ val lambda = implode_subst (map xml_to_hol_type lambda)
+ in
+ mk_proof (PInstT(p,lambda))
+ end
+ | x2p (Elem("psubst",[("i",is)],prf::prfs)) =
+ let
+ val tm = index_to_term is
+ val prf = x2p prf
+ val prfs = map x2p prfs
+ in
+ mk_proof (PSubst(prfs,tm,prf))
+ end
+ | x2p (Elem("pabs",[("i",is)],[prf])) =
+ let
+ val p = x2p prf
+ val t = index_to_term is
+ in
+ mk_proof (PAbs (p,t))
+ end
+ | x2p (Elem("pdisch",[("i",is)],[prf])) =
+ let
+ val p = x2p prf
+ val t = index_to_term is
+ in
+ mk_proof (PDisch (p,t))
+ end
+ | x2p (Elem("pmp",[],[prf1,prf2])) =
+ let
+ val p1 = x2p prf1
+ val p2 = x2p prf2
+ in
+ mk_proof (PMp(p1,p2))
+ end
+ | x2p (Elem("phyp",[("i",is)],[])) = mk_proof (PHyp (index_to_term is))
+ | x2p (Elem("paxiom",[("n",n),("i",is)],[])) =
+ mk_proof (PAxm(n,index_to_term is))
+ | x2p (Elem("pfact",atts,[])) =
+ let
+ val thyname = get_segment thyname atts
+ val thmname = protect_factname (get_name atts)
+ val p = mk_proof PDisk
+ val _ = set_disk_info_of p thyname thmname
+ in
+ p
+ end
+ | x2p (Elem("pdef",[("s",seg),("n",name),("i",is)],[])) =
+ mk_proof (PDef(seg,protect_constname name,index_to_term is))
+ | x2p (Elem("ptmspec",[("s",seg)],p::names)) =
+ let
+ val names = map (fn Elem("name",[("n",name)],[]) => name
+ | _ => raise ERR "x2p" "Bad proof (ptmspec)") names
+ in
+ mk_proof (PTmSpec(seg,names,x2p p))
+ end
+ | x2p (Elem("ptyintro",[("s",seg),("n",name),("a",abs_name),("r",rep_name)],[xP,xt,p])) =
+ let
+ val P = xml_to_term xP
+ val t = xml_to_term xt
+ in
+ mk_proof (PTyIntro(seg,protect_tyname name,protect_constname abs_name,protect_constname rep_name,P,t,x2p p))
+ end
+ | x2p (Elem("ptydef",[("s",seg),("n",name)],[p])) =
+ mk_proof (PTyDef(seg,protect_tyname name,x2p p))
+ | x2p (Elem("poracle",[],chldr)) =
+ let
+ val (oracles,terms) = List.partition (fn (Elem("oracle",_,_)) => true | _ => false) chldr
+ val ors = map (fn (Elem("oracle",[("n",name)],[])) => name | _ => raise ERR "x2p" "bad oracle") oracles
+ val (c,asl) = case terms of
+ [] => raise ERR "x2p" "Bad oracle description"
+ | (hd::tl) => (hd,tl)
+ val tg = fold_rev (Tag.merge o Tag.read) ors Tag.empty_tag
+ in
+ mk_proof (POracle(tg,map xml_to_term asl,xml_to_term c))
+ end
+ | x2p (Elem("pspec",[("i",is)],[prf])) =
+ let
+ val p = x2p prf
+ val tm = index_to_term is
+ in
+ mk_proof (PSpec(p,tm))
+ end
+ | x2p (Elem("pinst",[],p::theta)) =
+ let
+ val p = x2p p
+ val theta = implode_subst (map xml_to_term theta)
+ in
+ mk_proof (PInst(p,theta))
+ end
+ | x2p (Elem("pgen",[("i",is)],[prf])) =
+ let
+ val p = x2p prf
+ val tm = index_to_term is
+ in
+ mk_proof (PGen(p,tm))
+ end
+ | x2p (Elem("pgenabs",[],prf::tms)) =
+ let
+ val p = x2p prf
+ val tml = map xml_to_term tms
+ in
+ mk_proof (PGenAbs(p,NONE,tml))
+ end
+ | x2p (Elem("pgenabs",[("i",is)],prf::tms)) =
+ let
+ val p = x2p prf
+ val tml = map xml_to_term tms
+ in
+ mk_proof (PGenAbs(p,SOME (index_to_term is),tml))
+ end
+ | x2p (Elem("pimpas",[],[prf1,prf2])) =
+ let
+ val p1 = x2p prf1
+ val p2 = x2p prf2
+ in
+ mk_proof (PImpAS(p1,p2))
+ end
+ | x2p (Elem("psym",[],[prf])) =
+ let
+ val p = x2p prf
+ in
+ mk_proof (PSym p)
+ end
+ | x2p (Elem("ptrans",[],[prf1,prf2])) =
+ let
+ val p1 = x2p prf1
+ val p2 = x2p prf2
+ in
+ mk_proof (PTrans(p1,p2))
+ end
+ | x2p (Elem("pcomb",[],[prf1,prf2])) =
+ let
+ val p1 = x2p prf1
+ val p2 = x2p prf2
+ in
+ mk_proof (PComb(p1,p2))
+ end
+ | x2p (Elem("peqmp",[],[prf1,prf2])) =
+ let
+ val p1 = x2p prf1
+ val p2 = x2p prf2
+ in
+ mk_proof (PEqMp(p1,p2))
+ end
+ | x2p (Elem("peqimp",[],[prf])) =
+ let
+ val p = x2p prf
+ in
+ mk_proof (PEqImp p)
+ end
+ | x2p (Elem("pexists",[("e",ise),("w",isw)],[prf])) =
+ let
+ val p = x2p prf
+ val ex = index_to_term ise
+ val w = index_to_term isw
+ in
+ mk_proof (PExists(p,ex,w))
+ end
+ | x2p (Elem("pchoose",[("i",is)],[prf1,prf2])) =
+ let
+ val v = index_to_term is
+ val p1 = x2p prf1
+ val p2 = x2p prf2
+ in
+ mk_proof (PChoose(v,p1,p2))
+ end
+ | x2p (Elem("pconj",[],[prf1,prf2])) =
+ let
+ val p1 = x2p prf1
+ val p2 = x2p prf2
+ in
+ mk_proof (PConj(p1,p2))
+ end
+ | x2p (Elem("pconjunct1",[],[prf])) =
+ let
+ val p = x2p prf
+ in
+ mk_proof (PConjunct1 p)
+ end
+ | x2p (Elem("pconjunct2",[],[prf])) =
+ let
+ val p = x2p prf
+ in
+ mk_proof (PConjunct2 p)
+ end
+ | x2p (Elem("pdisj1",[("i",is)],[prf])) =
+ let
+ val p = x2p prf
+ val t = index_to_term is
+ in
+ mk_proof (PDisj1 (p,t))
+ end
+ | x2p (Elem("pdisj2",[("i",is)],[prf])) =
+ let
+ val p = x2p prf
+ val t = index_to_term is
+ in
+ mk_proof (PDisj2 (p,t))
+ end
+ | x2p (Elem("pdisjcases",[],[prf1,prf2,prf3])) =
+ let
+ val p1 = x2p prf1
+ val p2 = x2p prf2
+ val p3 = x2p prf3
+ in
+ mk_proof (PDisjCases(p1,p2,p3))
+ end
+ | x2p (Elem("pnoti",[],[prf])) =
+ let
+ val p = x2p prf
+ in
+ mk_proof (PNotI p)
+ end
+ | x2p (Elem("pnote",[],[prf])) =
+ let
+ val p = x2p prf
+ in
+ mk_proof (PNotE p)
+ end
+ | x2p (Elem("pcontr",[("i",is)],[prf])) =
+ let
+ val p = x2p prf
+ val t = index_to_term is
+ in
+ mk_proof (PContr (p,t))
+ end
+ | x2p _ = raise ERR "x2p" "Bad proof"
+ in
+ x2p prf
+ end
+
+fun import_proof_concl thyname thmname thy =
+ let
+ val is = TextIO.openIn(proof_file_name thyname thmname thy)
+ val proof_xml = xml_to_import_xml (XML.parse (TextIO.inputAll is))
+ val _ = TextIO.closeIn is
+ in
+ case proof_xml of
+ Elem("proof",[],xtypes::xterms::_::rest) =>
+ let
+ val types = TypeNet.input_types thyname xtypes
+ val terms = TermNet.input_terms thyname types xterms
+ fun f xtm thy = TermNet.get_term_from_xml thy thyname types terms xtm
+ in
+ case rest of
+ [] => NONE
+ | [xtm] => SOME (f xtm)
+ | _ => raise ERR "import_proof" "Bad argument list"
+ end
+ | _ => raise ERR "import_proof" "Bad proof"
+ end
+
+fun import_proof thyname thmname thy =
+ let
+ val is = TextIO.openIn(proof_file_name thyname thmname thy)
+ val proof_xml = xml_to_import_xml (XML.parse (TextIO.inputAll is))
+ val _ = TextIO.closeIn is
+ in
+ case proof_xml of
+ Elem("proof",[],xtypes::xterms::prf::rest) =>
+ let
+ val types = TypeNet.input_types thyname xtypes
+ val terms = TermNet.input_terms thyname types xterms
+ in
+ (case rest of
+ [] => NONE
+ | [xtm] => SOME (fn thy => TermNet.get_term_from_xml thy thyname types terms xtm)
+ | _ => raise ERR "import_proof" "Bad argument list",
+ xml_to_proof thyname types terms prf)
+ end
+ | _ => raise ERR "import_proof" "Bad proof"
+ end
+
+fun uniq_compose m th i st =
+ let
+ val res = Thm.bicompose false (false,th,m) i st
+ in
+ case Seq.pull res of
+ SOME (th,rest) => (case Seq.pull rest of
+ SOME _ => raise ERR "uniq_compose" "Not unique!"
+ | NONE => th)
+ | NONE => raise ERR "uniq_compose" "No result"
+ end
+
+val reflexivity_thm = @{thm refl}
+val mp_thm = @{thm mp}
+val imp_antisym_thm = @{thm light_imp_as}
+val disch_thm = @{thm impI}
+val ccontr_thm = @{thm ccontr}
+
+val meta_eq_to_obj_eq_thm = @{thm meta_eq_to_obj_eq}
+
+val gen_thm = @{thm HOLallI}
+val choose_thm = @{thm exE}
+val exists_thm = @{thm exI}
+val conj_thm = @{thm conjI}
+val conjunct1_thm = @{thm conjunct1}
+val conjunct2_thm = @{thm conjunct2}
+val spec_thm = @{thm spec}
+val disj_cases_thm = @{thm disjE}
+val disj1_thm = @{thm disjI1}
+val disj2_thm = @{thm disjI2}
+
+local
+ val th = @{thm not_def}
+ val thy = theory_of_thm th
+ val pp = Thm.reflexive (cterm_of thy (Const(@{const_name Trueprop},HOLogic.boolT-->propT)))
+in
+val not_elim_thm = Thm.combination pp th
+end
+
+val not_intro_thm = Thm.symmetric not_elim_thm
+val abs_thm = @{thm ext}
+val trans_thm = @{thm trans}
+val symmetry_thm = @{thm sym}
+val eqmp_thm = @{thm iffD1}
+val eqimp_thm = @{thm Importer.eq_imp}
+val comb_thm = @{thm cong}
+
+(* Beta-eta normalizes a theorem (only the conclusion, not the *
+hypotheses!) *)
+
+fun beta_eta_thm th =
+ let
+ val th1 = Thm.equal_elim (Thm.beta_conversion true (cprop_of th)) th
+ val th2 = Thm.equal_elim (Thm.eta_conversion (cprop_of th1)) th1
+ in
+ th2
+ end
+
+fun implies_elim_all th =
+ Library.foldl (fn (th,p) => Thm.implies_elim th (Thm.assume p)) (th,cprems_of th)
+
+fun norm_hyps th =
+ th |> beta_eta_thm
+ |> implies_elim_all
+ |> implies_intr_hyps
+
+fun mk_GEN v th sg =
+ let
+ val c = HOLogic.dest_Trueprop (concl_of th)
+ val cv = cterm_of sg v
+ val lc = Term.lambda v c
+ val clc = Thm.cterm_of sg lc
+ val cvty = ctyp_of_term cv
+ val th1 = implies_elim_all th
+ val th2 = beta_eta_thm (Thm.forall_intr cv th1)
+ val th3 = th2 COMP (beta_eta_thm (Drule.instantiate' [SOME cvty] [SOME clc] gen_thm))
+ val c = prop_of th3
+ val vname = fst(dest_Free v)
+ val (cold,cnew) = case c of
+ tpc $ (Const(@{const_name All},_) $ Abs(oldname,_,_)) =>
+ (Abs(oldname,dummyT,Bound 0),Abs(vname,dummyT,Bound 0))
+ | tpc $ (Const(@{const_name All},_) $ _) => (tpc,tpc)
+ | _ => raise ERR "mk_GEN" "Unknown conclusion"
+ val th4 = Thm.rename_boundvars cold cnew th3
+ val res = implies_intr_hyps th4
+ in
+ res
+ end
+
+fun rearrange sg tm th =
+ let
+ val tm' = Envir.beta_eta_contract tm
+ fun find [] n = Thm.permute_prems 0 1 (Thm.implies_intr (Thm.cterm_of sg tm) th)
+ | find (p::ps) n = if tm' aconv (Envir.beta_eta_contract p)
+ then Thm.permute_prems n 1 th
+ else find ps (n+1)
+ in
+ find (prems_of th) 0
+ end
+
+fun zip (x::xs) (y::ys) = (x,y)::(zip xs ys)
+ | zip [] [] = []
+ | zip _ _ = raise ERR "zip" "arguments not of same length"
+
+fun mk_INST dom rng th =
+ th |> forall_intr_list dom
+ |> forall_elim_list rng
+
+(* Code for disambiguating variablenames (wrt. types) *)
+
+val disamb_info_empty = {vars=[],rens=[]}
+
+fun rens_of { vars = _, rens = rens } = rens
+
+fun disamb_term_from info tm = (info, tm)
+
+fun has_ren (HOLThm _) = false
+
+fun disamb_thm_from info (HOLThm (_,thm)) = (info, thm)
+
+fun disamb_terms_from info tms = (info, tms)
+
+fun disamb_thms_from info hthms = (info, map hthm2thm hthms)
+
+fun disamb_term tm = disamb_term_from disamb_info_empty tm
+fun disamb_thm thm = disamb_thm_from disamb_info_empty thm
+fun disamb_thms thms = disamb_thms_from disamb_info_empty thms
+
+fun norm_hthm _ (hth as HOLThm _) = hth
+
+(* End of disambiguating code *)
+
+fun disambiguate_frees thm =
+ let
+ fun ERR s = error ("Drule.disambiguate_frees: "^s)
+ val ct = cprop_of thm
+ val t = term_of ct
+ val thy = theory_of_cterm ct
+ val frees = Misc_Legacy.term_frees t
+ val freenames = Term.add_free_names t []
+ val is_old_name = member (op =) freenames
+ fun name_of (Free (n, _)) = n
+ | name_of _ = ERR "name_of"
+ fun new_name' bump map n =
+ let val n' = n^bump in
+ if is_old_name n' orelse Symtab.lookup map n' <> NONE then
+ new_name' (Symbol.bump_string bump) map n
+ else
+ n'
+ end
+ val new_name = new_name' "a"
+ fun replace_name n' (Free (_, t)) = Free (n', t)
+ | replace_name _ _ = ERR "replace_name"
+ (* map: old or fresh name -> old free,
+ invmap: old free which has fresh name assigned to it -> fresh name *)
+ fun dis v (mapping as (map,invmap)) =
+ let val n = name_of v in
+ case Symtab.lookup map n of
+ NONE => (Symtab.update (n, v) map, invmap)
+ | SOME v' =>
+ if v=v' then
+ mapping
+ else
+ let val n' = new_name map n in
+ (Symtab.update (n', v) map,
+ Termtab.update (v, n') invmap)
+ end
+ end
+ in
+ if (length freenames = length frees) then
+ thm
+ else
+ let
+ val (_, invmap) =
+ fold dis frees (Symtab.empty, Termtab.empty)
+ fun make_subst (oldfree, newname) (intros, elims) =
+ (cterm_of thy oldfree :: intros,
+ cterm_of thy (replace_name newname oldfree) :: elims)
+ val (intros, elims) = fold make_subst (Termtab.dest invmap) ([], [])
+ in
+ forall_elim_list elims (forall_intr_list intros thm)
+ end
+ end
+
+val debug = Unsynchronized.ref false
+
+fun if_debug f x = if !debug then f x else ()
+val message = if_debug writeln
+
+fun get_importer_thm thyname thmname thy =
+ case get_importer_theorem thyname thmname thy of
+ SOME hth => SOME (HOLThm hth)
+ | NONE =>
+ let
+ val pending = Importer_Pending.get thy
+ in
+ case StringPair.lookup pending (thyname,thmname) of
+ SOME hth => SOME (HOLThm hth)
+ | NONE => NONE
+ end
+
+fun non_trivial_term_consts t = fold_aterms
+ (fn Const (c, _) =>
+ if c = @{const_name Trueprop} orelse c = @{const_name All}
+ orelse c = @{const_name HOL.implies} orelse c = @{const_name HOL.conj} orelse c = @{const_name HOL.eq}
+ then I else insert (op =) c
+ | _ => I) t [];
+
+fun split_name str =
+ let
+ val sub = Substring.full str
+ val (f,idx) = apsnd Substring.string (Substring.splitr Char.isDigit sub)
+ val (newstr,u) = pairself Substring.string (Substring.splitr (fn c => c = #"_") f)
+ in
+ if not (idx = "") andalso u = "_"
+ then SOME (newstr, the (Int.fromString idx))
+ else NONE
+ end
+ handle _ => NONE (* FIXME avoid handle _ *)
+
+fun rewrite_importer_term t thy =
+ let
+ val import_rews1 = map (Thm.transfer thy) (Importer_Rewrites.get thy)
+ val importerss = Simplifier.global_context thy empty_ss addsimps import_rews1
+ in
+ Thm.transfer thy (Simplifier.full_rewrite importerss (cterm_of thy t))
+ end
+
+fun get_isabelle_thm thyname thmname importerconc thy =
+ let
+ val (info,importerconc') = disamb_term importerconc
+ val i2h_conc = Thm.symmetric (rewrite_importer_term (HOLogic.mk_Trueprop importerconc') thy)
+ val isaconc =
+ case concl_of i2h_conc of
+ Const("==",_) $ lhs $ _ => lhs
+ | _ => error "get_isabelle_thm" "Bad rewrite rule"
+ val _ = (message "Original conclusion:";
+ if_debug prin importerconc';
+ message "Modified conclusion:";
+ if_debug prin isaconc)
+
+ fun mk_res th = HOLThm (rens_of info, Thm.equal_elim i2h_conc th)
+ in
+ case get_importer_mapping thyname thmname thy of
+ SOME (SOME thmname) =>
+ let
+ val th1 = (SOME (Global_Theory.get_thm thy thmname)
+ handle ERROR _ =>
+ (case split_name thmname of
+ SOME (listname,idx) => (SOME (nth (Global_Theory.get_thms thy listname) (idx - 1))
+ handle _ => NONE) (* FIXME avoid handle _ *)
+ | NONE => NONE))
+ in
+ case th1 of
+ SOME th2 =>
+ (case Shuffler.set_prop thy isaconc [(thmname,th2)] of
+ SOME (_,th) => (message "YES";(thy, SOME (mk_res th)))
+ | NONE => (message "NO2";error "get_isabelle_thm" "Bad mapping"))
+ | NONE => (message "NO1";error "get_isabelle_thm" "Bad mapping")
+ end
+ | SOME NONE => error ("Trying to access ignored theorem " ^ thmname)
+ | NONE =>
+ let
+ val _ = (message "Looking for conclusion:";
+ if_debug prin isaconc)
+ val cs = non_trivial_term_consts isaconc;
+ val _ = (message "Looking for consts:";
+ message (commas cs))
+ val pot_thms = Shuffler.find_potential thy isaconc
+ val _ = message (string_of_int (length pot_thms) ^ " potential theorems")
+ in
+ case Shuffler.set_prop thy isaconc pot_thms of
+ SOME (isaname,th) =>
+ let
+ val hth as HOLThm args = mk_res th
+ val thy' = thy |> add_importer_theorem thyname thmname args
+ |> add_importer_mapping thyname thmname isaname
+ in
+ (thy',SOME hth)
+ end
+ | NONE => (thy,NONE)
+ end
+ end
+ handle e =>
+ if Exn.is_interrupt e then reraise e
+ else
+ (if_debug (fn () =>
+ writeln ("Exception in get_isabelle_thm:\n" ^ ML_Compiler.exn_message e)) ();
+ (thy,NONE))
+
+fun get_isabelle_thm_and_warn thyname thmname importerconc thy =
+ let
+ val (a, b) = get_isabelle_thm thyname thmname importerconc thy
+ fun warn () =
+ let
+ val (_,importerconc') = disamb_term importerconc
+ val i2h_conc = Thm.symmetric (rewrite_importer_term (HOLogic.mk_Trueprop importerconc') thy)
+ in
+ case concl_of i2h_conc of
+ Const("==",_) $ lhs $ _ =>
+ (warning ("Failed lookup of theorem '"^thmname^"':");
+ writeln "Original conclusion:";
+ prin importerconc';
+ writeln "Modified conclusion:";
+ prin lhs)
+ | _ => ()
+ end
+ in
+ case b of
+ NONE => (warn () handle _ => (); (a,b)) (* FIXME avoid handle _ *)
+ | _ => (a, b)
+ end
+
+fun get_thm thyname thmname thy =
+ case get_importer_thm thyname thmname thy of
+ SOME hth => (thy,SOME hth)
+ | NONE => ((case import_proof_concl thyname thmname thy of
+ SOME f => get_isabelle_thm_and_warn thyname thmname (f thy) thy
+ | NONE => (message "No conclusion"; (thy,NONE)))
+ handle IO.Io _ => (message "IO exception"; (thy,NONE))
+ | PK _ => (message "PK exception"; (thy,NONE)))
+
+fun rename_const thyname thy name =
+ case get_importer_const_renaming thyname name thy of
+ SOME cname => cname
+ | NONE => name
+
+fun get_def thyname constname rhs thy =
+ let
+ val constname = rename_const thyname thy constname
+ val (thmname,thy') = get_defname thyname constname thy
+ val _ = message ("Looking for definition " ^ thyname ^ "." ^ thmname)
+ in
+ get_isabelle_thm_and_warn thyname thmname (mk_teq (thyname ^ "." ^ constname) rhs thy') thy'
+ end
+
+fun get_axiom thyname axname thy =
+ case get_thm thyname axname thy of
+ arg as (_,SOME _) => arg
+ | _ => raise ERR "get_axiom" ("Trying to retrieve axiom (" ^ axname ^ ")")
+
+fun intern_store_thm gen_output thyname thmname hth thy =
+ let
+ val (hth' as HOLThm (args as (_,th))) = norm_hthm thy hth
+ val rew = rewrite_importer_term (concl_of th) thy
+ val th = Thm.equal_elim rew th
+ val thy' = add_importer_pending thyname thmname args thy
+ val th = disambiguate_frees th
+ val th = Object_Logic.rulify th
+ val thy2 =
+ if gen_output
+ then
+ add_dump ("lemma " ^ (quotename thmname) ^ ": " ^
+ (smart_string_of_thm (Syntax.init_pretty_global thy') th) ^ "\n by (import " ^
+ thyname ^ " " ^ (quotename thmname) ^ ")") thy'
+ else thy'
+ in
+ (thy2,hth')
+ end
+
+val store_thm = intern_store_thm true
+
+fun mk_REFL ctm =
+ let
+ val cty = Thm.ctyp_of_term ctm
+ in
+ Drule.instantiate' [SOME cty] [SOME ctm] reflexivity_thm
+ end
+
+fun REFL tm thy =
+ let
+ val _ = message "REFL:"
+ val (info,tm') = disamb_term tm
+ val ctm = Thm.cterm_of thy tm'
+ val res = HOLThm(rens_of info,mk_REFL ctm)
+ val _ = if_debug pth res
+ in
+ (thy,res)
+ end
+
+fun ASSUME tm thy =
+ let
+ val _ = message "ASSUME:"
+ val (info,tm') = disamb_term tm
+ val ctm = Thm.cterm_of thy (HOLogic.mk_Trueprop tm')
+ val th = Thm.trivial ctm
+ val res = HOLThm(rens_of info,th)
+ val _ = if_debug pth res
+ in
+ (thy,res)
+ end
+
+fun INST_TYPE lambda (hth as HOLThm(_,th)) thy =
+ let
+ val _ = message "INST_TYPE:"
+ val _ = if_debug pth hth
+ val tys_before = Misc_Legacy.add_term_tfrees (prop_of th,[])
+ val th1 = Thm.varifyT_global th
+ val tys_after = Misc_Legacy.add_term_tvars (prop_of th1,[])
+ val tyinst = map (fn (bef, iS) =>
+ (case try (Lib.assoc (TFree bef)) lambda of
+ SOME ty => (ctyp_of thy (TVar iS), ctyp_of thy ty)
+ | NONE => (ctyp_of thy (TVar iS), ctyp_of thy (TFree bef))
+ ))
+ (zip tys_before tys_after)
+ val res = Drule.instantiate_normalize (tyinst,[]) th1
+ val hth = HOLThm([],res)
+ val res = norm_hthm thy hth
+ val _ = message "RESULT:"
+ val _ = if_debug pth res
+ in
+ (thy,res)
+ end
+
+fun INST sigma hth thy =
+ let
+ val _ = message "INST:"
+ val _ = if_debug (app (fn (x,y) => (prin x; prin y))) sigma
+ val _ = if_debug pth hth
+ val (sdom,srng) = ListPair.unzip (rev sigma)
+ val th = hthm2thm hth
+ val th1 = mk_INST (map (cterm_of thy) sdom) (map (cterm_of thy) srng) th
+ val res = HOLThm([],th1)
+ val _ = message "RESULT:"
+ val _ = if_debug pth res
+ in
+ (thy,res)
+ end
+
+fun EQ_IMP_RULE (hth as HOLThm(rens,th)) thy =
+ let
+ val _ = message "EQ_IMP_RULE:"
+ val _ = if_debug pth hth
+ val res = HOLThm(rens,th RS eqimp_thm)
+ val _ = message "RESULT:"
+ val _ = if_debug pth res
+ in
+ (thy,res)
+ end
+
+fun mk_EQ_MP th1 th2 = [beta_eta_thm th1, beta_eta_thm th2] MRS eqmp_thm
+
+fun EQ_MP hth1 hth2 thy =
+ let
+ val _ = message "EQ_MP:"
+ val _ = if_debug pth hth1
+ val _ = if_debug pth hth2
+ val (info,[th1,th2]) = disamb_thms [hth1,hth2]
+ val res = HOLThm(rens_of info,mk_EQ_MP th1 th2)
+ val _ = message "RESULT:"
+ val _ = if_debug pth res
+ in
+ (thy,res)
+ end
+
+fun mk_COMB th1 th2 thy =
+ let
+ val (f,g) = case concl_of th1 of
+ _ $ (Const(@{const_name HOL.eq},_) $ f $ g) => (f,g)
+ | _ => raise ERR "mk_COMB" "First theorem not an equality"
+ val (x,y) = case concl_of th2 of
+ _ $ (Const(@{const_name HOL.eq},_) $ x $ y) => (x,y)
+ | _ => raise ERR "mk_COMB" "Second theorem not an equality"
+ val fty = type_of f
+ val (fd,fr) = Term.dest_funT fty
+ val comb_thm' = Drule.instantiate'
+ [SOME (ctyp_of thy fd),SOME (ctyp_of thy fr)]
+ [SOME (cterm_of thy f),SOME (cterm_of thy g),
+ SOME (cterm_of thy x),SOME (cterm_of thy y)] comb_thm
+ in
+ [th1,th2] MRS comb_thm'
+ end
+
+fun SUBST rews ctxt hth thy =
+ let
+ val _ = message "SUBST:"
+ val _ = if_debug (app pth) rews
+ val _ = if_debug prin ctxt
+ val _ = if_debug pth hth
+ val (info,th) = disamb_thm hth
+ val (info1,ctxt') = disamb_term_from info ctxt
+ val (info2,rews') = disamb_thms_from info1 rews
+
+ val cctxt = cterm_of thy ctxt'
+ fun subst th [] = th
+ | subst th (rew::rews) = subst (mk_COMB th rew thy) rews
+ val res = HOLThm(rens_of info2,mk_EQ_MP (subst (mk_REFL cctxt) rews') th)
+ val _ = message "RESULT:"
+ val _ = if_debug pth res
+ in
+ (thy,res)
+ end
+
+fun DISJ_CASES hth hth1 hth2 thy =
+ let
+ val _ = message "DISJ_CASES:"
+ val _ = if_debug (app pth) [hth,hth1,hth2]
+ val (info,th) = disamb_thm hth
+ val (info1,th1) = disamb_thm_from info hth1
+ val (info2,th2) = disamb_thm_from info1 hth2
+ val th1 = norm_hyps th1
+ val th2 = norm_hyps th2
+ val (l,r) = case concl_of th of
+ _ $ (Const(@{const_name HOL.disj},_) $ l $ r) => (l,r)
+ | _ => raise ERR "DISJ_CASES" "Conclusion not a disjunction"
+ val th1' = rearrange thy (HOLogic.mk_Trueprop l) th1
+ val th2' = rearrange thy (HOLogic.mk_Trueprop r) th2
+ val res1 = th RS disj_cases_thm
+ val res2 = uniq_compose ((nprems_of th1')-1) th1' ((nprems_of th)+1) res1
+ val res3 = uniq_compose ((nprems_of th2')-1) th2' (nprems_of res2) res2
+ val res = HOLThm(rens_of info2,res3)
+ val _ = message "RESULT:"
+ val _ = if_debug pth res
+ in
+ (thy,res)
+ end
+
+fun DISJ1 hth tm thy =
+ let
+ val _ = message "DISJ1:"
+ val _ = if_debug pth hth
+ val _ = if_debug prin tm
+ val (info,th) = disamb_thm hth
+ val (info',tm') = disamb_term_from info tm
+ val ct = Thm.cterm_of thy tm'
+ val disj1_thm' = Drule.instantiate' [] [NONE,SOME ct] disj1_thm
+ val res = HOLThm(rens_of info',th RS disj1_thm')
+ val _ = message "RESULT:"
+ val _ = if_debug pth res
+ in
+ (thy,res)
+ end
+
+fun DISJ2 tm hth thy =
+ let
+ val _ = message "DISJ1:"
+ val _ = if_debug prin tm
+ val _ = if_debug pth hth
+ val (info,th) = disamb_thm hth
+ val (info',tm') = disamb_term_from info tm
+ val ct = Thm.cterm_of thy tm'
+ val disj2_thm' = Drule.instantiate' [] [NONE,SOME ct] disj2_thm
+ val res = HOLThm(rens_of info',th RS disj2_thm')
+ val _ = message "RESULT:"
+ val _ = if_debug pth res
+ in
+ (thy,res)
+ end
+
+fun IMP_ANTISYM hth1 hth2 thy =
+ let
+ val _ = message "IMP_ANTISYM:"
+ val _ = if_debug pth hth1
+ val _ = if_debug pth hth2
+ val (info,[th1,th2]) = disamb_thms [hth1,hth2]
+ val th = [beta_eta_thm th1,beta_eta_thm th2] MRS imp_antisym_thm
+ val res = HOLThm(rens_of info,th)
+ val _ = message "RESULT:"
+ val _ = if_debug pth res
+ in
+ (thy,res)
+ end
+
+fun SYM (hth as HOLThm(rens,th)) thy =
+ let
+ val _ = message "SYM:"
+ val _ = if_debug pth hth
+ val th = th RS symmetry_thm
+ val res = HOLThm(rens,th)
+ val _ = message "RESULT:"
+ val _ = if_debug pth res
+ in
+ (thy,res)
+ end
+
+fun MP hth1 hth2 thy =
+ let
+ val _ = message "MP:"
+ val _ = if_debug pth hth1
+ val _ = if_debug pth hth2
+ val (info,[th1,th2]) = disamb_thms [hth1,hth2]
+ val th = [beta_eta_thm th1,beta_eta_thm th2] MRS mp_thm
+ val res = HOLThm(rens_of info,th)
+ val _ = message "RESULT:"
+ val _ = if_debug pth res
+ in
+ (thy,res)
+ end
+
+fun CONJ hth1 hth2 thy =
+ let
+ val _ = message "CONJ:"
+ val _ = if_debug pth hth1
+ val _ = if_debug pth hth2
+ val (info,[th1,th2]) = disamb_thms [hth1,hth2]
+ val th = [th1,th2] MRS conj_thm
+ val res = HOLThm(rens_of info,th)
+ val _ = message "RESULT:"
+ val _ = if_debug pth res
+ in
+ (thy,res)
+ end
+
+fun CONJUNCT1 (hth as HOLThm(rens,th)) thy =
+ let
+ val _ = message "CONJUNCT1:"
+ val _ = if_debug pth hth
+ val res = HOLThm(rens,th RS conjunct1_thm)
+ val _ = message "RESULT:"
+ val _ = if_debug pth res
+ in
+ (thy,res)
+ end
+
+fun CONJUNCT2 (hth as HOLThm(rens,th)) thy =
+ let
+ val _ = message "CONJUNCT1:"
+ val _ = if_debug pth hth
+ val res = HOLThm(rens,th RS conjunct2_thm)
+ val _ = message "RESULT:"
+ val _ = if_debug pth res
+ in
+ (thy,res)
+ end
+
+fun EXISTS ex wit hth thy =
+ let
+ val _ = message "EXISTS:"
+ val _ = if_debug prin ex
+ val _ = if_debug prin wit
+ val _ = if_debug pth hth
+ val (info,th) = disamb_thm hth
+ val (info',[ex',wit']) = disamb_terms_from info [ex,wit]
+ val cwit = cterm_of thy wit'
+ val cty = ctyp_of_term cwit
+ val a = case ex' of
+ (Const(@{const_name Ex},_) $ a) => a
+ | _ => raise ERR "EXISTS" "Argument not existential"
+ val ca = cterm_of thy a
+ val exists_thm' = beta_eta_thm (Drule.instantiate' [SOME cty] [SOME ca,SOME cwit] exists_thm)
+ val th1 = beta_eta_thm th
+ val th2 = implies_elim_all th1
+ val th3 = th2 COMP exists_thm'
+ val th = implies_intr_hyps th3
+ val res = HOLThm(rens_of info',th)
+ val _ = message "RESULT:"
+ val _ = if_debug pth res
+ in
+ (thy,res)
+ end
+
+fun CHOOSE v hth1 hth2 thy =
+ let
+ val _ = message "CHOOSE:"
+ val _ = if_debug prin v
+ val _ = if_debug pth hth1
+ val _ = if_debug pth hth2
+ val (info,[th1,th2]) = disamb_thms [hth1,hth2]
+ val (info',v') = disamb_term_from info v
+ fun strip 0 _ th = th
+ | strip n (p::ps) th =
+ strip (n-1) ps (Thm.implies_elim th (Thm.assume p))
+ | strip _ _ _ = raise ERR "CHOOSE" "strip error"
+ val cv = cterm_of thy v'
+ val th2 = norm_hyps th2
+ val cvty = ctyp_of_term cv
+ val c = HOLogic.dest_Trueprop (concl_of th2)
+ val cc = cterm_of thy c
+ val a = case concl_of th1 of
+ _ $ (Const(@{const_name Ex},_) $ a) => a
+ | _ => raise ERR "CHOOSE" "Conclusion not existential"
+ val ca = cterm_of (theory_of_thm th1) a
+ val choose_thm' = beta_eta_thm (Drule.instantiate' [SOME cvty] [SOME ca,SOME cc] choose_thm)
+ val th21 = rearrange thy (HOLogic.mk_Trueprop (a $ v')) th2
+ val th22 = strip ((nprems_of th21)-1) (cprems_of th21) th21
+ val th23 = beta_eta_thm (Thm.forall_intr cv th22)
+ val th11 = implies_elim_all (beta_eta_thm th1)
+ val th' = th23 COMP (th11 COMP choose_thm')
+ val th = implies_intr_hyps th'
+ val res = HOLThm(rens_of info',th)
+ val _ = message "RESULT:"
+ val _ = if_debug pth res
+ in
+ (thy,res)
+ end
+
+fun GEN v hth thy =
+ let
+ val _ = message "GEN:"
+ val _ = if_debug prin v
+ val _ = if_debug pth hth
+ val (info,th) = disamb_thm hth
+ val (info',v') = disamb_term_from info v
+ val res = HOLThm(rens_of info',mk_GEN v' th thy)
+ val _ = message "RESULT:"
+ val _ = if_debug pth res
+ in
+ (thy,res)
+ end
+
+fun SPEC tm hth thy =
+ let
+ val _ = message "SPEC:"
+ val _ = if_debug prin tm
+ val _ = if_debug pth hth
+ val (info,th) = disamb_thm hth
+ val (info',tm') = disamb_term_from info tm
+ val ctm = Thm.cterm_of thy tm'
+ val cty = Thm.ctyp_of_term ctm
+ val spec' = Drule.instantiate' [SOME cty] [NONE,SOME ctm] spec_thm
+ val th = th RS spec'
+ val res = HOLThm(rens_of info',th)
+ val _ = message "RESULT:"
+ val _ = if_debug pth res
+ in
+ (thy,res)
+ end
+
+fun COMB hth1 hth2 thy =
+ let
+ val _ = message "COMB:"
+ val _ = if_debug pth hth1
+ val _ = if_debug pth hth2
+ val (info,[th1,th2]) = disamb_thms [hth1,hth2]
+ val res = HOLThm(rens_of info,mk_COMB th1 th2 thy)
+ val _ = message "RESULT:"
+ val _ = if_debug pth res
+ in
+ (thy,res)
+ end
+
+fun TRANS hth1 hth2 thy =
+ let
+ val _ = message "TRANS:"
+ val _ = if_debug pth hth1
+ val _ = if_debug pth hth2
+ val (info,[th1,th2]) = disamb_thms [hth1,hth2]
+ val th = [th1,th2] MRS trans_thm
+ val res = HOLThm(rens_of info,th)
+ val _ = message "RESULT:"
+ val _ = if_debug pth res
+ in
+ (thy,res)
+ end
+
+
+fun CCONTR tm hth thy =
+ let
+ val _ = message "SPEC:"
+ val _ = if_debug prin tm
+ val _ = if_debug pth hth
+ val (info,th) = disamb_thm hth
+ val (info',tm') = disamb_term_from info tm
+ val th = norm_hyps th
+ val ct = cterm_of thy tm'
+ val th1 = rearrange thy (HOLogic.mk_Trueprop (Const(@{const_name Not},HOLogic.boolT-->HOLogic.boolT) $ tm')) th
+ val ccontr_thm' = Drule.instantiate' [] [SOME ct] ccontr_thm
+ val res1 = uniq_compose ((nprems_of th1) - 1) th1 1 ccontr_thm'
+ val res = HOLThm(rens_of info',res1)
+ val _ = message "RESULT:"
+ val _ = if_debug pth res
+ in
+ (thy,res)
+ end
+
+fun mk_ABS v th thy =
+ let
+ val cv = cterm_of thy v
+ val th1 = implies_elim_all (beta_eta_thm th)
+ val (f,g) = case concl_of th1 of
+ _ $ (Const(@{const_name HOL.eq},_) $ f $ g) => (Term.lambda v f,Term.lambda v g)
+ | _ => raise ERR "mk_ABS" "Bad conclusion"
+ val (fd,fr) = Term.dest_funT (type_of f)
+ val abs_thm' = Drule.instantiate' [SOME (ctyp_of thy fd), SOME (ctyp_of thy fr)] [SOME (cterm_of thy f), SOME (cterm_of thy g)] abs_thm
+ val th2 = Thm.forall_intr cv th1
+ val th3 = th2 COMP abs_thm'
+ val res = implies_intr_hyps th3
+ in
+ res
+ end
+
+fun ABS v hth thy =
+ let
+ val _ = message "ABS:"
+ val _ = if_debug prin v
+ val _ = if_debug pth hth
+ val (info,th) = disamb_thm hth
+ val (info',v') = disamb_term_from info v
+ val res = HOLThm(rens_of info',mk_ABS v' th thy)
+ val _ = message "RESULT:"
+ val _ = if_debug pth res
+ in
+ (thy,res)
+ end
+
+fun GEN_ABS copt vlist hth thy =
+ let
+ val _ = message "GEN_ABS:"
+ val _ = case copt of
+ SOME c => if_debug prin c
+ | NONE => ()
+ val _ = if_debug (app prin) vlist
+ val _ = if_debug pth hth
+ val (info,th) = disamb_thm hth
+ val (info',vlist') = disamb_terms_from info vlist
+ val th1 =
+ case copt of
+ SOME (Const(cname,cty)) =>
+ let
+ fun inst_type ty1 ty2 (TVar _) = raise ERR "GEN_ABS" "Type variable found!"
+ | inst_type ty1 ty2 (ty as TFree _) = if ty1 = ty
+ then ty2
+ else ty
+ | inst_type ty1 ty2 (Type(name,tys)) =
+ Type(name,map (inst_type ty1 ty2) tys)
+ in
+ fold_rev (fn v => fn th =>
+ let
+ val cdom = fst (Term.dest_funT (fst (Term.dest_funT cty)))
+ val vty = type_of v
+ val newcty = inst_type cdom vty cty
+ val cc = cterm_of thy (Const(cname,newcty))
+ in
+ mk_COMB (mk_REFL cc) (mk_ABS v th thy) thy
+ end) vlist' th
+ end
+ | SOME _ => raise ERR "GEN_ABS" "Bad constant"
+ | NONE =>
+ fold_rev (fn v => fn th => mk_ABS v th thy) vlist' th
+ val res = HOLThm(rens_of info',th1)
+ val _ = message "RESULT:"
+ val _ = if_debug pth res
+ in
+ (thy,res)
+ end
+
+fun NOT_INTRO (hth as HOLThm(rens,th)) thy =
+ let
+ val _ = message "NOT_INTRO:"
+ val _ = if_debug pth hth
+ val th1 = implies_elim_all (beta_eta_thm th)
+ val a = case concl_of th1 of
+ _ $ (Const(@{const_name HOL.implies},_) $ a $ Const(@{const_name False},_)) => a
+ | _ => raise ERR "NOT_INTRO" "Conclusion of bad form"
+ val ca = cterm_of thy a
+ val th2 = Thm.equal_elim (Drule.instantiate' [] [SOME ca] not_intro_thm) th1
+ val res = HOLThm(rens,implies_intr_hyps th2)
+ val _ = message "RESULT:"
+ val _ = if_debug pth res
+ in
+ (thy,res)
+ end
+
+fun NOT_ELIM (hth as HOLThm(rens,th)) thy =
+ let
+ val _ = message "NOT_INTRO:"
+ val _ = if_debug pth hth
+ val th1 = implies_elim_all (beta_eta_thm th)
+ val a = case concl_of th1 of
+ _ $ (Const(@{const_name Not},_) $ a) => a
+ | _ => raise ERR "NOT_ELIM" "Conclusion of bad form"
+ val ca = cterm_of thy a
+ val th2 = Thm.equal_elim (Drule.instantiate' [] [SOME ca] not_elim_thm) th1
+ val res = HOLThm(rens,implies_intr_hyps th2)
+ val _ = message "RESULT:"
+ val _ = if_debug pth res
+ in
+ (thy,res)
+ end
+
+fun DISCH tm hth thy =
+ let
+ val _ = message "DISCH:"
+ val _ = if_debug prin tm
+ val _ = if_debug pth hth
+ val (info,th) = disamb_thm hth
+ val (info',tm') = disamb_term_from info tm
+ val th1 = beta_eta_thm th
+ val th2 = implies_elim_all th1
+ val th3 = Thm.implies_intr (cterm_of thy (HOLogic.mk_Trueprop tm')) th2
+ val th4 = th3 COMP disch_thm
+ val res = HOLThm (rens_of info', implies_intr_hyps th4)
+ val _ = message "RESULT:"
+ val _ = if_debug pth res
+ in
+ (thy,res)
+ end
+
+val spaces = space_implode " "
+
+fun new_definition thyname constname rhs thy =
+ let
+ val constname = rename_const thyname thy constname
+ val redeclared = is_some (Sign.const_type thy (Sign.intern_const thy constname));
+ val _ = warning ("Introducing constant " ^ constname)
+ val (thmname,thy) = get_defname thyname constname thy
+ val (_,rhs') = disamb_term rhs
+ val ctype = type_of rhs'
+ val csyn = mk_syn thy constname
+ val thy1 = case Importer_DefThy.get thy of
+ Replaying _ => thy
+ | _ => Sign.add_consts_i [(Binding.name constname,ctype,csyn)] thy
+ val eq = mk_defeq constname rhs' thy1
+ val (thms, thy2) = Global_Theory.add_defs false [((Binding.name thmname,eq),[])] thy1
+ val def_thm = hd thms
+ val thm' = def_thm RS meta_eq_to_obj_eq_thm
+ val (thy',th) = (thy2, thm')
+ val fullcname = Sign.intern_const thy' constname
+ val thy'' = add_importer_const_mapping thyname constname true fullcname thy'
+ val (linfo,tm24) = disamb_term (mk_teq constname rhs' thy'')
+ val rew = rewrite_importer_term eq thy''
+ val crhs = cterm_of thy'' (#2 (Logic.dest_equals (prop_of rew)))
+ val thy22 =
+ if Thm.def_name constname = thmname andalso not redeclared andalso csyn = NoSyn
+ then
+ let
+ val ctxt = Syntax.init_pretty_global thy''
+ val p1 = quotename constname
+ val p2 = Syntax.string_of_typ ctxt ctype
+ val p3 = string_of_mixfix csyn
+ val p4 = smart_string_of_cterm ctxt crhs
+ in
+ add_dump ("definition\n " ^ p1 ^ " :: \"" ^ p2 ^ "\" "^ p3 ^ " where\n " ^ p4) thy''
+ end
+ else
+ let val ctxt = Syntax.init_pretty_global thy'' in
+ add_dump ("consts\n " ^ quotename constname ^ " :: \"" ^
+ Syntax.string_of_typ ctxt ctype ^
+ "\" " ^ string_of_mixfix csyn ^ "\n\ndefs\n " ^
+ quotename thmname ^ ": " ^ smart_string_of_cterm ctxt crhs) thy''
+ end
+ val hth = case Shuffler.set_prop thy22 (HOLogic.mk_Trueprop tm24) [("",th)] of
+ SOME (_,res) => HOLThm(rens_of linfo,res)
+ | NONE => raise ERR "new_definition" "Bad conclusion"
+ val fullname = Sign.full_bname thy22 thmname
+ val thy22' = case opt_get_output_thy thy22 of
+ "" => add_importer_mapping thyname thmname fullname thy22
+ | output_thy =>
+ let
+ val moved_thmname = output_thy ^ "." ^ thyname ^ "." ^ thmname
+ in
+ thy22 |> add_importer_move fullname moved_thmname
+ |> add_importer_mapping thyname thmname moved_thmname
+ end
+ val _ = message "new_definition:"
+ val _ = if_debug pth hth
+ in
+ (thy22',hth)
+ end
+
+fun new_specification thyname thmname names hth thy =
+ case Importer_DefThy.get thy of
+ Replaying _ => (thy,hth)
+ | _ =>
+ let
+ val _ = message "NEW_SPEC:"
+ val _ = if_debug pth hth
+ val names = map (rename_const thyname thy) names
+ val _ = warning ("Introducing constants " ^ commas names)
+ val (HOLThm(rens,th)) = norm_hthm thy hth
+ val thy1 = case Importer_DefThy.get thy of
+ Replaying _ => thy
+ | _ =>
+ let
+ fun dest_eta_abs (Abs(x,xT,body)) = (x,xT,body)
+ | dest_eta_abs body =
+ let
+ val (dT,_) = Term.dest_funT (type_of body)
+ in
+ ("x",dT,body $ Bound 0)
+ end
+ handle TYPE _ => raise ERR "new_specification" "not an abstraction type"
+ fun dest_exists (Const(@{const_name Ex},_) $ abody) =
+ dest_eta_abs abody
+ | dest_exists _ =
+ raise ERR "new_specification" "Bad existential formula"
+
+ val (consts,_) = Library.foldl (fn ((cs,ex),cname) =>
+ let
+ val (_,cT,p) = dest_exists ex
+ in
+ ((cname,cT,mk_syn thy cname)::cs,p)
+ end) (([],HOLogic.dest_Trueprop (concl_of th)),names)
+ val str = Library.foldl (fn (acc, (c, T, csyn)) =>
+ acc ^ "\n " ^ quotename c ^ " :: \"" ^
+ Syntax.string_of_typ_global thy T ^ "\" " ^ string_of_mixfix csyn) ("consts", consts)
+ val thy' = add_dump str thy
+ in
+ Sign.add_consts_i (map (fn (c, T, mx) => (Binding.name c, T, mx)) consts) thy'
+ end
+
+ val thy1 = fold_rev (fn name => fn thy =>
+ snd (get_defname thyname name thy)) names thy1
+ fun new_name name = fst (get_defname thyname name thy1)
+ val names' = map (fn name => (new_name name,name,false)) names
+ val (thy',res) = Choice_Specification.add_specification NONE
+ names'
+ (thy1,th)
+ val res' = Thm.unvarify_global res
+ val hth = HOLThm(rens,res')
+ val rew = rewrite_importer_term (concl_of res') thy'
+ val th = Thm.equal_elim rew res'
+ fun handle_const (name,thy) =
+ let
+ val defname = Thm.def_name name
+ val (newname,thy') = get_defname thyname name thy
+ in
+ (if defname = newname
+ then quotename name
+ else (quotename newname) ^ ": " ^ (quotename name),thy')
+ end
+ val (new_names,thy') = fold_rev (fn name => fn (names, thy) =>
+ let
+ val (name',thy') = handle_const (name,thy)
+ in
+ (name'::names,thy')
+ end) names ([], thy')
+ val thy'' =
+ add_dump ("specification (" ^ (spaces new_names) ^ ") " ^ thmname ^ ": " ^
+ (smart_string_of_thm (Syntax.init_pretty_global thy') th) ^
+ "\n by (import " ^ thyname ^ " " ^ thmname ^ ")")
+ thy'
+ val _ = message "RESULT:"
+ val _ = if_debug pth hth
+ in
+ intern_store_thm false thyname thmname hth thy''
+ end
+
+fun new_axiom name _ _ = raise ERR "new_axiom" ("Oh, no you don't! (" ^ name ^ ")")
+
+fun to_isa_thm (hth as HOLThm(_,th)) =
+ let
+ val (HOLThm args) = norm_hthm (theory_of_thm th) hth
+ in
+ apsnd Thm.strip_shyps args
+ end
+
+fun to_isa_term tm = tm
+
+local
+ val light_nonempty = @{thm light_ex_imp_nonempty}
+ val ex_imp_nonempty = @{thm ex_imp_nonempty}
+ val typedef_hol2hol4 = @{thm typedef_hol2hol4}
+ val typedef_hol2hollight = @{thm typedef_hol2hollight}
+in
+fun new_type_definition thyname thmname tycname hth thy =
+ case Importer_DefThy.get thy of
+ Replaying _ => (thy,hth)
+ | _ =>
+ let
+ val _ = message "TYPE_DEF:"
+ val _ = if_debug pth hth
+ val _ = warning ("Introducing type " ^ tycname)
+ val (HOLThm(rens,td_th)) = norm_hthm thy hth
+ val th2 = beta_eta_thm (td_th RS ex_imp_nonempty)
+ val c = case concl_of th2 of
+ _ $ (Const(@{const_name Ex},_) $ Abs(_,_,Const(@{const_name Set.member},_) $ _ $ c)) => c
+ | _ => raise ERR "new_type_definition" "Bad type definition theorem"
+ val tfrees = Misc_Legacy.term_tfrees c
+ val tnames = map fst tfrees
+ val tsyn = mk_syn thy tycname
+ val ((_, typedef_info), thy') =
+ Typedef.add_typedef_global false (SOME (Binding.name thmname))
+ (Binding.name tycname, map (rpair dummyS) tnames, tsyn) c NONE (rtac th2 1) thy
+
+ val th3 = (#type_definition (#2 typedef_info)) RS typedef_hol2hol4
+
+ val fulltyname = Sign.intern_type thy' tycname
+ val thy'' = add_importer_type_mapping thyname tycname true fulltyname thy'
+
+ val (hth' as HOLThm args) = norm_hthm thy'' (HOLThm(rens,th3))
+ val _ = if has_ren hth' then warning ("Theorem " ^ thmname ^ " needs variable-disambiguating")
+ else ()
+ val thy4 = add_importer_pending thyname thmname args thy''
+
+ val rew = rewrite_importer_term (concl_of td_th) thy4
+ val th = Thm.equal_elim rew (Thm.transfer thy4 td_th)
+ val c = case HOLogic.dest_Trueprop (prop_of th) of
+ Const(@{const_name Ex},exT) $ P =>
+ let
+ val PT = domain_type exT
+ in
+ Const (@{const_name Collect},PT-->HOLogic.mk_setT (domain_type PT)) $ P
+ end
+ | _ => error "Internal error in ProofKernel.new_typedefinition"
+ val tnames_string = if null tnames
+ then ""
+ else "(" ^ commas tnames ^ ") "
+ val proc_prop =
+ smart_string_of_cterm
+ (Syntax.init_pretty_global thy4
+ |> not (null tnames) ? Config.put show_all_types true)
+ val thy5 = add_dump ("typedef (open) " ^ tnames_string ^ (quotename tycname) ^ " = " ^ (proc_prop (cterm_of thy4 c)) ^ " "
+ ^ (string_of_mixfix tsyn) ^ "\n by (rule typedef_helper,import " ^ thyname ^ " " ^ thmname ^ ")") thy4
+
+ val thy6 = add_dump ("lemmas " ^ thmname ^ " = typedef_hol2importer [OF type_definition_" ^ tycname ^ "]") thy5
+
+ val _ = message "RESULT:"
+ val _ = if_debug pth hth'
+ in
+ (thy6,hth')
+ end
+
+fun add_dump_syntax thy name =
+ let
+ val n = quotename name
+ val syn = string_of_mixfix (mk_syn thy name)
+ in
+ add_dump ("syntax\n "^n^" :: _ "^syn) thy
+ end
+
+fun type_introduction thyname thmname tycname abs_name rep_name (P,t) hth thy =
+ case Importer_DefThy.get thy of
+ Replaying _ => (thy,
+ HOLThm([], Global_Theory.get_thm thy (thmname^"_@intern")) handle ERROR _ => hth)
+ | _ =>
+ let
+ val _ = message "TYPE_INTRO:"
+ val _ = if_debug pth hth
+ val _ = warning ("Introducing type " ^ tycname ^ " (with morphisms " ^ abs_name ^ " and " ^ rep_name ^ ")")
+ val (HOLThm(rens,td_th)) = norm_hthm thy hth
+ val tT = type_of t
+ val light_nonempty' =
+ Drule.instantiate' [SOME (ctyp_of thy tT)]
+ [SOME (cterm_of thy P),
+ SOME (cterm_of thy t)] light_nonempty
+ val th2 = beta_eta_thm (td_th RS (beta_eta_thm light_nonempty'))
+ val c = case concl_of th2 of
+ _ $ (Const(@{const_name Ex},_) $ Abs(_,_,Const(@{const_name Set.member},_) $ _ $ c)) => c
+ | _ => raise ERR "type_introduction" "Bad type definition theorem"
+ val tfrees = Misc_Legacy.term_tfrees c
+ val tnames = sort_strings (map fst tfrees)
+ val tsyn = mk_syn thy tycname
+ val ((_, typedef_info), thy') =
+ Typedef.add_typedef_global false NONE
+ (Binding.name tycname, map (rpair dummyS) tnames, tsyn) c
+ (SOME(Binding.name rep_name,Binding.name abs_name)) (rtac th2 1) thy
+ val fulltyname = Sign.intern_type thy' tycname
+ val aty = Type (fulltyname, map mk_vartype tnames)
+ val typedef_hol2hollight' =
+ Drule.instantiate'
+ [SOME (ctyp_of thy' aty), SOME (ctyp_of thy' tT)]
+ [NONE, NONE, NONE, SOME (cterm_of thy' (Free ("a", aty))), SOME (cterm_of thy' (Free ("r", tT)))]
+ typedef_hol2hollight
+ val th4 = (#type_definition (#2 typedef_info)) RS typedef_hol2hollight'
+ val _ = null (Thm.fold_terms Term.add_tvars th4 []) orelse
+ raise ERR "type_introduction" "no type variables expected any more"
+ val _ = null (Thm.fold_terms Term.add_vars th4 []) orelse
+ raise ERR "type_introduction" "no term variables expected any more"
+ val _ = message ("step 3: thyname="^thyname^", tycname="^tycname^", fulltyname="^fulltyname)
+ val thy'' = add_importer_type_mapping thyname tycname true fulltyname thy'
+ val _ = message "step 4"
+ val (hth' as HOLThm args) = norm_hthm thy'' (HOLThm(rens,th4))
+ val thy4 = add_importer_pending thyname thmname args thy''
+
+ val P' = P (* why !? #2 (Logic.dest_equals (concl_of (rewrite_importer_term P thy4))) *)
+ val c =
+ let
+ val PT = type_of P'
+ in
+ Const (@{const_name Collect},PT-->HOLogic.mk_setT (domain_type PT)) $ P'
+ end
+
+ val tnames_string = if null tnames
+ then ""
+ else "(" ^ commas tnames ^ ") "
+ val proc_prop =
+ smart_string_of_cterm
+ (Syntax.init_pretty_global thy4
+ |> not (null tnames) ? Config.put show_all_types true)
+ val thy = add_dump ("typedef (open) " ^ tnames_string ^ (quotename tycname) ^
+ " = " ^ (proc_prop (cterm_of thy4 c)) ^ " " ^
+ (string_of_mixfix tsyn) ^ " morphisms "^
+ (quote rep_name)^" "^(quote abs_name)^"\n"^
+ (" apply (rule light_ex_imp_nonempty[where t="^
+ (proc_prop (cterm_of thy4 t))^"])\n"^
+ (" by (import " ^ thyname ^ " " ^ (quotename thmname) ^ ")"))) thy4
+ val str_aty = Syntax.string_of_typ_global thy aty
+ val thy = add_dump_syntax thy rep_name
+ val thy = add_dump_syntax thy abs_name
+ val thy = add_dump ("lemmas " ^ (quote (thmname^"_@intern")) ^
+ " = typedef_hol2hollight \n"^
+ " [where a=\"a :: "^str_aty^"\" and r=r" ^
+ " ,\n OF "^(quotename ("type_definition_" ^ tycname)) ^ "]") thy
+ val _ = message "RESULT:"
+ val _ = if_debug pth hth'
+ in
+ (thy,hth')
+ end
+end
+
+val prin = prin
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL4/replay.ML Sun Apr 01 14:50:47 2012 +0200
@@ -0,0 +1,344 @@
+(* Title: HOL/Import/replay.ML
+ Author: Sebastian Skalberg (TU Muenchen)
+*)
+
+structure Replay = (* FIXME proper signature *)
+struct
+
+open ProofKernel
+
+exception REPLAY of string * string
+fun ERR f mesg = REPLAY (f,mesg)
+fun NY f = raise ERR f "NOT YET!"
+
+fun replay_proof int_thms thyname thmname prf thy =
+ let
+ fun rp (PRefl tm) thy = ProofKernel.REFL tm thy
+ | rp (PInstT(p,lambda)) thy =
+ let
+ val (thy',th) = rp' p thy
+ in
+ ProofKernel.INST_TYPE lambda th thy'
+ end
+ | rp (PSubst(prfs,ctxt,prf)) thy =
+ let
+ val (thy',ths) = fold_rev (fn p => fn (thy, ths) =>
+ let
+ val (thy',th) = rp' p thy
+ in
+ (thy',th::ths)
+ end) prfs (thy,[])
+ val (thy'',th) = rp' prf thy'
+ in
+ ProofKernel.SUBST ths ctxt th thy''
+ end
+ | rp (PAbs(prf,v)) thy =
+ let
+ val (thy',th) = rp' prf thy
+ in
+ ProofKernel.ABS v th thy'
+ end
+ | rp (PDisch(prf,tm)) thy =
+ let
+ val (thy',th) = rp' prf thy
+ in
+ ProofKernel.DISCH tm th thy'
+ end
+ | rp (PMp(prf1,prf2)) thy =
+ let
+ val (thy1,th1) = rp' prf1 thy
+ val (thy2,th2) = rp' prf2 thy1
+ in
+ ProofKernel.MP th1 th2 thy2
+ end
+ | rp (PHyp tm) thy = ProofKernel.ASSUME tm thy
+ | rp (PDef(seg,name,rhs)) thy =
+ (case ProofKernel.get_def seg name rhs thy of
+ (thy',SOME res) => (thy',res)
+ | (thy',NONE) =>
+ if seg = thyname
+ then ProofKernel.new_definition seg name rhs thy'
+ else raise ERR "replay_proof" ("Too late for term definition: "^seg^" != "^thyname))
+ | rp (POracle _) thy = NY "ORACLE"
+ | rp (PSpec(prf,tm)) thy =
+ let
+ val (thy',th) = rp' prf thy
+ in
+ ProofKernel.SPEC tm th thy'
+ end
+ | rp (PInst(prf,theta)) thy =
+ let
+ val (thy',th) = rp' prf thy
+ in
+ ProofKernel.INST theta th thy'
+ end
+ | rp (PGen(prf,v)) thy =
+ let
+ val (thy',th) = rp' prf thy
+ val p = ProofKernel.GEN v th thy'
+ in
+ p
+ end
+ | rp (PGenAbs(prf,opt,vl)) thy =
+ let
+ val (thy',th) = rp' prf thy
+ in
+ ProofKernel.GEN_ABS opt vl th thy'
+ end
+ | rp (PImpAS(prf1,prf2)) thy =
+ let
+ val (thy1,th1) = rp' prf1 thy
+ val (thy2,th2) = rp' prf2 thy1
+ in
+ ProofKernel.IMP_ANTISYM th1 th2 thy2
+ end
+ | rp (PSym prf) thy =
+ let
+ val (thy1,th) = rp' prf thy
+ in
+ ProofKernel.SYM th thy1
+ end
+ | rp (PTrans(prf1,prf2)) thy =
+ let
+ val (thy1,th1) = rp' prf1 thy
+ val (thy2,th2) = rp' prf2 thy1
+ in
+ ProofKernel.TRANS th1 th2 thy2
+ end
+ | rp (PComb(prf1,prf2)) thy =
+ let
+ val (thy1,th1) = rp' prf1 thy
+ val (thy2,th2) = rp' prf2 thy1
+ in
+ ProofKernel.COMB th1 th2 thy2
+ end
+ | rp (PEqMp(prf1,prf2)) thy =
+ let
+ val (thy1,th1) = rp' prf1 thy
+ val (thy2,th2) = rp' prf2 thy1
+ in
+ ProofKernel.EQ_MP th1 th2 thy2
+ end
+ | rp (PEqImp prf) thy =
+ let
+ val (thy',th) = rp' prf thy
+ in
+ ProofKernel.EQ_IMP_RULE th thy'
+ end
+ | rp (PExists(prf,ex,wit)) thy =
+ let
+ val (thy',th) = rp' prf thy
+ in
+ ProofKernel.EXISTS ex wit th thy'
+ end
+ | rp (PChoose(v,prf1,prf2)) thy =
+ let
+ val (thy1,th1) = rp' prf1 thy
+ val (thy2,th2) = rp' prf2 thy1
+ in
+ ProofKernel.CHOOSE v th1 th2 thy2
+ end
+ | rp (PConj(prf1,prf2)) thy =
+ let
+ val (thy1,th1) = rp' prf1 thy
+ val (thy2,th2) = rp' prf2 thy1
+ in
+ ProofKernel.CONJ th1 th2 thy2
+ end
+ | rp (PConjunct1 prf) thy =
+ let
+ val (thy',th) = rp' prf thy
+ in
+ ProofKernel.CONJUNCT1 th thy'
+ end
+ | rp (PConjunct2 prf) thy =
+ let
+ val (thy',th) = rp' prf thy
+ in
+ ProofKernel.CONJUNCT2 th thy'
+ end
+ | rp (PDisj1(prf,tm)) thy =
+ let
+ val (thy',th) = rp' prf thy
+ in
+ ProofKernel.DISJ1 th tm thy'
+ end
+ | rp (PDisj2(prf,tm)) thy =
+ let
+ val (thy',th) = rp' prf thy
+ in
+ ProofKernel.DISJ2 tm th thy'
+ end
+ | rp (PDisjCases(prf,prf1,prf2)) thy =
+ let
+ val (thy',th) = rp' prf thy
+ val (thy1,th1) = rp' prf1 thy'
+ val (thy2,th2) = rp' prf2 thy1
+ in
+ ProofKernel.DISJ_CASES th th1 th2 thy2
+ end
+ | rp (PNotI prf) thy =
+ let
+ val (thy',th) = rp' prf thy
+ in
+ ProofKernel.NOT_INTRO th thy'
+ end
+ | rp (PNotE prf) thy =
+ let
+ val (thy',th) = rp' prf thy
+ in
+ ProofKernel.NOT_ELIM th thy'
+ end
+ | rp (PContr(prf,tm)) thy =
+ let
+ val (thy',th) = rp' prf thy
+ in
+ ProofKernel.CCONTR tm th thy'
+ end
+ | rp (PTmSpec _) _ = raise ERR "rp" "Shouldn't reach here (PTmSpec)"
+ | rp (PTyDef _) _ = raise ERR "rp" "Shouldn't reach here (PTyDef)"
+ | rp (PTyIntro _) _ = raise ERR "rp" "Shouldn't reach here (PTyIntro)"
+ | rp PDisk _ = raise ERR "rp" "Shouldn't reach here (PDisk)"
+ | rp _ _ = raise ERR "rp" "What the hell is this? Which case did I forget?"
+ and rp' p thy =
+ let
+ val pc = content_of p
+ in
+ case pc of
+ PDisk => (case disk_info_of p of
+ SOME(thyname',thmname) =>
+ (case Int.fromString thmname of
+ SOME i =>
+ if thyname' = thyname
+ then
+ (case Array.sub(int_thms,i-1) of
+ NONE =>
+ let
+ val (thy',th) = rp' (snd (import_proof thyname' thmname thy) thy) thy
+ val _ = Array.update(int_thms,i-1,SOME th)
+ in
+ (thy',th)
+ end
+ | SOME th => (thy,th))
+ else raise ERR "replay_proof" ("Library " ^ thyname' ^ " should be built before " ^ thyname ^ " (" ^ thmname ^ ")")
+ | NONE =>
+ (case ProofKernel.get_thm thyname' thmname thy of
+ (thy',SOME res) => (thy',res)
+ | (thy',NONE) =>
+ if thyname' = thyname
+ then
+ let
+ val _ = writeln ("Found no " ^ thmname ^ " theorem, replaying...")
+ val (_, prf) = import_proof thyname' thmname thy'
+ val prf = prf thy'
+ val (thy',th) = replay_proof int_thms thyname' thmname prf thy'
+ val _ = writeln ("Successfully finished replaying "^thmname^" !")
+ in
+ case content_of prf of
+ PTmSpec _ => (thy',th)
+ | PTyDef _ => (thy',th)
+ | PTyIntro _ => (thy',th)
+ | _ => ProofKernel.store_thm thyname' thmname th thy'
+ end
+ else raise ERR "replay_proof" ("Library " ^ thyname' ^ " should be built before " ^ thyname ^ " (" ^ thmname ^ ")")))
+ | NONE => raise ERR "rp'.PDisk" "Not enough information")
+ | PAxm(name,c) =>
+ (case ProofKernel.get_axiom thyname name thy of
+ (thy',SOME res) => (thy',res)
+ | (thy',NONE) => ProofKernel.new_axiom name c thy')
+ | PTmSpec(seg,names,prf') =>
+ let
+ val (thy',th) = rp' prf' thy
+ in
+ ProofKernel.new_specification seg thmname names th thy'
+ end
+ | PTyDef(seg,name,prf') =>
+ let
+ val (thy',th) = rp' prf' thy
+ in
+ ProofKernel.new_type_definition seg thmname name th thy'
+ end
+ | PTyIntro(seg,name,abs_name,rep_name,P,t,prf') =>
+ let
+ val (thy',th) = rp' prf' thy
+ in
+ ProofKernel.type_introduction seg thmname name abs_name rep_name (P,t) th thy'
+ end
+ | _ => rp pc thy
+ end
+ in
+ rp' prf thy
+ end
+
+fun setup_int_thms thyname thy =
+ let
+ val fname =
+ case ProofKernel.get_proof_dir thyname thy of
+ SOME p => OS.Path.joinDirFile {dir=p,file=OS.Path.joinBaseExt{base = "facts",ext=SOME "lst"}}
+ | NONE => error "Cannot find proof files"
+ val is = TextIO.openIn fname
+ val (num_int_thms,facts) =
+ let
+ fun get_facts facts =
+ case TextIO.inputLine is of
+ NONE => (case facts of
+ i::facts => (the (Int.fromString i),map ProofKernel.protect_factname (rev facts))
+ | _ => raise ERR "replay_thm" "Bad facts.lst file")
+ | SOME fact => get_facts ((String.substring(fact,0,String.size fact -1 ))::facts)
+ in
+ get_facts []
+ end
+ val _ = TextIO.closeIn is
+ val int_thms = Array.array(num_int_thms,NONE:thm option)
+ in
+ (int_thms,facts)
+ end
+
+fun import_single_thm thyname int_thms thmname thy =
+ let
+ fun replay_fact (thmname,thy) =
+ let
+ val prf = mk_proof PDisk
+ val _ = set_disk_info_of prf thyname thmname
+ val _ = writeln ("Replaying "^thmname^" ...")
+ val p = fst (replay_proof int_thms thyname thmname prf thy)
+ in
+ p
+ end
+ in
+ replay_fact (thmname,thy)
+ end
+
+fun import_thms thyname int_thms thmnames thy =
+ let
+ fun replay_fact thmname thy =
+ let
+ val prf = mk_proof PDisk
+ val _ = set_disk_info_of prf thyname thmname
+ val _ = writeln ("Replaying "^thmname^" ...")
+ val p = fst (replay_proof int_thms thyname thmname prf thy)
+ in
+ p
+ end
+ val res_thy = fold replay_fact thmnames thy
+ in
+ res_thy
+ end
+
+fun import_thm thyname thmname thy =
+ let
+ val int_thms = fst (setup_int_thms thyname thy)
+ fun replay_fact (thmname,thy) =
+ let
+ val prf = mk_proof PDisk
+ val _ = set_disk_info_of prf thyname thmname
+ val _ = writeln ("Replaying "^thmname^" ...")
+ val p = fst (replay_proof int_thms thyname thmname prf thy)
+ in
+ p
+ end
+ in
+ replay_fact (thmname,thy)
+ end
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/HOL4/shuffler.ML Sun Apr 01 14:50:47 2012 +0200
@@ -0,0 +1,585 @@
+(* Title: HOL/Import/shuffler.ML
+ Author: Sebastian Skalberg, TU Muenchen
+
+Package for proving two terms equal by normalizing (hence the
+"shuffler" name). Uses the simplifier for the normalization.
+*)
+
+signature Shuffler =
+sig
+ val debug : bool Unsynchronized.ref
+
+ val norm_term : theory -> term -> thm
+ val make_equal : theory -> term -> term -> thm option
+ val set_prop : theory -> term -> (string * thm) list -> (string * thm) option
+
+ val find_potential: theory -> term -> (string * thm) list
+
+ val gen_shuffle_tac: Proof.context -> bool -> (string * thm) list -> int -> tactic
+ val shuffle_tac: Proof.context -> thm list -> int -> tactic
+ val search_tac : Proof.context -> int -> tactic
+
+ val print_shuffles: theory -> unit
+
+ val add_shuffle_rule: thm -> theory -> theory
+ val shuffle_attr: attribute
+
+ val setup : theory -> theory
+end
+
+structure Shuffler : Shuffler =
+struct
+
+val debug = Unsynchronized.ref false
+
+fun if_debug f x = if !debug then f x else ()
+val message = if_debug writeln
+
+val string_of_thm = Print_Mode.setmp [] Display.string_of_thm_without_context;
+
+structure ShuffleData = Theory_Data
+(
+ type T = thm list
+ val empty = []
+ val extend = I
+ val merge = Thm.merge_thms
+)
+
+fun print_shuffles thy =
+ Pretty.writeln (Pretty.big_list "Shuffle theorems:"
+ (map (Display.pretty_thm_global thy) (ShuffleData.get thy)))
+
+val weaken =
+ let
+ val cert = cterm_of Pure.thy
+ val P = Free("P",propT)
+ val Q = Free("Q",propT)
+ val PQ = Logic.mk_implies(P,Q)
+ val PPQ = Logic.mk_implies(P,PQ)
+ val cP = cert P
+ val cPQ = cert PQ
+ val cPPQ = cert PPQ
+ val th1 = Thm.assume cPQ |> implies_intr_list [cPQ,cP]
+ val th3 = Thm.assume cP
+ val th4 = implies_elim_list (Thm.assume cPPQ) [th3,th3]
+ |> implies_intr_list [cPPQ,cP]
+ in
+ Thm.equal_intr th4 th1 |> Drule.export_without_context
+ end
+
+val imp_comm =
+ let
+ val cert = cterm_of Pure.thy
+ val P = Free("P",propT)
+ val Q = Free("Q",propT)
+ val R = Free("R",propT)
+ val PQR = Logic.mk_implies(P,Logic.mk_implies(Q,R))
+ val QPR = Logic.mk_implies(Q,Logic.mk_implies(P,R))
+ val cP = cert P
+ val cQ = cert Q
+ val cPQR = cert PQR
+ val cQPR = cert QPR
+ val th1 = implies_elim_list (Thm.assume cPQR) [Thm.assume cP,Thm.assume cQ]
+ |> implies_intr_list [cPQR,cQ,cP]
+ val th2 = implies_elim_list (Thm.assume cQPR) [Thm.assume cQ,Thm.assume cP]
+ |> implies_intr_list [cQPR,cP,cQ]
+ in
+ Thm.equal_intr th1 th2 |> Drule.export_without_context
+ end
+
+val all_comm =
+ let
+ val cert = cterm_of Pure.thy
+ val xT = TFree("'a",[])
+ val yT = TFree("'b",[])
+ val x = Free("x",xT)
+ val y = Free("y",yT)
+ val P = Free("P",xT-->yT-->propT)
+ val lhs = Logic.all x (Logic.all y (P $ x $ y))
+ val rhs = Logic.all y (Logic.all x (P $ x $ y))
+ val cl = cert lhs
+ val cr = cert rhs
+ val cx = cert x
+ val cy = cert y
+ val th1 = Thm.assume cr
+ |> forall_elim_list [cy,cx]
+ |> forall_intr_list [cx,cy]
+ |> Thm.implies_intr cr
+ val th2 = Thm.assume cl
+ |> forall_elim_list [cx,cy]
+ |> forall_intr_list [cy,cx]
+ |> Thm.implies_intr cl
+ in
+ Thm.equal_intr th1 th2 |> Drule.export_without_context
+ end
+
+val equiv_comm =
+ let
+ val cert = cterm_of Pure.thy
+ val T = TFree("'a",[])
+ val t = Free("t",T)
+ val u = Free("u",T)
+ val ctu = cert (Logic.mk_equals(t,u))
+ val cut = cert (Logic.mk_equals(u,t))
+ val th1 = Thm.assume ctu |> Thm.symmetric |> Thm.implies_intr ctu
+ val th2 = Thm.assume cut |> Thm.symmetric |> Thm.implies_intr cut
+ in
+ Thm.equal_intr th1 th2 |> Drule.export_without_context
+ end
+
+(* This simplification procedure rewrites !!x y. P x y
+deterministicly, in order for the normalization function, defined
+below, to handle nested quantifiers robustly *)
+
+local
+
+exception RESULT of int
+
+fun find_bound n (Bound i) = if i = n then raise RESULT 0
+ else if i = n+1 then raise RESULT 1
+ else ()
+ | find_bound n (t $ u) = (find_bound n t; find_bound n u)
+ | find_bound n (Abs(_,_,t)) = find_bound (n+1) t
+ | find_bound _ _ = ()
+
+fun swap_bound n (Bound i) = if i = n then Bound (n+1)
+ else if i = n+1 then Bound n
+ else Bound i
+ | swap_bound n (t $ u) = (swap_bound n t $ swap_bound n u)
+ | swap_bound n (Abs(x,xT,t)) = Abs(x,xT,swap_bound (n+1) t)
+ | swap_bound n t = t
+
+fun rew_th thy xv yv t =
+ let
+ val lhs = Logic.list_all ([xv,yv],t)
+ val rhs = Logic.list_all ([yv,xv],swap_bound 0 t)
+ val rew = Logic.mk_equals (lhs,rhs)
+ val init = Thm.trivial (cterm_of thy rew)
+ in
+ all_comm RS init
+ end
+
+fun quant_rewrite thy _ (t as Const("all",T1) $ (Abs(x,xT,Const("all",T2) $ Abs(y,yT,body)))) =
+ let
+ val res = (find_bound 0 body;2) handle RESULT i => i
+ in
+ case res of
+ 0 => SOME (rew_th thy (x,xT) (y,yT) body)
+ | 1 => if string_ord(y,x) = LESS
+ then
+ let
+ val newt = Const("all",T1) $ (Abs(y,xT,Const("all",T2) $ Abs(x,yT,body)))
+ val t_th = Thm.reflexive (cterm_of thy t)
+ val newt_th = Thm.reflexive (cterm_of thy newt)
+ in
+ SOME (Thm.transitive t_th newt_th)
+ end
+ else NONE
+ | _ => error "norm_term (quant_rewrite) internal error"
+ end
+ | quant_rewrite _ _ _ = (warning "quant_rewrite: Unknown lhs"; NONE)
+
+fun freeze_thaw_term t =
+ let
+ val tvars = Misc_Legacy.term_tvars t
+ val tfree_names = Misc_Legacy.add_term_tfree_names(t,[])
+ val (type_inst,_) =
+ fold (fn (w as (v,_), S) => fn (inst, used) =>
+ let
+ val v' = singleton (Name.variant_list used) v
+ in
+ ((w,TFree(v',S))::inst,v'::used)
+ end)
+ tvars ([], tfree_names)
+ val t' = subst_TVars type_inst t
+ in
+ (t', map (fn (w,TFree(v,S)) => (v,TVar(w,S))
+ | _ => error "Internal error in Shuffler.freeze_thaw") type_inst)
+ end
+
+fun inst_tfrees thy [] thm = thm
+ | inst_tfrees thy ((name,U)::rest) thm =
+ let
+ val cU = ctyp_of thy U
+ val tfrees = Misc_Legacy.add_term_tfrees (prop_of thm,[])
+ val (rens, thm') = Thm.varifyT_global'
+ (remove (op = o apsnd fst) name tfrees) thm
+ val mid =
+ case rens of
+ [] => thm'
+ | [((_, S), idx)] => Drule.instantiate_normalize
+ ([(ctyp_of thy (TVar (idx, S)), cU)], []) thm'
+ | _ => error "Shuffler.inst_tfrees internal error"
+ in
+ inst_tfrees thy rest mid
+ end
+
+fun is_Abs (Abs _) = true
+ | is_Abs _ = false
+
+fun eta_redex (t $ Bound 0) =
+ let
+ fun free n (Bound i) = i = n
+ | free n (t $ u) = free n t orelse free n u
+ | free n (Abs(_,_,t)) = free (n+1) t
+ | free n _ = false
+ in
+ not (free 0 t)
+ end
+ | eta_redex _ = false
+
+fun eta_contract thy _ origt =
+ let
+ val (typet,Tinst) = freeze_thaw_term origt
+ val (init,thaw) = Misc_Legacy.freeze_thaw (Thm.reflexive (cterm_of thy typet))
+ val final = inst_tfrees thy Tinst o thaw
+ val t = #1 (Logic.dest_equals (prop_of init))
+ val _ =
+ let
+ val lhs = #1 (Logic.dest_equals (prop_of (final init)))
+ in
+ if not (lhs aconv origt)
+ then
+ writeln (cat_lines
+ (["Something is utterly wrong: (orig, lhs, frozen type, t, tinst)",
+ Syntax.string_of_term_global thy origt,
+ Syntax.string_of_term_global thy lhs,
+ Syntax.string_of_term_global thy typet,
+ Syntax.string_of_term_global thy t] @
+ map (fn (n, T) => n ^ ": " ^ Syntax.string_of_typ_global thy T) Tinst))
+ else ()
+ end
+ in
+ case t of
+ Const("all",_) $ (Abs(x,xT,Const("==",_) $ P $ Q)) =>
+ (if eta_redex P andalso eta_redex Q
+ then
+ let
+ val cert = cterm_of thy
+ val v = Free (singleton (Name.variant_list (Term.add_free_names t [])) "v", xT)
+ val cv = cert v
+ val ct = cert t
+ val th = (Thm.assume ct)
+ |> Thm.forall_elim cv
+ |> Thm.abstract_rule x cv
+ val ext_th = Thm.eta_conversion (cert (Abs(x,xT,P)))
+ val th' = Thm.transitive (Thm.symmetric ext_th) th
+ val cu = cert (prop_of th')
+ val uth = Thm.combination (Thm.assume cu) (Thm.reflexive cv)
+ val uth' = (Thm.beta_conversion false (cert (Abs(x,xT,Q) $ v)))
+ |> Thm.transitive uth
+ |> Thm.forall_intr cv
+ |> Thm.implies_intr cu
+ val rew_th = Thm.equal_intr (th' |> Thm.implies_intr ct) uth'
+ val res = final rew_th
+ in
+ SOME res
+ end
+ else NONE)
+ | _ => NONE
+ end
+
+fun eta_expand thy _ origt =
+ let
+ val (typet,Tinst) = freeze_thaw_term origt
+ val (init,thaw) = Misc_Legacy.freeze_thaw (Thm.reflexive (cterm_of thy typet))
+ val final = inst_tfrees thy Tinst o thaw
+ val t = #1 (Logic.dest_equals (prop_of init))
+ val _ =
+ let
+ val lhs = #1 (Logic.dest_equals (prop_of (final init)))
+ in
+ if not (lhs aconv origt)
+ then
+ writeln (cat_lines
+ (["Something is utterly wrong: (orig, lhs, frozen type, t, tinst)",
+ Syntax.string_of_term_global thy origt,
+ Syntax.string_of_term_global thy lhs,
+ Syntax.string_of_term_global thy typet,
+ Syntax.string_of_term_global thy t] @
+ map (fn (n, T) => n ^ ": " ^ Syntax.string_of_typ_global thy T) Tinst))
+ else ()
+ end
+ in
+ case t of
+ Const("==",T) $ P $ Q =>
+ if is_Abs P orelse is_Abs Q
+ then (case domain_type T of
+ Type("fun",[aT,_]) =>
+ let
+ val cert = cterm_of thy
+ val vname = singleton (Name.variant_list (Term.add_free_names t [])) "v"
+ val v = Free(vname,aT)
+ val cv = cert v
+ val ct = cert t
+ val th1 = (Thm.combination (Thm.assume ct) (Thm.reflexive cv))
+ |> Thm.forall_intr cv
+ |> Thm.implies_intr ct
+ val concl = cert (concl_of th1)
+ val th2 = (Thm.assume concl)
+ |> Thm.forall_elim cv
+ |> Thm.abstract_rule vname cv
+ val (lhs,rhs) = Logic.dest_equals (prop_of th2)
+ val elhs = Thm.eta_conversion (cert lhs)
+ val erhs = Thm.eta_conversion (cert rhs)
+ val th2' = Thm.transitive
+ (Thm.transitive (Thm.symmetric elhs) th2)
+ erhs
+ val res = Thm.equal_intr th1 (th2' |> Thm.implies_intr concl)
+ val res' = final res
+ in
+ SOME res'
+ end
+ | _ => NONE)
+ else NONE
+ | _ => error ("Bad eta_expand argument" ^ Syntax.string_of_term_global thy t)
+ end;
+
+fun mk_tfree s = TFree("'"^s,[])
+fun mk_free s t = Free (s,t)
+val xT = mk_tfree "a"
+val yT = mk_tfree "b"
+val x = Free ("x", xT)
+val y = Free ("y", yT)
+val P = mk_free "P" (xT-->yT-->propT)
+val Q = mk_free "Q" (xT-->yT)
+val R = mk_free "R" (xT-->yT)
+in
+
+fun quant_simproc thy = Simplifier.simproc_global_i
+ thy
+ "Ordered rewriting of nested quantifiers"
+ [Logic.all x (Logic.all y (P $ x $ y))]
+ quant_rewrite
+fun eta_expand_simproc thy = Simplifier.simproc_global_i
+ thy
+ "Smart eta-expansion by equivalences"
+ [Logic.mk_equals(Q,R)]
+ eta_expand
+fun eta_contract_simproc thy = Simplifier.simproc_global_i
+ thy
+ "Smart handling of eta-contractions"
+ [Logic.all x (Logic.mk_equals (Q $ x, R $ x))]
+ eta_contract
+end
+
+(* Disambiguates the names of bound variables in a term, returning t
+== t' where all the names of bound variables in t' are unique *)
+
+fun disamb_bound thy t =
+ let
+
+ fun F (t $ u,idx) =
+ let
+ val (t',idx') = F (t,idx)
+ val (u',idx'') = F (u,idx')
+ in
+ (t' $ u',idx'')
+ end
+ | F (Abs(_,xT,t),idx) =
+ let
+ val x' = "x" ^ string_of_int idx
+ val (t',idx') = F (t,idx+1)
+ in
+ (Abs(x',xT,t'),idx')
+ end
+ | F arg = arg
+ val (t',_) = F (t,0)
+ val ct = cterm_of thy t
+ val ct' = cterm_of thy t'
+ val res = Thm.transitive (Thm.reflexive ct) (Thm.reflexive ct')
+ val _ = message ("disamb_term: " ^ (string_of_thm res))
+ in
+ res
+ end
+
+(* Transforms a term t to some normal form t', returning the theorem t
+== t'. This is originally a help function for make_equal, but might
+be handy in its own right, for example for indexing terms. *)
+
+fun norm_term thy t =
+ let
+ val norms = ShuffleData.get thy
+ val ss = Simplifier.global_context thy empty_ss
+ addsimps (map (Thm.transfer thy) norms)
+ addsimprocs [quant_simproc thy, eta_expand_simproc thy,eta_contract_simproc thy]
+ fun chain f th =
+ let
+ val rhs = Thm.rhs_of th
+ in
+ Thm.transitive th (f rhs)
+ end
+ val th =
+ t |> disamb_bound thy
+ |> chain (Simplifier.full_rewrite ss)
+ |> chain Thm.eta_conversion
+ |> Thm.strip_shyps
+ val _ = message ("norm_term: " ^ (string_of_thm th))
+ in
+ th
+ end
+
+
+(* Closes a theorem with respect to free and schematic variables (does
+not touch type variables, though). *)
+
+fun close_thm th =
+ let
+ val thy = Thm.theory_of_thm th
+ val c = prop_of th
+ val vars = Misc_Legacy.add_term_frees (c, Misc_Legacy.add_term_vars(c,[]))
+ in
+ Drule.forall_intr_list (map (cterm_of thy) vars) th
+ end
+
+
+(* Normalizes a theorem's conclusion using norm_term. *)
+
+fun norm_thm thy th =
+ let
+ val c = prop_of th
+ in
+ Thm.equal_elim (norm_term thy c) th
+ end
+
+(* make_equal thy t u tries to construct the theorem t == u under the
+signature thy. If it succeeds, SOME (t == u) is returned, otherwise
+NONE is returned. *)
+
+fun make_equal thy t u =
+ let
+ val t_is_t' = norm_term thy t
+ val u_is_u' = norm_term thy u
+ val th = Thm.transitive t_is_t' (Thm.symmetric u_is_u')
+ val _ = message ("make_equal: SOME " ^ (string_of_thm th))
+ in
+ SOME th
+ end
+ handle THM _ => (message "make_equal: NONE";NONE)
+
+fun match_consts ignore t (* th *) =
+ let
+ fun add_consts (Const (c, _), cs) =
+ if member (op =) ignore c
+ then cs
+ else insert (op =) c cs
+ | add_consts (t $ u, cs) = add_consts (t, add_consts (u, cs))
+ | add_consts (Abs (_, _, t), cs) = add_consts (t, cs)
+ | add_consts (_, cs) = cs
+ val t_consts = add_consts(t,[])
+ in
+ fn (_,th) =>
+ let
+ val th_consts = add_consts(prop_of th,[])
+ in
+ eq_set (op =) (t_consts, th_consts)
+ end
+ end
+
+val collect_ignored = fold_rev (fn thm => fn cs =>
+ let
+ val (lhs, rhs) = Logic.dest_equals (prop_of thm);
+ val consts_lhs = Term.add_const_names lhs [];
+ val consts_rhs = Term.add_const_names rhs [];
+ val ignore_lhs = subtract (op =) consts_rhs consts_lhs;
+ val ignore_rhs = subtract (op =) consts_lhs consts_rhs;
+ in
+ fold_rev (insert (op =)) cs (ignore_lhs @ ignore_rhs)
+ end)
+
+(* set_prop t thms tries to make a theorem with the proposition t from
+one of the theorems thms, by shuffling the propositions around. If it
+succeeds, SOME theorem is returned, otherwise NONE. *)
+
+fun set_prop thy t =
+ let
+ val vars = Misc_Legacy.add_term_frees (t, Misc_Legacy.add_term_vars (t,[]))
+ val closed_t = fold_rev Logic.all vars t
+ val rew_th = norm_term thy closed_t
+ val rhs = Thm.rhs_of rew_th
+
+ fun process [] = NONE
+ | process ((name,th)::thms) =
+ let
+ val norm_th = Thm.varifyT_global (norm_thm thy (close_thm (Thm.transfer thy th)))
+ val triv_th = Thm.trivial rhs
+ val _ = message ("Shuffler.set_prop: Gluing together " ^ (string_of_thm norm_th) ^ " and " ^ (string_of_thm triv_th))
+ val mod_th = case Seq.pull (Thm.bicompose false (*true*) (false,norm_th,0) 1 triv_th) of
+ SOME(th,_) => SOME th
+ | NONE => NONE
+ in
+ case mod_th of
+ SOME mod_th =>
+ let
+ val closed_th = Thm.equal_elim (Thm.symmetric rew_th) mod_th
+ in
+ message ("Shuffler.set_prop succeeded by " ^ name);
+ SOME (name,forall_elim_list (map (cterm_of thy) vars) closed_th)
+ end
+ | NONE => process thms
+ end
+ handle THM _ => process thms
+ in
+ fn thms =>
+ case process thms of
+ res as SOME (_,th) => if (prop_of th) aconv t
+ then res
+ else error "Internal error in set_prop"
+ | NONE => NONE
+ end
+
+fun find_potential thy t =
+ let
+ val shuffles = ShuffleData.get thy
+ val ignored = collect_ignored shuffles []
+ val all_thms =
+ map (`Thm.get_name_hint) (maps #2 (Facts.dest_static [] (Global_Theory.facts_of thy)))
+ in
+ filter (match_consts ignored t) all_thms
+ end
+
+fun gen_shuffle_tac ctxt search thms = SUBGOAL (fn (t, i) =>
+ let
+ val thy = Proof_Context.theory_of ctxt
+ val set = set_prop thy t
+ fun process_tac thms st =
+ case set thms of
+ SOME (_,th) => Seq.of_list (compose (th,i,st))
+ | NONE => Seq.empty
+ in
+ process_tac thms APPEND
+ (if search then process_tac (find_potential thy t) else no_tac)
+ end)
+
+fun shuffle_tac ctxt thms =
+ gen_shuffle_tac ctxt false (map (pair "") thms);
+
+fun search_tac ctxt =
+ gen_shuffle_tac ctxt true (map (pair "premise") (Assumption.all_prems_of ctxt));
+
+fun add_shuffle_rule thm thy =
+ let
+ val shuffles = ShuffleData.get thy
+ in
+ if exists (curry Thm.eq_thm thm) shuffles
+ then (warning ((string_of_thm thm) ^ " already known to the shuffler");
+ thy)
+ else ShuffleData.put (thm::shuffles) thy
+ end
+
+val shuffle_attr = Thm.declaration_attribute (fn th => Context.mapping (add_shuffle_rule th) I);
+
+val setup =
+ Method.setup @{binding shuffle_tac}
+ (Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD' (shuffle_tac ctxt ths)))
+ "solve goal by shuffling terms around" #>
+ Method.setup @{binding search_tac}
+ (Scan.succeed (SIMPLE_METHOD' o search_tac)) "search for suitable theorems" #>
+ add_shuffle_rule weaken #>
+ add_shuffle_rule equiv_comm #>
+ add_shuffle_rule imp_comm #>
+ add_shuffle_rule Drule.norm_hhf_eq #>
+ add_shuffle_rule Drule.triv_forall_equality #>
+ Attrib.setup @{binding shuffle_rule} (Scan.succeed shuffle_attr) "declare rule for shuffler";
+
+end
--- a/src/HOL/Import/HOL_Light/Compatibility.thy Sun Apr 01 09:12:03 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,350 +0,0 @@
-(* Title: HOL/Import/HOL_Light/Compatibility.thy
- Author: Steven Obua and Sebastian Skalberg, TU Muenchen
- Author: Cezary Kaliszyk
-*)
-
-theory Compatibility
-imports Main Fact Parity "~~/src/HOL/Library/Infinite_Set"
- HOLLightList HOLLightReal HOLLightInt "~~/src/HOL/Import/Importer"
-begin
-
-(* list *)
-lemmas [import_rew] = list_el_def member_def list_mem_def
-(* int *)
-lemmas [import_rew] = int_coprime.simps int_gcd.simps hl_mod_def hl_div_def int_mod_def eqeq_def
-(* real *)
-lemma [import_rew]:
- "real (0::nat) = 0" "real (1::nat) = 1" "real (2::nat) = 2"
- by simp_all
-
-lemma one:
- "\<forall>v. v = ()"
- by simp
-
-lemma num_Axiom:
- "\<exists>!fn. fn 0 = e \<and> (\<forall>n. fn (Suc n) = f (fn n) n)"
- apply (rule ex1I[of _ "nat_rec e (%n e. f e n)"])
- apply (auto simp add: fun_eq_iff)
- apply (induct_tac x)
- apply simp_all
- done
-
-lemma SUC_INJ:
- "\<forall>m n. Suc m = Suc n \<longleftrightarrow> m = n"
- by simp
-
-lemma PAIR:
- "(fst x, snd x) = x"
- by simp
-
-lemma EXISTS_UNIQUE_THM:
- "(Ex1 P) = (Ex P & (\<forall>x y. P x & P y --> (x = y)))"
- by auto
-
-lemma DEF__star_:
- "op * = (SOME mult. (\<forall>n. mult 0 n = 0) \<and> (\<forall>m n. mult (Suc m) n = mult m n + n))"
- apply (rule some_equality[symmetric])
- apply (auto simp add: fun_eq_iff)
- apply (induct_tac x)
- apply auto
- done
-
-lemma DEF__slash__backslash_:
- "(t1 \<and> t2) = ((\<lambda>f. f t1 t2 :: bool) = (\<lambda>f. f True True))"
- unfolding fun_eq_iff
- by (intro iffI, simp_all) (erule allE[of _ "(%a b. a \<and> b)"], simp)
-
-lemma DEF__lessthan__equal_:
- "op \<le> = (SOME u. (\<forall>m. u m 0 = (m = 0)) \<and> (\<forall>m n. u m (Suc n) = (m = Suc n \<or> u m n)))"
- apply (rule some_equality[symmetric])
- apply auto[1]
- apply (simp add: fun_eq_iff)
- apply (intro allI)
- apply (induct_tac xa)
- apply auto
- done
-
-lemma DEF__lessthan_:
- "op < = (SOME u. (\<forall>m. u m 0 = False) \<and> (\<forall>m n. u m (Suc n) = (m = n \<or> u m n)))"
- apply (rule some_equality[symmetric])
- apply auto[1]
- apply (simp add: fun_eq_iff)
- apply (intro allI)
- apply (induct_tac xa)
- apply auto
- done
-
-lemma DEF__greaterthan__equal_:
- "(op \<ge>) = (%u ua. ua \<le> u)"
- by (simp)
-
-lemma DEF__greaterthan_:
- "(op >) = (%u ua. ua < u)"
- by (simp)
-
-lemma DEF__equal__equal__greaterthan_:
- "(t1 \<longrightarrow> t2) = ((t1 \<and> t2) = t1)"
- by auto
-
-lemma DEF_WF:
- "wfP = (\<lambda>u. \<forall>P. (\<exists>x. P x) \<longrightarrow> (\<exists>x. P x \<and> (\<forall>y. u y x \<longrightarrow> \<not> P y)))"
- unfolding fun_eq_iff
-proof (intro allI iffI impI wfI_min[to_pred], elim exE wfE_min[to_pred])
- fix P :: "'a \<Rightarrow> bool" and xa :: "'a"
- assume "P xa"
- then show "xa \<in> Collect P" by simp
-next
- fix x P xa z
- assume "P xa" "z \<in> {a. P a}" "\<And>y. x y z \<Longrightarrow> y \<notin> {a. P a}"
- then show "\<exists>xa. P xa \<and> (\<forall>y. x y xa \<longrightarrow> \<not> P y)" by auto
-next
- fix x :: "'a \<Rightarrow> 'a \<Rightarrow> bool" and xa :: "'a" and Q
- assume a: "xa \<in> Q"
- assume b: "\<forall>P. Ex P \<longrightarrow> (\<exists>xa. P xa \<and> (\<forall>y. x y xa \<longrightarrow> \<not> P y))"
- then have "Ex (\<lambda>x. x \<in> Q) \<longrightarrow> (\<exists>xa. (\<lambda>x. x \<in> Q) xa \<and> (\<forall>y. x y xa \<longrightarrow> \<not> (\<lambda>x. x \<in> Q) y))" by auto
- then show "\<exists>z\<in>Q. \<forall>y. x y z \<longrightarrow> y \<notin> Q" using a by auto
-qed
-
-lemma DEF_UNIV: "top = (%x. True)"
- by (rule ext) (simp add: top1I)
-
-lemma DEF_UNIONS:
- "Sup = (\<lambda>u. {ua. \<exists>x. (\<exists>ua. ua \<in> u \<and> x \<in> ua) \<and> ua = x})"
- by (auto simp add: Union_eq)
-
-lemma DEF_UNION:
- "op \<union> = (\<lambda>u ua. {ub. \<exists>x. (x \<in> u \<or> x \<in> ua) \<and> ub = x})"
- by auto
-
-lemma DEF_SUBSET: "op \<subseteq> = (\<lambda>u ua. \<forall>x. x \<in> u \<longrightarrow> x \<in> ua)"
- by (metis set_rev_mp subsetI)
-
-lemma DEF_SND:
- "snd = (\<lambda>p. SOME x. EX y. p = (y, x))"
- unfolding fun_eq_iff
- by (rule someI2) (auto intro: snd_conv[symmetric] someI2)
-
-definition [simp, import_rew]: "SETSPEC x P y \<longleftrightarrow> P & x = y"
-
-lemma DEF_PSUBSET: "op \<subset> = (\<lambda>u ua. u \<subseteq> ua & u \<noteq> ua)"
- by (metis psubset_eq)
-
-definition [import_rew]: "Pred n = n - (Suc 0)"
-
-lemma DEF_PRE: "Pred = (SOME PRE. PRE 0 = 0 & (\<forall>n. PRE (Suc n) = n))"
- apply (rule some_equality[symmetric])
- apply (simp add: Pred_def)
- apply (rule ext)
- apply (induct_tac x)
- apply (auto simp add: Pred_def)
- done
-
-lemma DEF_ONE_ONE:
- "inj = (\<lambda>u. \<forall>x1 x2. u x1 = u x2 \<longrightarrow> x1 = x2)"
- by (simp add: inj_on_def)
-
-definition ODD'[import_rew]: "(ODD :: nat \<Rightarrow> bool) = odd"
-
-lemma DEF_ODD:
- "odd = (SOME ODD. ODD 0 = False \<and> (\<forall>n. ODD (Suc n) = (\<not> ODD n)))"
- apply (rule some_equality[symmetric])
- apply simp
- unfolding fun_eq_iff
- apply (intro allI)
- apply (induct_tac x)
- apply simp_all
- done
-
-definition [import_rew, simp]: "NUMERAL (x :: nat) = x"
-
-lemma DEF_MOD:
- "op mod = (SOME r. \<forall>m n. if n = (0 :: nat) then m div n = 0 \<and>
- r m n = m else m = m div n * n + r m n \<and> r m n < n)"
- apply (rule some_equality[symmetric])
- apply (auto simp add: fun_eq_iff)
- apply (case_tac "xa = 0")
- apply auto
- apply (drule_tac x="x" in spec)
- apply (drule_tac x="xa" in spec)
- apply auto
- by (metis mod_less mod_mult_self2 nat_add_commute nat_mult_commute)
-
-definition "MEASURE = (%u x y. (u x :: nat) < u y)"
-
-lemma [import_rew]:
- "MEASURE u = (%a b. (a, b) \<in> measure u)"
- unfolding MEASURE_def measure_def
- by simp
-
-definition
- "LET f s = f s"
-
-lemma [import_rew]:
- "LET f s = Let s f"
- by (simp add: LET_def Let_def)
-
-lemma DEF_INTERS:
- "Inter = (\<lambda>u. {ua. \<exists>x. (\<forall>ua. ua \<in> u \<longrightarrow> x \<in> ua) \<and> ua = x})"
- by auto
-
-lemma DEF_INTER:
- "op \<inter> = (\<lambda>u ua. {ub. \<exists>x. (x \<in> u \<and> x \<in> ua) \<and> ub = x})"
- by auto
-
-lemma DEF_INSERT:
- "insert = (\<lambda>u ua. {y. y \<in> ua | y = u})"
- by auto
-
-lemma DEF_IMAGE:
- "op ` = (\<lambda>u ua. {ub. \<exists>y. (\<exists>x. x \<in> ua \<and> y = u x) \<and> ub = y})"
- by (simp add: fun_eq_iff image_def Bex_def)
-
-lemma DEF_GEQ:
- "(op =) = (op =)"
- by simp
-
-lemma DEF_GABS:
- "Eps = Eps"
- by simp
-
-lemma DEF_FST:
- "fst = (%p. SOME x. EX y. p = (x, y))"
- unfolding fun_eq_iff
- by (rule someI2) (auto intro: fst_conv[symmetric] someI2)
-
-lemma DEF_FINITE:
- "finite = (\<lambda>a. \<forall>FP. (\<forall>a. a = {} \<or> (\<exists>x s. a = insert x s \<and> FP s) \<longrightarrow> FP a) \<longrightarrow> FP a)"
- unfolding fun_eq_iff
- apply (intro allI iffI impI)
- apply (erule finite_induct)
- apply auto[2]
- apply (drule_tac x="finite" in spec)
- by (metis finite_insert infinite_imp_nonempty infinite_super predicate1I)
-
-lemma DEF_FACT:
- "fact = (SOME FACT. FACT 0 = 1 & (\<forall>n. FACT (Suc n) = Suc n * FACT n))"
- apply (rule some_equality[symmetric])
- apply (simp_all)
- unfolding fun_eq_iff
- apply (intro allI)
- apply (induct_tac x)
- apply simp_all
- done
-
-lemma DEF_EXP:
- "power = (SOME EXP. (\<forall>m. EXP m 0 = 1) \<and> (\<forall>m n. EXP m (Suc n) = m * EXP m n))"
- apply (rule some_equality[symmetric])
- apply (simp_all)
- unfolding fun_eq_iff
- apply (intro allI)
- apply (induct_tac xa)
- apply simp_all
- done
-
-lemma DEF_EVEN:
- "even = (SOME EVEN. EVEN 0 = True \<and> (\<forall>n. EVEN (Suc n) = (\<not> EVEN n)))"
- apply (rule some_equality[symmetric])
- apply simp
- unfolding fun_eq_iff
- apply (intro allI)
- apply (induct_tac x)
- apply simp_all
- done
-
-lemma DEF_EMPTY: "bot = (%x. False)"
- by (rule ext) auto
-
-lemma DEF_DIV:
- "op div = (SOME q. \<exists>r. \<forall>m n. if n = (0 :: nat) then q m n = 0 \<and> r m n = m
- else m = q m n * n + r m n \<and> r m n < n)"
- apply (rule some_equality[symmetric])
- apply (rule_tac x="op mod" in exI)
- apply (auto simp add: fun_eq_iff)
- apply (case_tac "xa = 0")
- apply auto
- apply (drule_tac x="x" in spec)
- apply (drule_tac x="xa" in spec)
- apply auto
- by (metis div_mult_self2 gr_implies_not0 mod_div_trivial mod_less
- nat_add_commute nat_mult_commute plus_nat.add_0)
-
-definition [import_rew]: "DISJOINT a b \<longleftrightarrow> a \<inter> b = {}"
-
-lemma DEF_DISJOINT:
- "DISJOINT = (%u ua. u \<inter> ua = {})"
- by (auto simp add: DISJOINT_def [abs_def])
-
-lemma DEF_DIFF:
- "op - = (\<lambda>u ua. {ub. \<exists>x. (x \<in> u \<and> x \<notin> ua) \<and> ub = x})"
- by (metis set_diff_eq)
-
-definition [import_rew]: "DELETE s e = s - {e}"
-
-lemma DEF_DELETE:
- "DELETE = (\<lambda>u ua. {ub. \<exists>y. (y \<in> u \<and> y \<noteq> ua) \<and> ub = y})"
- by (auto simp add: DELETE_def [abs_def])
-
-lemma COND_DEF:
- "(if b then t else f) = (SOME x. (b = True \<longrightarrow> x = t) \<and> (b = False \<longrightarrow> x = f))"
- by auto
-
-definition [simp]: "NUMERAL_BIT1 n = n + (n + Suc 0)"
-
-lemma BIT1_DEF:
- "NUMERAL_BIT1 = (%u. Suc (2 * u))"
- by (auto)
-
-definition [simp]: "NUMERAL_BIT0 (n :: nat) = n + n"
-
-lemma BIT0_DEF:
- "NUMERAL_BIT0 = (SOME BIT0. BIT0 0 = 0 \<and> (\<forall>n. BIT0 (Suc n) = Suc (Suc (BIT0 n))))"
- apply (rule some_equality[symmetric])
- apply auto
- apply (rule ext)
- apply (induct_tac x)
- apply auto
- done
-
-lemma [import_rew]:
- "NUMERAL_BIT0 n = 2 * n"
- "NUMERAL_BIT1 n = 2 * n + 1"
- "2 * 0 = (0 :: nat)"
- "2 * 1 = (2 :: nat)"
- "0 + 1 = (1 :: nat)"
- by simp_all
-
-lemma DEF_MINUS: "op - = (SOME sub. (\<forall>m. sub m 0 = m) & (\<forall>m n. sub m (Suc n) = sub m n - Suc 0))"
- apply (rule some_equality[symmetric])
- apply auto
- apply (rule ext)+
- apply (induct_tac xa)
- apply auto
- done
-
-lemma DEF_PLUS: "op + = (SOME add. (\<forall>n. add 0 n = n) & (\<forall>m n. add (Suc m) n = Suc (add m n)))"
- apply (rule some_equality[symmetric])
- apply auto
- apply (rule ext)+
- apply (induct_tac x)
- apply auto
- done
-
-lemmas [import_rew] = id_apply
-
-lemma DEF_CHOICE: "Eps = (%u. SOME x. u x)"
- by simp
-
-definition dotdot :: "nat => nat => nat set"
- where "dotdot u ua = {ub. EX x. (u <= x & x <= ua) & ub = x}"
-
-lemma [import_rew]: "dotdot a b = {a..b}"
- unfolding fun_eq_iff atLeastAtMost_def atLeast_def atMost_def dotdot_def
- by (simp add: Collect_conj_eq)
-
-definition [import_rew,simp]: "INFINITE S \<longleftrightarrow> \<not> finite S"
-
-lemma DEF_INFINITE: "INFINITE = (\<lambda>u. \<not>finite u)"
- by (simp add: INFINITE_def [abs_def])
-
-end
-
--- a/src/HOL/Import/HOL_Light/Generate.thy Sun Apr 01 09:12:03 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,5 +0,0 @@
-theory Generate
-imports "Template/GenHOLLight"
-begin
-
-end
--- a/src/HOL/Import/HOL_Light/Generated/HOLLight.thy Sun Apr 01 09:12:03 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,7472 +0,0 @@
-(* AUTOMATICALLY GENERATED, DO NOT EDIT! *)
-
-theory HOLLight
-imports "../../Importer" "../Compatibility"
-begin
-
-setup_theory "~~/src/HOL/Import/HOL_Light/Generated" hollight
-
-consts
- "_FALSITY_" :: "bool" ("'_FALSITY'_")
-
-defs
- "_FALSITY__def": "_FALSITY_ == False"
-
-lemma DEF__FALSITY_: "_FALSITY_ = False"
- by (import hollight DEF__FALSITY_)
-
-lemma CONJ_ACI: "(p & q) = (q & p) &
-((p & q) & r) = (p & q & r) &
-(p & q & r) = (q & p & r) & (p & p) = p & (p & p & q) = (p & q)"
- by (import hollight CONJ_ACI)
-
-lemma DISJ_ACI: "(p | q) = (q | p) &
-((p | q) | r) = (p | q | r) &
-(p | q | r) = (q | p | r) & (p | p) = p & (p | p | q) = (p | q)"
- by (import hollight DISJ_ACI)
-
-lemma IMP_CONJ_ALT: "(p & q --> r) = (q --> p --> r)"
- by (import hollight IMP_CONJ_ALT)
-
-lemma EQ_CLAUSES: "(True = t) = t & (t = True) = t & (False = t) = (~ t) & (t = False) = (~ t)"
- by (import hollight EQ_CLAUSES)
-
-lemma NOT_CLAUSES_WEAK: "(~ True) = False & (~ False) = True"
- by (import hollight NOT_CLAUSES_WEAK)
-
-lemma AND_CLAUSES: "(True & t) = t &
-(t & True) = t & (False & t) = False & (t & False) = False & (t & t) = t"
- by (import hollight AND_CLAUSES)
-
-lemma OR_CLAUSES: "(True | t) = True &
-(t | True) = True & (False | t) = t & (t | False) = t & (t | t) = t"
- by (import hollight OR_CLAUSES)
-
-lemma IMP_CLAUSES: "(True --> t) = t &
-(t --> True) = True &
-(False --> t) = True & (t --> t) = True & (t --> False) = (~ t)"
- by (import hollight IMP_CLAUSES)
-
-lemma IMP_EQ_CLAUSE: "((x::'q_851) = x --> (p::bool)) = p"
- by (import hollight IMP_EQ_CLAUSE)
-
-lemma TRIV_EXISTS_AND_THM: "(EX x::'A. (P::bool) & (Q::bool)) = ((EX x::'A. P) & (EX x::'A. Q))"
- by (import hollight TRIV_EXISTS_AND_THM)
-
-lemma TRIV_AND_EXISTS_THM: "((EX x::'A. (P::bool)) & (EX x::'A. (Q::bool))) = (EX x::'A. P & Q)"
- by (import hollight TRIV_AND_EXISTS_THM)
-
-lemma TRIV_FORALL_OR_THM: "(ALL x::'A. (P::bool) | (Q::bool)) = ((ALL x::'A. P) | (ALL x::'A. Q))"
- by (import hollight TRIV_FORALL_OR_THM)
-
-lemma TRIV_OR_FORALL_THM: "((ALL x::'A. (P::bool)) | (ALL x::'A. (Q::bool))) = (ALL x::'A. P | Q)"
- by (import hollight TRIV_OR_FORALL_THM)
-
-lemma TRIV_FORALL_IMP_THM: "(ALL x::'A. (P::bool) --> (Q::bool)) = ((EX x::'A. P) --> (ALL x::'A. Q))"
- by (import hollight TRIV_FORALL_IMP_THM)
-
-lemma TRIV_EXISTS_IMP_THM: "(EX x::'A. (P::bool) --> (Q::bool)) = ((ALL x::'A. P) --> (EX x::'A. Q))"
- by (import hollight TRIV_EXISTS_IMP_THM)
-
-lemma EXISTS_UNIQUE_ALT: "Ex1 (P::'A => bool) = (EX x::'A. ALL y::'A. P y = (x = y))"
- by (import hollight EXISTS_UNIQUE_ALT)
-
-lemma SELECT_UNIQUE: "(!!y::'A. (P::'A => bool) y = (y = (x::'A))) ==> Eps P = x"
- by (import hollight SELECT_UNIQUE)
-
-lemma EXCLUDED_MIDDLE: "t | ~ t"
- by (import hollight EXCLUDED_MIDDLE)
-
-lemma COND_CLAUSES: "(if True then x::'A else (xa::'A)) = x & (if False then x else xa) = xa"
- by (import hollight COND_CLAUSES)
-
-lemma COND_EXPAND: "(if b then t1 else t2) = ((~ b | t1) & (b | t2))"
- by (import hollight COND_EXPAND)
-
-lemma COND_RATOR: "(if b::bool then f::'A => 'B else (g::'A => 'B)) (x::'A) =
-(if b then f x else g x)"
- by (import hollight COND_RATOR)
-
-lemma COND_ABS: "(%x::'A. if b::bool then (f::'A => 'B) x else (g::'A => 'B) x) =
-(if b then f else g)"
- by (import hollight COND_ABS)
-
-lemma MONO_COND: "[| (A --> B) & (C --> D); if b then A else C |] ==> if b then B else D"
- by (import hollight MONO_COND)
-
-lemma SKOLEM_THM: "(ALL x::'A. Ex ((P::'A => 'B => bool) x)) =
-(EX x::'A => 'B. ALL xa::'A. P xa (x xa))"
- by (import hollight SKOLEM_THM)
-
-lemma UNIQUE_SKOLEM_ALT: "(ALL x::'A. Ex1 ((P::'A => 'B => bool) x)) =
-(EX f::'A => 'B. ALL (x::'A) y::'B. P x y = (f x = y))"
- by (import hollight UNIQUE_SKOLEM_ALT)
-
-lemma COND_EQ_CLAUSE: "(if (x::'q_2963) = x then y::'q_2956 else (z::'q_2956)) = y"
- by (import hollight COND_EQ_CLAUSE)
-
-lemma bool_RECURSION: "EX x::bool => 'A. x False = (a::'A) & x True = (b::'A)"
- by (import hollight bool_RECURSION)
-
-lemma o_ASSOC: "(f::'C => 'D) o ((g::'B => 'C) o (h::'A => 'B)) = f o g o h"
- by (import hollight o_ASSOC)
-
-lemma I_O_ID: "id o (f::'A => 'B) = f & f o id = f"
- by (import hollight I_O_ID)
-
-lemma EXISTS_ONE_REP: "EX x. x"
- by (import hollight EXISTS_ONE_REP)
-
-lemma one_axiom: "(f::'A => unit) = (x::'A => unit)"
- by (import hollight one_axiom)
-
-lemma one_RECURSION: "EX x::unit => 'A. x () = (e::'A)"
- by (import hollight one_RECURSION)
-
-lemma one_Axiom: "EX! fn::unit => 'A. fn () = (e::'A)"
- by (import hollight one_Axiom)
-
-lemma th_cond: "(b = False --> x = x0) & (b = True --> x = x1) ==> x = (b & x1 | ~ b & x0)"
- by (import hollight th_cond)
-
-definition
- LET_END :: "'A => 'A" where
- "LET_END == %t::'A. t"
-
-lemma DEF_LET_END: "LET_END = (%t::'A. t)"
- by (import hollight DEF_LET_END)
-
-consts
- "_SEQPATTERN" :: "('q_4007 => 'q_4004 => bool)
-=> ('q_4007 => 'q_4004 => bool) => 'q_4007 => 'q_4004 => bool" ("'_SEQPATTERN")
-
-defs
- "_SEQPATTERN_def": "_SEQPATTERN ==
-%(r::'q_4007 => 'q_4004 => bool) (s::'q_4007 => 'q_4004 => bool) x::'q_4007.
- if Ex (r x) then r x else s x"
-
-lemma DEF__SEQPATTERN: "_SEQPATTERN =
-(%(r::'q_4007 => 'q_4004 => bool) (s::'q_4007 => 'q_4004 => bool)
- x::'q_4007. if Ex (r x) then r x else s x)"
- by (import hollight DEF__SEQPATTERN)
-
-consts
- "_UNGUARDED_PATTERN" :: "bool => bool => bool" ("'_UNGUARDED'_PATTERN")
-
-defs
- "_UNGUARDED_PATTERN_def": "_UNGUARDED_PATTERN == op &"
-
-lemma DEF__UNGUARDED_PATTERN: "_UNGUARDED_PATTERN = op &"
- by (import hollight DEF__UNGUARDED_PATTERN)
-
-consts
- "_GUARDED_PATTERN" :: "bool => bool => bool => bool" ("'_GUARDED'_PATTERN")
-
-defs
- "_GUARDED_PATTERN_def": "_GUARDED_PATTERN == %p g r. p & g & r"
-
-lemma DEF__GUARDED_PATTERN: "_GUARDED_PATTERN = (%p g r. p & g & r)"
- by (import hollight DEF__GUARDED_PATTERN)
-
-consts
- "_MATCH" :: "'q_4049 => ('q_4049 => 'q_4053 => bool) => 'q_4053" ("'_MATCH")
-
-defs
- "_MATCH_def": "_MATCH ==
-%(e::'q_4049) r::'q_4049 => 'q_4053 => bool.
- if Ex1 (r e) then Eps (r e) else SOME z::'q_4053. False"
-
-lemma DEF__MATCH: "_MATCH =
-(%(e::'q_4049) r::'q_4049 => 'q_4053 => bool.
- if Ex1 (r e) then Eps (r e) else SOME z::'q_4053. False)"
- by (import hollight DEF__MATCH)
-
-consts
- "_FUNCTION" :: "('q_4071 => 'q_4075 => bool) => 'q_4071 => 'q_4075" ("'_FUNCTION")
-
-defs
- "_FUNCTION_def": "_FUNCTION ==
-%(r::'q_4071 => 'q_4075 => bool) x::'q_4071.
- if Ex1 (r x) then Eps (r x) else SOME z::'q_4075. False"
-
-lemma DEF__FUNCTION: "_FUNCTION =
-(%(r::'q_4071 => 'q_4075 => bool) x::'q_4071.
- if Ex1 (r x) then Eps (r x) else SOME z::'q_4075. False)"
- by (import hollight DEF__FUNCTION)
-
-lemma PAIR_EXISTS_THM: "EX (x::'A => 'B => bool) (a::'A) b::'B. x = Pair_Rep a b"
- by (import hollight PAIR_EXISTS_THM)
-
-lemma pair_RECURSION: "EX x::'A * 'B => 'C.
- ALL (a0::'A) a1::'B. x (a0, a1) = (PAIR'::'A => 'B => 'C) a0 a1"
- by (import hollight pair_RECURSION)
-
-definition
- UNCURRY :: "('A => 'B => 'C) => 'A * 'B => 'C" where
- "UNCURRY == %(u::'A => 'B => 'C) ua::'A * 'B. u (fst ua) (snd ua)"
-
-lemma DEF_UNCURRY: "UNCURRY = (%(u::'A => 'B => 'C) ua::'A * 'B. u (fst ua) (snd ua))"
- by (import hollight DEF_UNCURRY)
-
-definition
- PASSOC :: "(('A * 'B) * 'C => 'D) => 'A * 'B * 'C => 'D" where
- "PASSOC ==
-%(u::('A * 'B) * 'C => 'D) ua::'A * 'B * 'C.
- u ((fst ua, fst (snd ua)), snd (snd ua))"
-
-lemma DEF_PASSOC: "PASSOC =
-(%(u::('A * 'B) * 'C => 'D) ua::'A * 'B * 'C.
- u ((fst ua, fst (snd ua)), snd (snd ua)))"
- by (import hollight DEF_PASSOC)
-
-lemma LAMBDA_PAIR_THM: "(x::'q_4547 * 'q_4546 => 'q_4539) =
-(SOME f::'q_4547 * 'q_4546 => 'q_4539.
- ALL (xa::'q_4547) y::'q_4546. f (xa, y) = x (xa, y))"
- by (import hollight LAMBDA_PAIR_THM)
-
-lemma FORALL_PAIRED_THM: "All (SOME f::'q_4576 * 'q_4575 => bool.
- ALL (x::'q_4576) y::'q_4575.
- f (x, y) = (P::'q_4576 => 'q_4575 => bool) x y) =
-(ALL x::'q_4576. All (P x))"
- by (import hollight FORALL_PAIRED_THM)
-
-lemma EXISTS_PAIRED_THM: "Ex (SOME f::'q_4612 * 'q_4611 => bool.
- ALL (x::'q_4612) y::'q_4611.
- f (x, y) = (P::'q_4612 => 'q_4611 => bool) x y) =
-(EX x::'q_4612. Ex (P x))"
- by (import hollight EXISTS_PAIRED_THM)
-
-lemma FORALL_TRIPLED_THM: "All (SOME f::'q_4649 * 'q_4648 * 'q_4647 => bool.
- ALL (x::'q_4649) (y::'q_4648) z::'q_4647.
- f (x, y, z) = (P::'q_4649 => 'q_4648 => 'q_4647 => bool) x y z) =
-(ALL (x::'q_4649) y::'q_4648. All (P x y))"
- by (import hollight FORALL_TRIPLED_THM)
-
-lemma EXISTS_TRIPLED_THM: "Ex (SOME f::'q_4695 * 'q_4694 * 'q_4693 => bool.
- ALL (x::'q_4695) (y::'q_4694) z::'q_4693.
- f (x, y, z) = (P::'q_4695 => 'q_4694 => 'q_4693 => bool) x y z) =
-(EX (x::'q_4695) y::'q_4694. Ex (P x y))"
- by (import hollight EXISTS_TRIPLED_THM)
-
-lemma IND_SUC_0_EXISTS: "EX (x::ind => ind) z::ind.
- (ALL (x1::ind) x2::ind. (x x1 = x x2) = (x1 = x2)) &
- (ALL xa::ind. x xa ~= z)"
- by (import hollight IND_SUC_0_EXISTS)
-
-definition
- IND_SUC :: "ind => ind" where
- "IND_SUC ==
-SOME f. EX z. (ALL x1 x2. (f x1 = f x2) = (x1 = x2)) & (ALL x. f x ~= z)"
-
-lemma DEF_IND_SUC: "IND_SUC =
-(SOME f. EX z. (ALL x1 x2. (f x1 = f x2) = (x1 = x2)) & (ALL x. f x ~= z))"
- by (import hollight DEF_IND_SUC)
-
-definition
- IND_0 :: "ind" where
- "IND_0 ==
-SOME z.
- (ALL x1 x2. (IND_SUC x1 = IND_SUC x2) = (x1 = x2)) &
- (ALL x. IND_SUC x ~= z)"
-
-lemma DEF_IND_0: "IND_0 =
-(SOME z.
- (ALL x1 x2. (IND_SUC x1 = IND_SUC x2) = (x1 = x2)) &
- (ALL x. IND_SUC x ~= z))"
- by (import hollight DEF_IND_0)
-
-definition
- NUM_REP :: "ind => bool" where
- "NUM_REP ==
-%a. ALL NUM_REP'.
- (ALL a.
- a = IND_0 | (EX i. a = IND_SUC i & NUM_REP' i) -->
- NUM_REP' a) -->
- NUM_REP' a"
-
-lemma DEF_NUM_REP: "NUM_REP =
-(%a. ALL NUM_REP'.
- (ALL a.
- a = IND_0 | (EX i. a = IND_SUC i & NUM_REP' i) -->
- NUM_REP' a) -->
- NUM_REP' a)"
- by (import hollight DEF_NUM_REP)
-
-lemma num_RECURSION_STD: "EX fn::nat => 'Z.
- fn (0::nat) = (e::'Z) &
- (ALL n::nat. fn (Suc n) = (f::nat => 'Z => 'Z) n (fn n))"
- by (import hollight num_RECURSION_STD)
-
-lemma ADD_CLAUSES: "(ALL x::nat. (0::nat) + x = x) &
-(ALL x::nat. x + (0::nat) = x) &
-(ALL (x::nat) xa::nat. Suc x + xa = Suc (x + xa)) &
-(ALL (x::nat) xa::nat. x + Suc xa = Suc (x + xa))"
- by (import hollight ADD_CLAUSES)
-
-lemma ADD_AC: "(m::nat) + (n::nat) = n + m &
-m + n + (p::nat) = m + (n + p) & m + (n + p) = n + (m + p)"
- by (import hollight ADD_AC)
-
-lemma EQ_ADD_LCANCEL_0: "((m::nat) + (n::nat) = m) = (n = (0::nat))"
- by (import hollight EQ_ADD_LCANCEL_0)
-
-lemma EQ_ADD_RCANCEL_0: "((x::nat) + (xa::nat) = xa) = (x = (0::nat))"
- by (import hollight EQ_ADD_RCANCEL_0)
-
-lemma BIT1: "2 * x + 1 = Suc (x + x)"
- by (import hollight BIT1)
-
-lemma BIT1_THM: "2 * x + 1 = Suc (x + x)"
- by (import hollight BIT1_THM)
-
-lemma TWO: "2 = Suc 1"
- by (import hollight TWO)
-
-lemma MULT_CLAUSES: "(ALL x::nat. (0::nat) * x = (0::nat)) &
-(ALL x::nat. x * (0::nat) = (0::nat)) &
-(ALL x::nat. (1::nat) * x = x) &
-(ALL x::nat. x * (1::nat) = x) &
-(ALL (x::nat) xa::nat. Suc x * xa = x * xa + xa) &
-(ALL (x::nat) xa::nat. x * Suc xa = x + x * xa)"
- by (import hollight MULT_CLAUSES)
-
-lemma MULT_AC: "(m::nat) * (n::nat) = n * m &
-m * n * (p::nat) = m * (n * p) & m * (n * p) = n * (m * p)"
- by (import hollight MULT_AC)
-
-lemma EXP_EQ_1: "((x::nat) ^ (n::nat) = (1::nat)) = (x = (1::nat) | n = (0::nat))"
- by (import hollight EXP_EQ_1)
-
-lemma LT_ANTISYM: "~ ((m::nat) < (n::nat) & n < m)"
- by (import hollight LT_ANTISYM)
-
-lemma LET_ANTISYM: "~ ((m::nat) <= (n::nat) & n < m)"
- by (import hollight LET_ANTISYM)
-
-lemma LTE_ANTISYM: "~ ((x::nat) < (xa::nat) & xa <= x)"
- by (import hollight LTE_ANTISYM)
-
-lemma LT_CASES: "(m::nat) < (n::nat) | n < m | m = n"
- by (import hollight LT_CASES)
-
-lemma LTE_CASES: "(x::nat) < (xa::nat) | xa <= x"
- by (import hollight LTE_CASES)
-
-lemma LE_1: "(ALL x::nat. x ~= (0::nat) --> (0::nat) < x) &
-(ALL x::nat. x ~= (0::nat) --> (1::nat) <= x) &
-(ALL x>0::nat. x ~= (0::nat)) &
-(ALL x>0::nat. (1::nat) <= x) &
-(ALL x>=1::nat. (0::nat) < x) & (ALL x>=1::nat. x ~= (0::nat))"
- by (import hollight LE_1)
-
-lemma LT_EXISTS: "(m < n) = (EX d. n = m + Suc d)"
- by (import hollight LT_EXISTS)
-
-lemma LT_ADD: "((m::nat) < m + (n::nat)) = ((0::nat) < n)"
- by (import hollight LT_ADD)
-
-lemma LT_ADDR: "((xa::nat) < (x::nat) + xa) = ((0::nat) < x)"
- by (import hollight LT_ADDR)
-
-lemma LT_LMULT: "(m::nat) ~= (0::nat) & (n::nat) < (p::nat) ==> m * n < m * p"
- by (import hollight LT_LMULT)
-
-lemma LE_MULT_LCANCEL: "((m::nat) * (n::nat) <= m * (p::nat)) = (m = (0::nat) | n <= p)"
- by (import hollight LE_MULT_LCANCEL)
-
-lemma LE_MULT_RCANCEL: "((x::nat) * (xb::nat) <= (xa::nat) * xb) = (x <= xa | xb = (0::nat))"
- by (import hollight LE_MULT_RCANCEL)
-
-lemma LT_MULT_LCANCEL: "((m::nat) * (n::nat) < m * (p::nat)) = (m ~= (0::nat) & n < p)"
- by (import hollight LT_MULT_LCANCEL)
-
-lemma LT_MULT_RCANCEL: "((x::nat) * (xb::nat) < (xa::nat) * xb) = (x < xa & xb ~= (0::nat))"
- by (import hollight LT_MULT_RCANCEL)
-
-lemma LT_MULT2: "(m::nat) < (n::nat) & (p::nat) < (q::nat) ==> m * p < n * q"
- by (import hollight LT_MULT2)
-
-lemma WLOG_LE: "(ALL (m::nat) n::nat. (P::nat => nat => bool) m n = P n m) &
-(ALL (m::nat) n::nat. m <= n --> P m n)
-==> P (m::nat) (x::nat)"
- by (import hollight WLOG_LE)
-
-lemma WLOG_LT: "(ALL m::nat. (P::nat => nat => bool) m m) &
-(ALL (m::nat) n::nat. P m n = P n m) &
-(ALL (m::nat) n::nat. m < n --> P m n)
-==> P (m::nat) (x::nat)"
- by (import hollight WLOG_LT)
-
-lemma num_WOP: "Ex (P::nat => bool) = (EX n::nat. P n & (ALL m<n. ~ P m))"
- by (import hollight num_WOP)
-
-lemma num_MAX: "(Ex (P::nat => bool) & (EX M::nat. ALL x::nat. P x --> x <= M)) =
-(EX m::nat. P m & (ALL x::nat. P x --> x <= m))"
- by (import hollight num_MAX)
-
-lemma NOT_EVEN: "odd (n::nat) = odd n"
- by (import hollight NOT_EVEN)
-
-lemma NOT_ODD: "(~ odd (n::nat)) = even n"
- by (import hollight NOT_ODD)
-
-lemma EVEN_OR_ODD: "even (n::nat) | odd n"
- by (import hollight EVEN_OR_ODD)
-
-lemma EVEN_AND_ODD: "~ (even (x::nat) & odd x)"
- by (import hollight EVEN_AND_ODD)
-
-lemma EVEN_EXP: "even ((m::nat) ^ (n::nat)) = (even m & n ~= (0::nat))"
- by (import hollight EVEN_EXP)
-
-lemma ODD_MULT: "odd ((m::nat) * (n::nat)) = (odd m & odd n)"
- by (import hollight ODD_MULT)
-
-lemma ODD_EXP: "odd ((m::nat) ^ (n::nat)) = (odd m | n = (0::nat))"
- by (import hollight ODD_EXP)
-
-lemma EVEN_DOUBLE: "even ((2::nat) * (n::nat))"
- by (import hollight EVEN_DOUBLE)
-
-lemma ODD_DOUBLE: "odd (Suc (2 * x))"
- by (import hollight ODD_DOUBLE)
-
-lemma EVEN_EXISTS_LEMMA: "(even n --> (EX m. n = 2 * m)) & (odd n --> (EX m. n = Suc (2 * m)))"
- by (import hollight EVEN_EXISTS_LEMMA)
-
-lemma EVEN_ODD_DECOMPOSITION: "(EX (k::nat) m::nat. odd m & (n::nat) = (2::nat) ^ k * m) = (n ~= (0::nat))"
- by (import hollight EVEN_ODD_DECOMPOSITION)
-
-lemma SUB_0: "(0::nat) - (x::nat) = (0::nat) & x - (0::nat) = x"
- by (import hollight SUB_0)
-
-lemma SUB_PRESUC: "Suc m - n - Suc 0 = m - n"
- by (import hollight SUB_PRESUC)
-
-lemma ADD_SUBR: "(xa::nat) - ((x::nat) + xa) = (0::nat)"
- by (import hollight ADD_SUBR)
-
-lemma EVEN_SUB: "even ((m::nat) - (n::nat)) = (m <= n | even m = even n)"
- by (import hollight EVEN_SUB)
-
-lemma ODD_SUB: "odd ((x::nat) - (xa::nat)) = (xa < x & odd x ~= odd xa)"
- by (import hollight ODD_SUB)
-
-lemma EXP_LT_0: "((0::nat) < (xa::nat) ^ (x::nat)) = (xa ~= (0::nat) | x = (0::nat))"
- by (import hollight EXP_LT_0)
-
-lemma LT_EXP: "((x::nat) ^ (m::nat) < x ^ (n::nat)) =
-((2::nat) <= x & m < n | x = (0::nat) & m ~= (0::nat) & n = (0::nat))"
- by (import hollight LT_EXP)
-
-lemma LE_EXP: "((x::nat) ^ (m::nat) <= x ^ (n::nat)) =
-(if x = (0::nat) then m = (0::nat) --> n = (0::nat)
- else x = (1::nat) | m <= n)"
- by (import hollight LE_EXP)
-
-lemma EQ_EXP: "((x::nat) ^ (m::nat) = x ^ (n::nat)) =
-(if x = (0::nat) then (m = (0::nat)) = (n = (0::nat))
- else x = (1::nat) | m = n)"
- by (import hollight EQ_EXP)
-
-lemma EXP_MONO_LE_IMP: "(x::nat) <= (xa::nat) ==> x ^ (xb::nat) <= xa ^ xb"
- by (import hollight EXP_MONO_LE_IMP)
-
-lemma EXP_MONO_LT_IMP: "(x::nat) < (y::nat) & (n::nat) ~= (0::nat) ==> x ^ n < y ^ n"
- by (import hollight EXP_MONO_LT_IMP)
-
-lemma EXP_MONO_LE: "((x::nat) ^ (n::nat) <= (y::nat) ^ n) = (x <= y | n = (0::nat))"
- by (import hollight EXP_MONO_LE)
-
-lemma EXP_MONO_LT: "((x::nat) ^ (xb::nat) < (xa::nat) ^ xb) = (x < xa & xb ~= (0::nat))"
- by (import hollight EXP_MONO_LT)
-
-lemma EXP_MONO_EQ: "((x::nat) ^ (xb::nat) = (xa::nat) ^ xb) = (x = xa | xb = (0::nat))"
- by (import hollight EXP_MONO_EQ)
-
-lemma DIVMOD_EXIST: "(n::nat) ~= (0::nat) ==> EX (q::nat) r::nat. (m::nat) = q * n + r & r < n"
- by (import hollight DIVMOD_EXIST)
-
-lemma DIVMOD_EXIST_0: "EX (x::nat) xa::nat.
- if (n::nat) = (0::nat) then x = (0::nat) & xa = (m::nat)
- else m = x * n + xa & xa < n"
- by (import hollight DIVMOD_EXIST_0)
-
-lemma DIVISION: "(n::nat) ~= (0::nat) ==> (m::nat) = m div n * n + m mod n & m mod n < n"
- by (import hollight DIVISION)
-
-lemma DIVMOD_UNIQ_LEMMA: "((m::nat) = (q1::nat) * (n::nat) + (r1::nat) & r1 < n) &
-m = (q2::nat) * n + (r2::nat) & r2 < n
-==> q1 = q2 & r1 = r2"
- by (import hollight DIVMOD_UNIQ_LEMMA)
-
-lemma DIVMOD_UNIQ: "(m::nat) = (q::nat) * (n::nat) + (r::nat) & r < n
-==> m div n = q & m mod n = r"
- by (import hollight DIVMOD_UNIQ)
-
-lemma MOD_UNIQ: "(m::nat) = (q::nat) * (n::nat) + (r::nat) & r < n ==> m mod n = r"
- by (import hollight MOD_UNIQ)
-
-lemma DIV_UNIQ: "(m::nat) = (q::nat) * (n::nat) + (r::nat) & r < n ==> m div n = q"
- by (import hollight DIV_UNIQ)
-
-lemma MOD_EQ: "(m::nat) = (n::nat) + (q::nat) * (p::nat) ==> m mod p = n mod p"
- by (import hollight MOD_EQ)
-
-lemma DIV_LE: "(n::nat) ~= (0::nat) ==> (m::nat) div n <= m"
- by (import hollight DIV_LE)
-
-lemma DIV_MUL_LE: "(n::nat) * ((m::nat) div n) <= m"
- by (import hollight DIV_MUL_LE)
-
-lemma MOD_MOD: "(n::nat) * (p::nat) ~= (0::nat) ==> (m::nat) mod (n * p) mod n = m mod n"
- by (import hollight MOD_MOD)
-
-lemma MOD_MOD_REFL: "(n::nat) ~= (0::nat) ==> (m::nat) mod n mod n = m mod n"
- by (import hollight MOD_MOD_REFL)
-
-lemma DIV_MULT2: "(x::nat) * (xb::nat) ~= (0::nat) ==> x * (xa::nat) div (x * xb) = xa div xb"
- by (import hollight DIV_MULT2)
-
-lemma MOD_MULT2: "(x::nat) * (xb::nat) ~= (0::nat)
-==> x * (xa::nat) mod (x * xb) = x * (xa mod xb)"
- by (import hollight MOD_MULT2)
-
-lemma MOD_EXISTS: "(EX q::nat. (m::nat) = (n::nat) * q) =
-(if n = (0::nat) then m = (0::nat) else m mod n = (0::nat))"
- by (import hollight MOD_EXISTS)
-
-lemma LE_RDIV_EQ: "(a::nat) ~= (0::nat) ==> ((n::nat) <= (b::nat) div a) = (a * n <= b)"
- by (import hollight LE_RDIV_EQ)
-
-lemma LE_LDIV_EQ: "(a::nat) ~= (0::nat)
-==> ((b::nat) div a <= (n::nat)) = (b < a * (n + (1::nat)))"
- by (import hollight LE_LDIV_EQ)
-
-lemma LE_LDIV: "(x::nat) ~= (0::nat) & (xa::nat) <= x * (xb::nat) ==> xa div x <= xb"
- by (import hollight LE_LDIV)
-
-lemma DIV_MONO: "(p::nat) ~= (0::nat) & (m::nat) <= (n::nat) ==> m div p <= n div p"
- by (import hollight DIV_MONO)
-
-lemma DIV_MONO_LT: "(p::nat) ~= (0::nat) & (m::nat) + p <= (n::nat) ==> m div p < n div p"
- by (import hollight DIV_MONO_LT)
-
-lemma DIV_EQ_0: "(n::nat) ~= (0::nat) ==> ((m::nat) div n = (0::nat)) = (m < n)"
- by (import hollight DIV_EQ_0)
-
-lemma MOD_EQ_0: "(n::nat) ~= (0::nat)
-==> ((m::nat) mod n = (0::nat)) = (EX q::nat. m = q * n)"
- by (import hollight MOD_EQ_0)
-
-lemma EVEN_MOD: "even (n::nat) = (n mod (2::nat) = (0::nat))"
- by (import hollight EVEN_MOD)
-
-lemma ODD_MOD: "odd (n::nat) = (n mod (2::nat) = (1::nat))"
- by (import hollight ODD_MOD)
-
-lemma MOD_MULT_RMOD: "(n::nat) ~= (0::nat) ==> (m::nat) * ((p::nat) mod n) mod n = m * p mod n"
- by (import hollight MOD_MULT_RMOD)
-
-lemma MOD_MULT_LMOD: "(xa::nat) ~= (0::nat) ==> (x::nat) mod xa * (xb::nat) mod xa = x * xb mod xa"
- by (import hollight MOD_MULT_LMOD)
-
-lemma MOD_MULT_MOD2: "(xa::nat) ~= (0::nat)
-==> (x::nat) mod xa * ((xb::nat) mod xa) mod xa = x * xb mod xa"
- by (import hollight MOD_MULT_MOD2)
-
-lemma MOD_EXP_MOD: "(n::nat) ~= (0::nat) ==> ((m::nat) mod n) ^ (p::nat) mod n = m ^ p mod n"
- by (import hollight MOD_EXP_MOD)
-
-lemma MOD_ADD_MOD: "(n::nat) ~= (0::nat)
-==> ((a::nat) mod n + (b::nat) mod n) mod n = (a + b) mod n"
- by (import hollight MOD_ADD_MOD)
-
-lemma DIV_ADD_MOD: "(n::nat) ~= (0::nat)
-==> (((a::nat) + (b::nat)) mod n = a mod n + b mod n) =
- ((a + b) div n = a div n + b div n)"
- by (import hollight DIV_ADD_MOD)
-
-lemma MOD_LE: "(n::nat) ~= (0::nat) ==> (m::nat) mod n <= m"
- by (import hollight MOD_LE)
-
-lemma DIV_MONO2: "(p::nat) ~= (0::nat) & p <= (m::nat) ==> (n::nat) div m <= n div p"
- by (import hollight DIV_MONO2)
-
-lemma DIV_LE_EXCLUSION: "(b::nat) ~= (0::nat) & b * (c::nat) < ((a::nat) + (1::nat)) * (d::nat)
-==> c div d <= a div b"
- by (import hollight DIV_LE_EXCLUSION)
-
-lemma DIV_EQ_EXCLUSION: "(b::nat) * (c::nat) < ((a::nat) + (1::nat)) * (d::nat) &
-a * d < (c + (1::nat)) * b
-==> a div b = c div d"
- by (import hollight DIV_EQ_EXCLUSION)
-
-lemma MULT_DIV_LE: "(p::nat) ~= (0::nat) ==> (m::nat) * ((n::nat) div p) <= m * n div p"
- by (import hollight MULT_DIV_LE)
-
-lemma DIV_DIV: "(xa::nat) * (xb::nat) ~= (0::nat)
-==> (x::nat) div xa div xb = x div (xa * xb)"
- by (import hollight DIV_DIV)
-
-lemma DIV_MOD: "(xa::nat) * (xb::nat) ~= (0::nat)
-==> (x::nat) div xa mod xb = x mod (xa * xb) div xa"
- by (import hollight DIV_MOD)
-
-lemma PRE_ELIM_THM: "P (n - Suc 0) = (ALL m. n = Suc m | m = 0 & n = 0 --> P m)"
- by (import hollight PRE_ELIM_THM)
-
-lemma SUB_ELIM_THM: "(P::nat => bool) ((a::nat) - (b::nat)) =
-(ALL d::nat. a = b + d | a < b & d = (0::nat) --> P d)"
- by (import hollight SUB_ELIM_THM)
-
-lemma DIVMOD_ELIM_THM: "(P::nat => nat => bool) ((m::nat) div (n::nat)) (m mod n) =
-(ALL (x::nat) xa::nat.
- n = (0::nat) & x = (0::nat) & xa = m | m = x * n + xa & xa < n -->
- P x xa)"
- by (import hollight DIVMOD_ELIM_THM)
-
-definition
- minimal :: "(nat => bool) => nat" where
- "minimal == %u. SOME n. u n & (ALL m<n. ~ u m)"
-
-lemma DEF_minimal: "minimal = (%u. SOME n. u n & (ALL m<n. ~ u m))"
- by (import hollight DEF_minimal)
-
-lemma MINIMAL: "Ex P = (P (minimal P) & (ALL x<minimal P. ~ P x))"
- by (import hollight MINIMAL)
-
-lemma TRANSITIVE_STEPWISE_LT_EQ: "(!!x y z. R x y & R y z ==> R x z)
-==> (ALL m n. m < n --> R m n) = (ALL n. R n (Suc n))"
- by (import hollight TRANSITIVE_STEPWISE_LT_EQ)
-
-lemma TRANSITIVE_STEPWISE_LT: "[| (ALL x y z. R x y & R y z --> R x z) & (ALL n. R n (Suc n)); m < n |]
-==> R m n"
- by (import hollight TRANSITIVE_STEPWISE_LT)
-
-lemma TRANSITIVE_STEPWISE_LE_EQ: "(ALL x. R x x) & (ALL x y z. R x y & R y z --> R x z)
-==> (ALL m n. m <= n --> R m n) = (ALL n. R n (Suc n))"
- by (import hollight TRANSITIVE_STEPWISE_LE_EQ)
-
-lemma TRANSITIVE_STEPWISE_LE: "[| (ALL x. R x x) &
- (ALL x y z. R x y & R y z --> R x z) & (ALL n. R n (Suc n));
- m <= n |]
-==> R m n"
- by (import hollight TRANSITIVE_STEPWISE_LE)
-
-lemma WF_EQ: "wfP (u_556::'A => 'A => bool) =
-(ALL P::'A => bool.
- Ex P = (EX x::'A. P x & (ALL y::'A. u_556 y x --> ~ P y)))"
- by (import hollight WF_EQ)
-
-lemma WF_IND: "wfP (u_556::'A => 'A => bool) =
-(ALL P::'A => bool.
- (ALL x::'A. (ALL y::'A. u_556 y x --> P y) --> P x) --> All P)"
- by (import hollight WF_IND)
-
-lemma WF_DCHAIN: "wfP (u_556::'A => 'A => bool) =
-(~ (EX s::nat => 'A. ALL n::nat. u_556 (s (Suc n)) (s n)))"
- by (import hollight WF_DCHAIN)
-
-lemma WF_UREC: "[| wfP (u_556::'A => 'A => bool);
- !!(f::'A => 'B) (g::'A => 'B) x::'A.
- (!!z::'A. u_556 z x ==> f z = g z)
- ==> (H::('A => 'B) => 'A => 'B) f x = H g x;
- (ALL x::'A. (f::'A => 'B) x = H f x) &
- (ALL x::'A. (g::'A => 'B) x = H g x) |]
-==> f = g"
- by (import hollight WF_UREC)
-
-lemma WF_UREC_WF: "(!!(H::('A => bool) => 'A => bool) (f::'A => bool) g::'A => bool.
- [| !!(f::'A => bool) (g::'A => bool) x::'A.
- (!!z::'A. (u_556::'A => 'A => bool) z x ==> f z = g z)
- ==> H f x = H g x;
- (ALL x::'A. f x = H f x) & (ALL x::'A. g x = H g x) |]
- ==> f = g)
-==> wfP u_556"
- by (import hollight WF_UREC_WF)
-
-lemma WF_REC_INVARIANT: "[| wfP (u_556::'A => 'A => bool);
- !!(f::'A => 'B) (g::'A => 'B) x::'A.
- (!!z::'A. u_556 z x ==> f z = g z & (S::'A => 'B => bool) z (f z))
- ==> (H::('A => 'B) => 'A => 'B) f x = H g x & S x (H f x) |]
-==> EX f::'A => 'B. ALL x::'A. f x = H f x"
- by (import hollight WF_REC_INVARIANT)
-
-lemma WF_REC: "[| wfP (u_556::'A => 'A => bool);
- !!(f::'A => 'B) (g::'A => 'B) x::'A.
- (!!z::'A. u_556 z x ==> f z = g z)
- ==> (H::('A => 'B) => 'A => 'B) f x = H g x |]
-==> EX f::'A => 'B. ALL x::'A. f x = H f x"
- by (import hollight WF_REC)
-
-lemma WF_REC_WF: "(!!H::('A => nat) => 'A => nat.
- (!!(f::'A => nat) (g::'A => nat) x::'A.
- (!!z::'A. (u_556::'A => 'A => bool) z x ==> f z = g z)
- ==> H f x = H g x)
- ==> EX f::'A => nat. ALL x::'A. f x = H f x)
-==> wfP u_556"
- by (import hollight WF_REC_WF)
-
-lemma WF_EREC: "[| wfP (u_556::'A => 'A => bool);
- !!(f::'A => 'B) (g::'A => 'B) x::'A.
- (!!z::'A. u_556 z x ==> f z = g z)
- ==> (H::('A => 'B) => 'A => 'B) f x = H g x |]
-==> EX! f::'A => 'B. ALL x::'A. f x = H f x"
- by (import hollight WF_EREC)
-
-lemma WF_SUBSET: "(ALL (x::'A) y::'A.
- (u_556::'A => 'A => bool) x y --> (u_670::'A => 'A => bool) x y) &
-wfP u_670
-==> wfP u_556"
- by (import hollight WF_SUBSET)
-
-lemma WF_MEASURE_GEN: "wfP (u_556::'B => 'B => bool)
-==> wfP (%(x::'A) x'::'A. u_556 ((m::'A => 'B) x) (m x'))"
- by (import hollight WF_MEASURE_GEN)
-
-lemma WF_LEX_DEPENDENT: "wfP (R::'A => 'A => bool) & (ALL x::'A. wfP ((S::'A => 'B => 'B => bool) x))
-==> wfP (SOME f::'A * 'B => 'A * 'B => bool.
- ALL (r1::'A) s1::'B.
- f (r1, s1) =
- (SOME f::'A * 'B => bool.
- ALL (r2::'A) s2::'B.
- f (r2, s2) = (R r1 r2 | r1 = r2 & S r1 s1 s2)))"
- by (import hollight WF_LEX_DEPENDENT)
-
-lemma WF_LEX: "wfP (x::'A => 'A => bool) & wfP (xa::'B => 'B => bool)
-==> wfP (SOME f::'A * 'B => 'A * 'B => bool.
- ALL (r1::'A) s1::'B.
- f (r1, s1) =
- (SOME f::'A * 'B => bool.
- ALL (r2::'A) s2::'B.
- f (r2, s2) = (x r1 r2 | r1 = r2 & xa s1 s2)))"
- by (import hollight WF_LEX)
-
-lemma WF_POINTWISE: "wfP (u_556::'A => 'A => bool) & wfP (u_670::'B => 'B => bool)
-==> wfP (SOME f::'A * 'B => 'A * 'B => bool.
- ALL (x1::'A) y1::'B.
- f (x1, y1) =
- (SOME f::'A * 'B => bool.
- ALL (x2::'A) y2::'B.
- f (x2, y2) = (u_556 x1 x2 & u_670 y1 y2)))"
- by (import hollight WF_POINTWISE)
-
-lemma WF_num: "(wfP::(nat => nat => bool) => bool) (op <::nat => nat => bool)"
- by (import hollight WF_num)
-
-lemma WF_REC_num: "(!!(f::nat => 'A) (g::nat => 'A) x::nat.
- (!!z::nat. z < x ==> f z = g z)
- ==> (H::(nat => 'A) => nat => 'A) f x = H g x)
-==> EX f::nat => 'A. ALL x::nat. f x = H f x"
- by (import hollight WF_REC_num)
-
-lemma WF_MEASURE: "wfP (%(a::'A) b::'A. (a, b) : measure (m::'A => nat))"
- by (import hollight WF_MEASURE)
-
-lemma MEASURE_LE: "(ALL x::'q_12099.
- (x, a::'q_12099) : measure (m::'q_12099 => nat) -->
- (x, b::'q_12099) : measure m) =
-(m a <= m b)"
- by (import hollight MEASURE_LE)
-
-lemma WF_REFL: "wfP (u_556::'A => 'A => bool) ==> ~ u_556 (x::'A) x"
- by (import hollight WF_REFL)
-
-lemma WF_REC_TAIL: "EX f::'A => 'B.
- ALL x::'A.
- f x =
- (if (P::'A => bool) x then f ((g::'A => 'A) x) else (h::'A => 'B) x)"
- by (import hollight WF_REC_TAIL)
-
-lemma WF_REC_TAIL_GENERAL: "wfP (u_556::'A => 'A => bool) &
-(ALL (f::'A => 'B) (g::'A => 'B) x::'A.
- (ALL z::'A. u_556 z x --> f z = g z) -->
- (P::('A => 'B) => 'A => bool) f x = P g x &
- (G::('A => 'B) => 'A => 'A) f x = G g x &
- (H::('A => 'B) => 'A => 'B) f x = H g x) &
-(ALL (f::'A => 'B) (g::'A => 'B) x::'A.
- (ALL z::'A. u_556 z x --> f z = g z) --> H f x = H g x) &
-(ALL (f::'A => 'B) (x::'A) y::'A. P f x & u_556 y (G f x) --> u_556 y x)
-==> EX f::'A => 'B. ALL x::'A. f x = (if P f x then f (G f x) else H f x)"
- by (import hollight WF_REC_TAIL_GENERAL)
-
-lemma ARITH_ZERO: "(0::nat) = (0::nat) & (0::nat) = (0::nat)"
- by (import hollight ARITH_ZERO)
-
-lemma ARITH_SUC: "(ALL x. Suc x = Suc x) &
-Suc 0 = 1 &
-(ALL x. Suc (2 * x) = 2 * x + 1) & (ALL x. Suc (2 * x + 1) = 2 * Suc x)"
- by (import hollight ARITH_SUC)
-
-lemma ARITH_PRE: "(ALL x. x - Suc 0 = x - Suc 0) &
-0 - Suc 0 = 0 &
-(ALL x. 2 * x - Suc 0 = (if x = 0 then 0 else 2 * (x - Suc 0) + 1)) &
-(ALL x. 2 * x + 1 - Suc 0 = 2 * x)"
- by (import hollight ARITH_PRE)
-
-lemma ARITH_ADD: "(ALL (x::nat) xa::nat. x + xa = x + xa) &
-(0::nat) + (0::nat) = (0::nat) &
-(ALL x::nat. (0::nat) + (2::nat) * x = (2::nat) * x) &
-(ALL x::nat.
- (0::nat) + ((2::nat) * x + (1::nat)) = (2::nat) * x + (1::nat)) &
-(ALL x::nat. (2::nat) * x + (0::nat) = (2::nat) * x) &
-(ALL x::nat. (2::nat) * x + (1::nat) + (0::nat) = (2::nat) * x + (1::nat)) &
-(ALL (x::nat) xa::nat. (2::nat) * x + (2::nat) * xa = (2::nat) * (x + xa)) &
-(ALL (x::nat) xa::nat.
- (2::nat) * x + ((2::nat) * xa + (1::nat)) =
- (2::nat) * (x + xa) + (1::nat)) &
-(ALL (x::nat) xa::nat.
- (2::nat) * x + (1::nat) + (2::nat) * xa =
- (2::nat) * (x + xa) + (1::nat)) &
-(ALL (x::nat) xa::nat.
- (2::nat) * x + (1::nat) + ((2::nat) * xa + (1::nat)) =
- (2::nat) * Suc (x + xa))"
- by (import hollight ARITH_ADD)
-
-lemma ARITH_MULT: "(ALL (x::nat) xa::nat. x * xa = x * xa) &
-(0::nat) * (0::nat) = (0::nat) &
-(ALL x::nat. (0::nat) * ((2::nat) * x) = (0::nat)) &
-(ALL x::nat. (0::nat) * ((2::nat) * x + (1::nat)) = (0::nat)) &
-(ALL x::nat. (2::nat) * x * (0::nat) = (0::nat)) &
-(ALL x::nat. ((2::nat) * x + (1::nat)) * (0::nat) = (0::nat)) &
-(ALL (x::nat) xa::nat.
- (2::nat) * x * ((2::nat) * xa) = (2::nat) * ((2::nat) * (x * xa))) &
-(ALL (x::nat) xa::nat.
- (2::nat) * x * ((2::nat) * xa + (1::nat)) =
- (2::nat) * x + (2::nat) * ((2::nat) * (x * xa))) &
-(ALL (x::nat) xa::nat.
- ((2::nat) * x + (1::nat)) * ((2::nat) * xa) =
- (2::nat) * xa + (2::nat) * ((2::nat) * (x * xa))) &
-(ALL (x::nat) xa::nat.
- ((2::nat) * x + (1::nat)) * ((2::nat) * xa + (1::nat)) =
- (2::nat) * x + (1::nat) +
- ((2::nat) * xa + (2::nat) * ((2::nat) * (x * xa))))"
- by (import hollight ARITH_MULT)
-
-lemma ARITH_EXP: "(ALL (x::nat) xa::nat. x ^ xa = x ^ xa) &
-(0::nat) ^ (0::nat) = (1::nat) &
-(ALL m::nat. ((2::nat) * m) ^ (0::nat) = (1::nat)) &
-(ALL m::nat. ((2::nat) * m + (1::nat)) ^ (0::nat) = (1::nat)) &
-(ALL n::nat. (0::nat) ^ ((2::nat) * n) = (0::nat) ^ n * (0::nat) ^ n) &
-(ALL (m::nat) n::nat.
- ((2::nat) * m) ^ ((2::nat) * n) =
- ((2::nat) * m) ^ n * ((2::nat) * m) ^ n) &
-(ALL (m::nat) n::nat.
- ((2::nat) * m + (1::nat)) ^ ((2::nat) * n) =
- ((2::nat) * m + (1::nat)) ^ n * ((2::nat) * m + (1::nat)) ^ n) &
-(ALL n::nat. (0::nat) ^ ((2::nat) * n + (1::nat)) = (0::nat)) &
-(ALL (m::nat) n::nat.
- ((2::nat) * m) ^ ((2::nat) * n + (1::nat)) =
- (2::nat) * m * (((2::nat) * m) ^ n * ((2::nat) * m) ^ n)) &
-(ALL (m::nat) n::nat.
- ((2::nat) * m + (1::nat)) ^ ((2::nat) * n + (1::nat)) =
- ((2::nat) * m + (1::nat)) *
- (((2::nat) * m + (1::nat)) ^ n * ((2::nat) * m + (1::nat)) ^ n))"
- by (import hollight ARITH_EXP)
-
-lemma ARITH_EVEN: "(ALL x::nat. even x = even x) &
-even (0::nat) = True &
-(ALL x::nat. even ((2::nat) * x) = True) &
-(ALL x::nat. even ((2::nat) * x + (1::nat)) = False)"
- by (import hollight ARITH_EVEN)
-
-lemma ARITH_ODD: "(ALL x::nat. odd x = odd x) &
-odd (0::nat) = False &
-(ALL x::nat. odd ((2::nat) * x) = False) &
-(ALL x::nat. odd ((2::nat) * x + (1::nat)) = True)"
- by (import hollight ARITH_ODD)
-
-lemma ARITH_LE: "(ALL (x::nat) xa::nat. (x <= xa) = (x <= xa)) &
-((0::nat) <= (0::nat)) = True &
-(ALL x::nat. ((2::nat) * x <= (0::nat)) = (x <= (0::nat))) &
-(ALL x::nat. ((2::nat) * x + (1::nat) <= (0::nat)) = False) &
-(ALL x::nat. ((0::nat) <= (2::nat) * x) = True) &
-(ALL x::nat. ((0::nat) <= (2::nat) * x + (1::nat)) = True) &
-(ALL (x::nat) xa::nat. ((2::nat) * x <= (2::nat) * xa) = (x <= xa)) &
-(ALL (x::nat) xa::nat.
- ((2::nat) * x <= (2::nat) * xa + (1::nat)) = (x <= xa)) &
-(ALL (x::nat) xa::nat.
- ((2::nat) * x + (1::nat) <= (2::nat) * xa) = (x < xa)) &
-(ALL (x::nat) xa::nat.
- ((2::nat) * x + (1::nat) <= (2::nat) * xa + (1::nat)) = (x <= xa))"
- by (import hollight ARITH_LE)
-
-lemma ARITH_LT: "(ALL (x::nat) xa::nat. (x < xa) = (x < xa)) &
-((0::nat) < (0::nat)) = False &
-(ALL x::nat. ((2::nat) * x < (0::nat)) = False) &
-(ALL x::nat. ((2::nat) * x + (1::nat) < (0::nat)) = False) &
-(ALL x::nat. ((0::nat) < (2::nat) * x) = ((0::nat) < x)) &
-(ALL x::nat. ((0::nat) < (2::nat) * x + (1::nat)) = True) &
-(ALL (x::nat) xa::nat. ((2::nat) * x < (2::nat) * xa) = (x < xa)) &
-(ALL (x::nat) xa::nat.
- ((2::nat) * x < (2::nat) * xa + (1::nat)) = (x <= xa)) &
-(ALL (x::nat) xa::nat.
- ((2::nat) * x + (1::nat) < (2::nat) * xa) = (x < xa)) &
-(ALL (x::nat) xa::nat.
- ((2::nat) * x + (1::nat) < (2::nat) * xa + (1::nat)) = (x < xa))"
- by (import hollight ARITH_LT)
-
-lemma ARITH_EQ: "(ALL (x::nat) xa::nat. (x = xa) = (x = xa)) &
-((0::nat) = (0::nat)) = True &
-(ALL x::nat. ((2::nat) * x = (0::nat)) = (x = (0::nat))) &
-(ALL x::nat. ((2::nat) * x + (1::nat) = (0::nat)) = False) &
-(ALL x::nat. ((0::nat) = (2::nat) * x) = ((0::nat) = x)) &
-(ALL x::nat. ((0::nat) = (2::nat) * x + (1::nat)) = False) &
-(ALL (x::nat) xa::nat. ((2::nat) * x = (2::nat) * xa) = (x = xa)) &
-(ALL (x::nat) xa::nat. ((2::nat) * x = (2::nat) * xa + (1::nat)) = False) &
-(ALL (x::nat) xa::nat. ((2::nat) * x + (1::nat) = (2::nat) * xa) = False) &
-(ALL (x::nat) xa::nat.
- ((2::nat) * x + (1::nat) = (2::nat) * xa + (1::nat)) = (x = xa))"
- by (import hollight ARITH_EQ)
-
-lemma ARITH_SUB: "(ALL (x::nat) xa::nat. x - xa = x - xa) &
-(0::nat) - (0::nat) = (0::nat) &
-(ALL x::nat. (0::nat) - (2::nat) * x = (0::nat)) &
-(ALL x::nat. (0::nat) - ((2::nat) * x + (1::nat)) = (0::nat)) &
-(ALL x::nat. (2::nat) * x - (0::nat) = (2::nat) * x) &
-(ALL x::nat. (2::nat) * x + (1::nat) - (0::nat) = (2::nat) * x + (1::nat)) &
-(ALL (m::nat) n::nat. (2::nat) * m - (2::nat) * n = (2::nat) * (m - n)) &
-(ALL (m::nat) n::nat.
- (2::nat) * m - ((2::nat) * n + (1::nat)) =
- (2::nat) * (m - n) - Suc (0::nat)) &
-(ALL (m::nat) n::nat.
- (2::nat) * m + (1::nat) - (2::nat) * n =
- (if n <= m then (2::nat) * (m - n) + (1::nat) else (0::nat))) &
-(ALL (m::nat) n::nat.
- (2::nat) * m + (1::nat) - ((2::nat) * n + (1::nat)) =
- (2::nat) * (m - n))"
- by (import hollight ARITH_SUB)
-
-lemma right_th: "(s::nat) * ((2::nat) * (x::nat) + (1::nat)) = s + (2::nat) * (s * x)"
- by (import hollight right_th)
-
-lemma SEMIRING_PTHS: "(ALL (x::'A) (y::'A) z::'A.
- (add::'A => 'A => 'A) x (add y z) = add (add x y) z) &
-(ALL (x::'A) y::'A. add x y = add y x) &
-(ALL x::'A. add (r0::'A) x = x) &
-(ALL (x::'A) (y::'A) z::'A.
- (mul::'A => 'A => 'A) x (mul y z) = mul (mul x y) z) &
-(ALL (x::'A) y::'A. mul x y = mul y x) &
-(ALL x::'A. mul (r1::'A) x = x) &
-(ALL x::'A. mul r0 x = r0) &
-(ALL (x::'A) (y::'A) z::'A. mul x (add y z) = add (mul x y) (mul x z)) &
-(ALL x::'A. (pwr::'A => nat => 'A) x (0::nat) = r1) &
-(ALL (x::'A) n::nat. pwr x (Suc n) = mul x (pwr x n))
-==> mul r1 (x::'A) = x &
- add (mul (a::'A) (m::'A)) (mul (b::'A) m) = mul (add a b) m &
- add (mul a m) m = mul (add a r1) m &
- add m (mul a m) = mul (add a r1) m &
- add m m = mul (add r1 r1) m &
- mul r0 m = r0 &
- add r0 a = a &
- add a r0 = a &
- mul a b = mul b a &
- mul (add a b) (c::'A) = add (mul a c) (mul b c) &
- mul r0 a = r0 &
- mul a r0 = r0 &
- mul r1 a = a &
- mul a r1 = a &
- mul (mul (lx::'A) (ly::'A)) (mul (rx::'A) (ry::'A)) =
- mul (mul lx rx) (mul ly ry) &
- mul (mul lx ly) (mul rx ry) = mul lx (mul ly (mul rx ry)) &
- mul (mul lx ly) (mul rx ry) = mul rx (mul (mul lx ly) ry) &
- mul (mul lx ly) rx = mul (mul lx rx) ly &
- mul (mul lx ly) rx = mul lx (mul ly rx) &
- mul lx rx = mul rx lx &
- mul lx (mul rx ry) = mul (mul lx rx) ry &
- mul lx (mul rx ry) = mul rx (mul lx ry) &
- add (add a b) (add c (d::'A)) = add (add a c) (add b d) &
- add (add a b) c = add a (add b c) &
- add a (add c d) = add c (add a d) &
- add (add a b) c = add (add a c) b &
- add a c = add c a &
- add a (add c d) = add (add a c) d &
- mul (pwr x (p::nat)) (pwr x (q::nat)) = pwr x (p + q) &
- mul x (pwr x q) = pwr x (Suc q) &
- mul (pwr x q) x = pwr x (Suc q) &
- mul x x = pwr x (2::nat) &
- pwr (mul x (y::'A)) q = mul (pwr x q) (pwr y q) &
- pwr (pwr x p) q = pwr x (p * q) &
- pwr x (0::nat) = r1 &
- pwr x (1::nat) = x &
- mul x (add y (z::'A)) = add (mul x y) (mul x z) &
- pwr x (Suc q) = mul x (pwr x q)"
- by (import hollight SEMIRING_PTHS)
-
-lemma NUM_INTEGRAL_LEMMA: "(w::nat) = (x::nat) + (d::nat) & (y::nat) = (z::nat) + (e::nat)
-==> (w * y + x * z = w * z + x * y) = (w = x | y = z)"
- by (import hollight NUM_INTEGRAL_LEMMA)
-
-lemma NUM_INTEGRAL: "(ALL x::nat. (0::nat) * x = (0::nat)) &
-(ALL (x::nat) (xa::nat) xb::nat. (x + xa = x + xb) = (xa = xb)) &
-(ALL (w::nat) (x::nat) (y::nat) z::nat.
- (w * y + x * z = w * z + x * y) = (w = x | y = z))"
- by (import hollight NUM_INTEGRAL)
-
-lemma INJ_INVERSE2: "(!!(x1::'A) (y1::'B) (x2::'A) y2::'B.
- ((P::'A => 'B => 'C) x1 y1 = P x2 y2) = (x1 = x2 & y1 = y2))
-==> EX (x::'C => 'A) Y::'C => 'B.
- ALL (xa::'A) y::'B. x (P xa y) = xa & Y (P xa y) = y"
- by (import hollight INJ_INVERSE2)
-
-definition
- NUMPAIR :: "nat => nat => nat" where
- "NUMPAIR == %u ua. 2 ^ u * (2 * ua + 1)"
-
-lemma DEF_NUMPAIR: "NUMPAIR = (%u ua. 2 ^ u * (2 * ua + 1))"
- by (import hollight DEF_NUMPAIR)
-
-lemma NUMPAIR_INJ_LEMMA: "NUMPAIR x xa = NUMPAIR xb xc ==> x = xb"
- by (import hollight NUMPAIR_INJ_LEMMA)
-
-lemma NUMPAIR_INJ: "(NUMPAIR x1 y1 = NUMPAIR x2 y2) = (x1 = x2 & y1 = y2)"
- by (import hollight NUMPAIR_INJ)
-
-definition
- NUMFST :: "nat => nat" where
- "NUMFST == SOME X. EX Y. ALL x y. X (NUMPAIR x y) = x & Y (NUMPAIR x y) = y"
-
-lemma DEF_NUMFST: "NUMFST = (SOME X. EX Y. ALL x y. X (NUMPAIR x y) = x & Y (NUMPAIR x y) = y)"
- by (import hollight DEF_NUMFST)
-
-definition
- NUMSND :: "nat => nat" where
- "NUMSND == SOME Y. ALL x y. NUMFST (NUMPAIR x y) = x & Y (NUMPAIR x y) = y"
-
-lemma DEF_NUMSND: "NUMSND = (SOME Y. ALL x y. NUMFST (NUMPAIR x y) = x & Y (NUMPAIR x y) = y)"
- by (import hollight DEF_NUMSND)
-
-definition
- NUMSUM :: "bool => nat => nat" where
- "NUMSUM == %u ua. if u then Suc (2 * ua) else 2 * ua"
-
-lemma DEF_NUMSUM: "NUMSUM = (%u ua. if u then Suc (2 * ua) else 2 * ua)"
- by (import hollight DEF_NUMSUM)
-
-lemma NUMSUM_INJ: "(NUMSUM b1 x1 = NUMSUM b2 x2) = (b1 = b2 & x1 = x2)"
- by (import hollight NUMSUM_INJ)
-
-definition
- NUMLEFT :: "nat => bool" where
- "NUMLEFT == SOME X. EX Y. ALL x y. X (NUMSUM x y) = x & Y (NUMSUM x y) = y"
-
-lemma DEF_NUMLEFT: "NUMLEFT = (SOME X. EX Y. ALL x y. X (NUMSUM x y) = x & Y (NUMSUM x y) = y)"
- by (import hollight DEF_NUMLEFT)
-
-definition
- NUMRIGHT :: "nat => nat" where
- "NUMRIGHT == SOME Y. ALL x y. NUMLEFT (NUMSUM x y) = x & Y (NUMSUM x y) = y"
-
-lemma DEF_NUMRIGHT: "NUMRIGHT = (SOME Y. ALL x y. NUMLEFT (NUMSUM x y) = x & Y (NUMSUM x y) = y)"
- by (import hollight DEF_NUMRIGHT)
-
-definition
- INJN :: "nat => nat => 'A => bool" where
- "INJN == %(u::nat) (n::nat) a::'A. n = u"
-
-lemma DEF_INJN: "INJN = (%(u::nat) (n::nat) a::'A. n = u)"
- by (import hollight DEF_INJN)
-
-lemma INJN_INJ: "(op =::bool => bool => bool)
- ((op =::(nat => 'A::type => bool) => (nat => 'A::type => bool) => bool)
- ((INJN::nat => nat => 'A::type => bool) (n1::nat))
- ((INJN::nat => nat => 'A::type => bool) (n2::nat)))
- ((op =::nat => nat => bool) n1 n2)"
- by (import hollight INJN_INJ)
-
-definition
- INJA :: "'A => nat => 'A => bool" where
- "INJA == %(u::'A) (n::nat) b::'A. b = u"
-
-lemma DEF_INJA: "INJA = (%(u::'A) (n::nat) b::'A. b = u)"
- by (import hollight DEF_INJA)
-
-lemma INJA_INJ: "(INJA (a1::'A) = INJA (a2::'A)) = (a1 = a2)"
- by (import hollight INJA_INJ)
-
-definition
- INJF :: "(nat => nat => 'A => bool) => nat => 'A => bool" where
- "INJF == %(u::nat => nat => 'A => bool) n::nat. u (NUMFST n) (NUMSND n)"
-
-lemma DEF_INJF: "INJF = (%(u::nat => nat => 'A => bool) n::nat. u (NUMFST n) (NUMSND n))"
- by (import hollight DEF_INJF)
-
-lemma INJF_INJ: "(INJF (f1::nat => nat => 'A => bool) =
- INJF (f2::nat => nat => 'A => bool)) =
-(f1 = f2)"
- by (import hollight INJF_INJ)
-
-definition
- INJP :: "(nat => 'A => bool) => (nat => 'A => bool) => nat => 'A => bool" where
- "INJP ==
-%(u::nat => 'A => bool) (ua::nat => 'A => bool) (n::nat) a::'A.
- if NUMLEFT n then u (NUMRIGHT n) a else ua (NUMRIGHT n) a"
-
-lemma DEF_INJP: "INJP =
-(%(u::nat => 'A => bool) (ua::nat => 'A => bool) (n::nat) a::'A.
- if NUMLEFT n then u (NUMRIGHT n) a else ua (NUMRIGHT n) a)"
- by (import hollight DEF_INJP)
-
-lemma INJP_INJ: "(INJP (f1::nat => 'A => bool) (f2::nat => 'A => bool) =
- INJP (f1'::nat => 'A => bool) (f2'::nat => 'A => bool)) =
-(f1 = f1' & f2 = f2')"
- by (import hollight INJP_INJ)
-
-definition
- ZCONSTR :: "nat => 'A => (nat => nat => 'A => bool) => nat => 'A => bool" where
- "ZCONSTR ==
-%(u::nat) (ua::'A) ub::nat => nat => 'A => bool.
- INJP (INJN (Suc u)) (INJP (INJA ua) (INJF ub))"
-
-lemma DEF_ZCONSTR: "ZCONSTR =
-(%(u::nat) (ua::'A) ub::nat => nat => 'A => bool.
- INJP (INJN (Suc u)) (INJP (INJA ua) (INJF ub)))"
- by (import hollight DEF_ZCONSTR)
-
-definition
- ZBOT :: "nat => 'A => bool" where
- "ZBOT == INJP (INJN (0::nat)) (SOME z::nat => 'A => bool. True)"
-
-lemma DEF_ZBOT: "ZBOT = INJP (INJN (0::nat)) (SOME z::nat => 'A => bool. True)"
- by (import hollight DEF_ZBOT)
-
-lemma ZCONSTR_ZBOT: "ZCONSTR (x::nat) (xa::'A) (xb::nat => nat => 'A => bool) ~= ZBOT"
- by (import hollight ZCONSTR_ZBOT)
-
-definition
- ZRECSPACE :: "(nat => 'A => bool) => bool" where
- "ZRECSPACE ==
-%a::nat => 'A => bool.
- ALL ZRECSPACE'::(nat => 'A => bool) => bool.
- (ALL a::nat => 'A => bool.
- a = ZBOT |
- (EX (c::nat) (i::'A) r::nat => nat => 'A => bool.
- a = ZCONSTR c i r & (ALL n::nat. ZRECSPACE' (r n))) -->
- ZRECSPACE' a) -->
- ZRECSPACE' a"
-
-lemma DEF_ZRECSPACE: "ZRECSPACE =
-(%a::nat => 'A => bool.
- ALL ZRECSPACE'::(nat => 'A => bool) => bool.
- (ALL a::nat => 'A => bool.
- a = ZBOT |
- (EX (c::nat) (i::'A) r::nat => nat => 'A => bool.
- a = ZCONSTR c i r & (ALL n::nat. ZRECSPACE' (r n))) -->
- ZRECSPACE' a) -->
- ZRECSPACE' a)"
- by (import hollight DEF_ZRECSPACE)
-
-typedef (open) 'a recspace = "Collect ZRECSPACE :: (nat \<Rightarrow> 'a \<Rightarrow> bool) set"
- morphisms "_dest_rec" "_mk_rec"
- apply (rule light_ex_imp_nonempty [where t="ZBOT"])
- by (import hollight TYDEF_recspace)
-
-syntax
- "_dest_rec" :: _ ("'_dest'_rec")
-
-syntax
- "_mk_rec" :: _ ("'_mk'_rec")
-
-lemmas "TYDEF_recspace_@intern" = typedef_hol2hollight
- [where a="a :: 'A recspace" and r=r ,
- OF type_definition_recspace]
-
-definition
- BOTTOM :: "'A recspace" where
- "BOTTOM == _mk_rec ZBOT"
-
-lemma DEF_BOTTOM: "BOTTOM = _mk_rec ZBOT"
- by (import hollight DEF_BOTTOM)
-
-definition
- CONSTR :: "nat => 'A => (nat => 'A recspace) => 'A recspace" where
- "CONSTR ==
-%(u::nat) (ua::'A::type) ub::nat => 'A::type recspace.
- _mk_rec (ZCONSTR u ua (%n::nat. _dest_rec (ub n)))"
-
-lemma DEF_CONSTR: "CONSTR =
-(%(u::nat) (ua::'A::type) ub::nat => 'A::type recspace.
- _mk_rec (ZCONSTR u ua (%n::nat. _dest_rec (ub n))))"
- by (import hollight DEF_CONSTR)
-
-lemma MK_REC_INJ: "[| _mk_rec (x::nat => 'A::type => bool) =
- _mk_rec (y::nat => 'A::type => bool);
- ZRECSPACE x & ZRECSPACE y |]
-==> x = y"
- by (import hollight MK_REC_INJ)
-
-lemma CONSTR_BOT: "CONSTR (c::nat) (i::'A) (r::nat => 'A recspace) ~= BOTTOM"
- by (import hollight CONSTR_BOT)
-
-lemma CONSTR_INJ: "(CONSTR (c1::nat) (i1::'A) (r1::nat => 'A recspace) =
- CONSTR (c2::nat) (i2::'A) (r2::nat => 'A recspace)) =
-(c1 = c2 & i1 = i2 & r1 = r2)"
- by (import hollight CONSTR_INJ)
-
-lemma CONSTR_IND: "(P::'A recspace => bool) BOTTOM &
-(ALL (c::nat) (i::'A) r::nat => 'A recspace.
- (ALL n::nat. P (r n)) --> P (CONSTR c i r))
-==> P (x::'A recspace)"
- by (import hollight CONSTR_IND)
-
-lemma CONSTR_REC: "EX f::'A recspace => 'B.
- ALL (c::nat) (i::'A) r::nat => 'A recspace.
- f (CONSTR c i r) =
- (Fn::nat => 'A => (nat => 'A recspace) => (nat => 'B) => 'B) c i r
- (%n::nat. f (r n))"
- by (import hollight CONSTR_REC)
-
-definition
- FCONS :: "'A => (nat => 'A) => nat => 'A" where
- "FCONS ==
-SOME FCONS::'A => (nat => 'A) => nat => 'A.
- (ALL (a::'A) f::nat => 'A. FCONS a f (0::nat) = a) &
- (ALL (a::'A) (f::nat => 'A) n::nat. FCONS a f (Suc n) = f n)"
-
-lemma DEF_FCONS: "FCONS =
-(SOME FCONS::'A => (nat => 'A) => nat => 'A.
- (ALL (a::'A) f::nat => 'A. FCONS a f (0::nat) = a) &
- (ALL (a::'A) (f::nat => 'A) n::nat. FCONS a f (Suc n) = f n))"
- by (import hollight DEF_FCONS)
-
-lemma FCONS_UNDO: "(f::nat => 'A) = FCONS (f (0::nat)) (f o Suc)"
- by (import hollight FCONS_UNDO)
-
-definition
- FNIL :: "nat => 'A" where
- "FNIL == %u::nat. SOME x::'A. True"
-
-lemma DEF_FNIL: "FNIL = (%u::nat. SOME x::'A. True)"
- by (import hollight DEF_FNIL)
-
-definition
- ISO :: "('A => 'B) => ('B => 'A) => bool" where
- "ISO ==
-%(u::'A => 'B) ua::'B => 'A.
- (ALL x::'B. u (ua x) = x) & (ALL y::'A. ua (u y) = y)"
-
-lemma DEF_ISO: "ISO =
-(%(u::'A => 'B) ua::'B => 'A.
- (ALL x::'B. u (ua x) = x) & (ALL y::'A. ua (u y) = y))"
- by (import hollight DEF_ISO)
-
-lemma ISO_REFL: "ISO (%x::'A. x) (%x::'A. x)"
- by (import hollight ISO_REFL)
-
-lemma ISO_FUN: "ISO (f::'A => 'A') (f'::'A' => 'A) & ISO (g::'B => 'B') (g'::'B' => 'B)
-==> ISO (%(h::'A => 'B) a'::'A'. g (h (f' a')))
- (%(h::'A' => 'B') a::'A. g' (h (f a)))"
- by (import hollight ISO_FUN)
-
-lemma ISO_USAGE: "ISO (f::'q_17485 => 'q_17482) (g::'q_17482 => 'q_17485)
-==> (ALL P::'q_17485 => bool. All P = (ALL x::'q_17482. P (g x))) &
- (ALL P::'q_17485 => bool. Ex P = (EX x::'q_17482. P (g x))) &
- (ALL (a::'q_17485) b::'q_17482. (a = g b) = (f a = b))"
- by (import hollight ISO_USAGE)
-
-typedef (open) char = "{a. ALL char'.
- (ALL a.
- (EX a0 a1 a2 a3 a4 a5 a6 a7.
- a =
- CONSTR (NUMERAL 0) (a0 :: bool, a1 :: bool, a2 :: bool, a3 :: bool, a4 :: bool, a5 :: bool, a6 :: bool, a7:: bool)
- (%n. BOTTOM)) -->
- char' a) -->
- char' a}"
- morphisms "_dest_char" "_mk_char"
- apply (rule light_ex_imp_nonempty [where t="CONSTR (NUMERAL 0) (a0, a1, a2, a3, a4, a5, a6, a7) (%n. BOTTOM)"])
- by (import hollight TYDEF_char)
-
-syntax
- "_dest_char" :: _ ("'_dest'_char")
-
-syntax
- "_mk_char" :: _ ("'_mk'_char")
-
-lemmas "TYDEF_char_@intern" = typedef_hol2hollight
- [where a="a :: hollight.char" and r=r ,
- OF type_definition_char]
-
-consts
- "_11937" :: "bool
-=> bool => bool => bool => bool => bool => bool => bool => hollight.char" ("'_11937")
-
-defs
- "_11937_def": "_11937 ==
-%(a0::bool) (a1::bool) (a2::bool) (a3::bool) (a4::bool) (a5::bool)
- (a6::bool) a7::bool.
- _mk_char
- (CONSTR (0::nat) (a0, a1, a2, a3, a4, a5, a6, a7) (%n::nat. BOTTOM))"
-
-lemma DEF__11937: "_11937 =
-(%(a0::bool) (a1::bool) (a2::bool) (a3::bool) (a4::bool) (a5::bool)
- (a6::bool) a7::bool.
- _mk_char
- (CONSTR (0::nat) (a0, a1, a2, a3, a4, a5, a6, a7) (%n::nat. BOTTOM)))"
- by (import hollight DEF__11937)
-
-definition
- ASCII :: "bool
-=> bool => bool => bool => bool => bool => bool => bool => hollight.char" where
- "ASCII == _11937"
-
-lemma DEF_ASCII: "ASCII = _11937"
- by (import hollight DEF_ASCII)
-
-consts
- dist :: "nat * nat => nat"
-
-defs
- dist_def: "hollight.dist == %u. fst u - snd u + (snd u - fst u)"
-
-lemma DEF_dist: "hollight.dist = (%u. fst u - snd u + (snd u - fst u))"
- by (import hollight DEF_dist)
-
-lemma DIST_REFL: "hollight.dist (x, x) = 0"
- by (import hollight DIST_REFL)
-
-lemma DIST_LZERO: "hollight.dist (0, x) = x"
- by (import hollight DIST_LZERO)
-
-lemma DIST_RZERO: "hollight.dist (x, 0) = x"
- by (import hollight DIST_RZERO)
-
-lemma DIST_SYM: "hollight.dist (x, xa) = hollight.dist (xa, x)"
- by (import hollight DIST_SYM)
-
-lemma DIST_LADD: "hollight.dist (x + xb, x + xa) = hollight.dist (xb, xa)"
- by (import hollight DIST_LADD)
-
-lemma DIST_RADD: "hollight.dist (x + xa, xb + xa) = hollight.dist (x, xb)"
- by (import hollight DIST_RADD)
-
-lemma DIST_LADD_0: "hollight.dist (x + xa, x) = xa"
- by (import hollight DIST_LADD_0)
-
-lemma DIST_RADD_0: "hollight.dist (x, x + xa) = xa"
- by (import hollight DIST_RADD_0)
-
-lemma DIST_LMUL: "x * hollight.dist (xa, xb) = hollight.dist (x * xa, x * xb)"
- by (import hollight DIST_LMUL)
-
-lemma DIST_RMUL: "hollight.dist (x, xa) * xb = hollight.dist (x * xb, xa * xb)"
- by (import hollight DIST_RMUL)
-
-lemma DIST_EQ_0: "(hollight.dist (x, xa) = 0) = (x = xa)"
- by (import hollight DIST_EQ_0)
-
-lemma DIST_ELIM_THM: "P (hollight.dist (x, y)) =
-(ALL d. (x = y + d --> P d) & (y = x + d --> P d))"
- by (import hollight DIST_ELIM_THM)
-
-lemma DIST_LE_CASES: "(hollight.dist (m, n) <= p) = (m <= n + p & n <= m + p)"
- by (import hollight DIST_LE_CASES)
-
-lemma DIST_TRIANGLE_LE: "hollight.dist (m, n) + hollight.dist (n, p) <= q
-==> hollight.dist (m, p) <= q"
- by (import hollight DIST_TRIANGLE_LE)
-
-lemma DIST_TRIANGLES_LE: "hollight.dist (m, n) <= r & hollight.dist (p, q) <= s
-==> hollight.dist (m, p) <= hollight.dist (n, q) + (r + s)"
- by (import hollight DIST_TRIANGLES_LE)
-
-lemma BOUNDS_LINEAR: "(ALL n::nat. (A::nat) * n <= (B::nat) * n + (C::nat)) = (A <= B)"
- by (import hollight BOUNDS_LINEAR)
-
-lemma BOUNDS_LINEAR_0: "(ALL n::nat. (A::nat) * n <= (B::nat)) = (A = (0::nat))"
- by (import hollight BOUNDS_LINEAR_0)
-
-lemma BOUNDS_DIVIDED: "(EX B::nat. ALL n::nat. (P::nat => nat) n <= B) =
-(EX (x::nat) B::nat. ALL n::nat. n * P n <= x * n + B)"
- by (import hollight BOUNDS_DIVIDED)
-
-lemma BOUNDS_NOTZERO: "(P::nat => nat => nat) (0::nat) (0::nat) = (0::nat) &
-(ALL (m::nat) n::nat. P m n <= (A::nat) * (m + n) + (B::nat))
-==> EX x::nat. ALL (m::nat) n::nat. P m n <= x * (m + n)"
- by (import hollight BOUNDS_NOTZERO)
-
-lemma BOUNDS_IGNORE: "(EX B::nat. ALL i::nat. (P::nat => nat) i <= (Q::nat => nat) i + B) =
-(EX (x::nat) N::nat. ALL i>=N. P i <= Q i + x)"
- by (import hollight BOUNDS_IGNORE)
-
-definition
- is_nadd :: "(nat => nat) => bool" where
- "is_nadd ==
-%u. EX B. ALL m n. hollight.dist (m * u n, n * u m) <= B * (m + n)"
-
-lemma DEF_is_nadd: "is_nadd =
-(%u. EX B. ALL m n. hollight.dist (m * u n, n * u m) <= B * (m + n))"
- by (import hollight DEF_is_nadd)
-
-lemma is_nadd_0: "is_nadd (%n. 0)"
- by (import hollight is_nadd_0)
-
-typedef (open) nadd = "Collect is_nadd" morphisms "dest_nadd" "mk_nadd"
- apply (rule light_ex_imp_nonempty[where t="%n. NUMERAL 0"])
- by (import hollight TYDEF_nadd)
-
-syntax
- dest_nadd :: _
-
-syntax
- mk_nadd :: _
-
-lemmas "TYDEF_nadd_@intern" = typedef_hol2hollight
- [where a="a :: nadd" and r=r ,
- OF type_definition_nadd]
-
-lemma NADD_CAUCHY: "EX xa.
- ALL xb xc.
- hollight.dist (xb * dest_nadd x xc, xc * dest_nadd x xb)
- <= xa * (xb + xc)"
- by (import hollight NADD_CAUCHY)
-
-lemma NADD_BOUND: "EX xa B. ALL n. dest_nadd x n <= xa * n + B"
- by (import hollight NADD_BOUND)
-
-lemma NADD_MULTIPLICATIVE: "EX xa.
- ALL m n.
- hollight.dist (dest_nadd x (m * n), m * dest_nadd x n) <= xa * m + xa"
- by (import hollight NADD_MULTIPLICATIVE)
-
-lemma NADD_ADDITIVE: "EX xa.
- ALL m n.
- hollight.dist (dest_nadd x (m + n), dest_nadd x m + dest_nadd x n)
- <= xa"
- by (import hollight NADD_ADDITIVE)
-
-lemma NADD_SUC: "EX xa. ALL n. hollight.dist (dest_nadd x (Suc n), dest_nadd x n) <= xa"
- by (import hollight NADD_SUC)
-
-lemma NADD_DIST_LEMMA: "EX xa. ALL m n. hollight.dist (dest_nadd x (m + n), dest_nadd x m) <= xa * n"
- by (import hollight NADD_DIST_LEMMA)
-
-lemma NADD_DIST: "EX xa.
- ALL m n.
- hollight.dist (dest_nadd x m, dest_nadd x n)
- <= xa * hollight.dist (m, n)"
- by (import hollight NADD_DIST)
-
-lemma NADD_ALTMUL: "EX A B.
- ALL n.
- hollight.dist
- (n * dest_nadd x (dest_nadd y n), dest_nadd x n * dest_nadd y n)
- <= A * n + B"
- by (import hollight NADD_ALTMUL)
-
-definition
- nadd_eq :: "nadd => nadd => bool" where
- "nadd_eq ==
-%u ua. EX B. ALL n. hollight.dist (dest_nadd u n, dest_nadd ua n) <= B"
-
-lemma DEF_nadd_eq: "nadd_eq =
-(%u ua. EX B. ALL n. hollight.dist (dest_nadd u n, dest_nadd ua n) <= B)"
- by (import hollight DEF_nadd_eq)
-
-lemma NADD_EQ_REFL: "nadd_eq x x"
- by (import hollight NADD_EQ_REFL)
-
-lemma NADD_EQ_SYM: "nadd_eq x y = nadd_eq y x"
- by (import hollight NADD_EQ_SYM)
-
-lemma NADD_EQ_TRANS: "nadd_eq x y & nadd_eq y z ==> nadd_eq x z"
- by (import hollight NADD_EQ_TRANS)
-
-definition
- nadd_of_num :: "nat => nadd" where
- "nadd_of_num == %u. mk_nadd (op * u)"
-
-lemma DEF_nadd_of_num: "nadd_of_num = (%u. mk_nadd (op * u))"
- by (import hollight DEF_nadd_of_num)
-
-lemma NADD_OF_NUM: "dest_nadd (nadd_of_num x) = op * x"
- by (import hollight NADD_OF_NUM)
-
-lemma NADD_OF_NUM_WELLDEF: "m = n ==> nadd_eq (nadd_of_num m) (nadd_of_num n)"
- by (import hollight NADD_OF_NUM_WELLDEF)
-
-lemma NADD_OF_NUM_EQ: "nadd_eq (nadd_of_num m) (nadd_of_num n) = (m = n)"
- by (import hollight NADD_OF_NUM_EQ)
-
-definition
- nadd_le :: "nadd => nadd => bool" where
- "nadd_le == %u ua. EX B. ALL n. dest_nadd u n <= dest_nadd ua n + B"
-
-lemma DEF_nadd_le: "nadd_le = (%u ua. EX B. ALL n. dest_nadd u n <= dest_nadd ua n + B)"
- by (import hollight DEF_nadd_le)
-
-lemma NADD_LE_WELLDEF_LEMMA: "nadd_eq x x' & nadd_eq y y' & nadd_le x y ==> nadd_le x' y'"
- by (import hollight NADD_LE_WELLDEF_LEMMA)
-
-lemma NADD_LE_WELLDEF: "nadd_eq x x' & nadd_eq y y' ==> nadd_le x y = nadd_le x' y'"
- by (import hollight NADD_LE_WELLDEF)
-
-lemma NADD_LE_REFL: "nadd_le x x"
- by (import hollight NADD_LE_REFL)
-
-lemma NADD_LE_TRANS: "nadd_le x y & nadd_le y z ==> nadd_le x z"
- by (import hollight NADD_LE_TRANS)
-
-lemma NADD_LE_ANTISYM: "(nadd_le x y & nadd_le y x) = nadd_eq x y"
- by (import hollight NADD_LE_ANTISYM)
-
-lemma NADD_LE_TOTAL_LEMMA: "~ nadd_le x y ==> EX n. n ~= 0 & dest_nadd y n + B < dest_nadd x n"
- by (import hollight NADD_LE_TOTAL_LEMMA)
-
-lemma NADD_LE_TOTAL: "nadd_le x y | nadd_le y x"
- by (import hollight NADD_LE_TOTAL)
-
-lemma NADD_ARCH: "EX xa. nadd_le x (nadd_of_num xa)"
- by (import hollight NADD_ARCH)
-
-lemma NADD_OF_NUM_LE: "nadd_le (nadd_of_num m) (nadd_of_num n) = (m <= n)"
- by (import hollight NADD_OF_NUM_LE)
-
-definition
- nadd_add :: "nadd => nadd => nadd" where
- "nadd_add == %u ua. mk_nadd (%n. dest_nadd u n + dest_nadd ua n)"
-
-lemma DEF_nadd_add: "nadd_add = (%u ua. mk_nadd (%n. dest_nadd u n + dest_nadd ua n))"
- by (import hollight DEF_nadd_add)
-
-lemma NADD_ADD: "dest_nadd (nadd_add x y) = (%n. dest_nadd x n + dest_nadd y n)"
- by (import hollight NADD_ADD)
-
-lemma NADD_ADD_WELLDEF: "nadd_eq x x' & nadd_eq y y' ==> nadd_eq (nadd_add x y) (nadd_add x' y')"
- by (import hollight NADD_ADD_WELLDEF)
-
-lemma NADD_ADD_SYM: "nadd_eq (nadd_add x y) (nadd_add y x)"
- by (import hollight NADD_ADD_SYM)
-
-lemma NADD_ADD_ASSOC: "nadd_eq (nadd_add x (nadd_add y z)) (nadd_add (nadd_add x y) z)"
- by (import hollight NADD_ADD_ASSOC)
-
-lemma NADD_ADD_LID: "nadd_eq (nadd_add (nadd_of_num 0) x) x"
- by (import hollight NADD_ADD_LID)
-
-lemma NADD_ADD_LCANCEL: "nadd_eq (nadd_add x y) (nadd_add x z) ==> nadd_eq y z"
- by (import hollight NADD_ADD_LCANCEL)
-
-lemma NADD_LE_ADD: "nadd_le x (nadd_add x y)"
- by (import hollight NADD_LE_ADD)
-
-lemma NADD_LE_EXISTS: "nadd_le x y ==> EX d. nadd_eq y (nadd_add x d)"
- by (import hollight NADD_LE_EXISTS)
-
-lemma NADD_OF_NUM_ADD: "nadd_eq (nadd_add (nadd_of_num x) (nadd_of_num xa)) (nadd_of_num (x + xa))"
- by (import hollight NADD_OF_NUM_ADD)
-
-definition
- nadd_mul :: "nadd => nadd => nadd" where
- "nadd_mul == %u ua. mk_nadd (%n. dest_nadd u (dest_nadd ua n))"
-
-lemma DEF_nadd_mul: "nadd_mul = (%u ua. mk_nadd (%n. dest_nadd u (dest_nadd ua n)))"
- by (import hollight DEF_nadd_mul)
-
-lemma NADD_MUL: "dest_nadd (nadd_mul x y) = (%n. dest_nadd x (dest_nadd y n))"
- by (import hollight NADD_MUL)
-
-lemma NADD_MUL_SYM: "nadd_eq (nadd_mul x y) (nadd_mul y x)"
- by (import hollight NADD_MUL_SYM)
-
-lemma NADD_MUL_ASSOC: "nadd_eq (nadd_mul x (nadd_mul y z)) (nadd_mul (nadd_mul x y) z)"
- by (import hollight NADD_MUL_ASSOC)
-
-lemma NADD_MUL_LID: "nadd_eq (nadd_mul (nadd_of_num 1) x) x"
- by (import hollight NADD_MUL_LID)
-
-lemma NADD_LDISTRIB: "nadd_eq (nadd_mul x (nadd_add y z)) (nadd_add (nadd_mul x y) (nadd_mul x z))"
- by (import hollight NADD_LDISTRIB)
-
-lemma NADD_MUL_WELLDEF_LEMMA: "nadd_eq y y' ==> nadd_eq (nadd_mul x y) (nadd_mul x y')"
- by (import hollight NADD_MUL_WELLDEF_LEMMA)
-
-lemma NADD_MUL_WELLDEF: "nadd_eq x x' & nadd_eq y y' ==> nadd_eq (nadd_mul x y) (nadd_mul x' y')"
- by (import hollight NADD_MUL_WELLDEF)
-
-lemma NADD_OF_NUM_MUL: "nadd_eq (nadd_mul (nadd_of_num x) (nadd_of_num xa)) (nadd_of_num (x * xa))"
- by (import hollight NADD_OF_NUM_MUL)
-
-lemma NADD_LE_0: "nadd_le (nadd_of_num 0) x"
- by (import hollight NADD_LE_0)
-
-lemma NADD_EQ_IMP_LE: "nadd_eq x y ==> nadd_le x y"
- by (import hollight NADD_EQ_IMP_LE)
-
-lemma NADD_LE_LMUL: "nadd_le y z ==> nadd_le (nadd_mul x y) (nadd_mul x z)"
- by (import hollight NADD_LE_LMUL)
-
-lemma NADD_LE_RMUL: "nadd_le x y ==> nadd_le (nadd_mul x z) (nadd_mul y z)"
- by (import hollight NADD_LE_RMUL)
-
-lemma NADD_LE_RADD: "nadd_le (nadd_add x z) (nadd_add y z) = nadd_le x y"
- by (import hollight NADD_LE_RADD)
-
-lemma NADD_LE_LADD: "nadd_le (nadd_add x y) (nadd_add x z) = nadd_le y z"
- by (import hollight NADD_LE_LADD)
-
-lemma NADD_RDISTRIB: "nadd_eq (nadd_mul (nadd_add x y) z) (nadd_add (nadd_mul x z) (nadd_mul y z))"
- by (import hollight NADD_RDISTRIB)
-
-lemma NADD_ARCH_MULT: "~ nadd_eq x (nadd_of_num 0)
-==> EX xa. nadd_le (nadd_of_num k) (nadd_mul (nadd_of_num xa) x)"
- by (import hollight NADD_ARCH_MULT)
-
-lemma NADD_ARCH_ZERO: "(!!n. nadd_le (nadd_mul (nadd_of_num n) x) k) ==> nadd_eq x (nadd_of_num 0)"
- by (import hollight NADD_ARCH_ZERO)
-
-lemma NADD_ARCH_LEMMA: "(!!n. nadd_le (nadd_mul (nadd_of_num n) x)
- (nadd_add (nadd_mul (nadd_of_num n) y) z))
-==> nadd_le x y"
- by (import hollight NADD_ARCH_LEMMA)
-
-lemma NADD_COMPLETE: "Ex P & (EX M. ALL x. P x --> nadd_le x M)
-==> EX M. (ALL x. P x --> nadd_le x M) &
- (ALL M'. (ALL x. P x --> nadd_le x M') --> nadd_le M M')"
- by (import hollight NADD_COMPLETE)
-
-lemma NADD_UBOUND: "EX xa N. ALL n>=N. dest_nadd x n <= xa * n"
- by (import hollight NADD_UBOUND)
-
-lemma NADD_NONZERO: "~ nadd_eq x (nadd_of_num 0) ==> EX N. ALL n>=N. dest_nadd x n ~= 0"
- by (import hollight NADD_NONZERO)
-
-lemma NADD_LBOUND: "~ nadd_eq x (nadd_of_num 0) ==> EX A N. ALL n>=N. n <= A * dest_nadd x n"
- by (import hollight NADD_LBOUND)
-
-definition
- nadd_rinv :: "nadd => nat => nat" where
- "nadd_rinv == %u n. n * n div dest_nadd u n"
-
-lemma DEF_nadd_rinv: "nadd_rinv = (%u n. n * n div dest_nadd u n)"
- by (import hollight DEF_nadd_rinv)
-
-lemma NADD_MUL_LINV_LEMMA0: "~ nadd_eq x (nadd_of_num 0) ==> EX xa B. ALL i. nadd_rinv x i <= xa * i + B"
- by (import hollight NADD_MUL_LINV_LEMMA0)
-
-lemma NADD_MUL_LINV_LEMMA1: "dest_nadd x n ~= 0
-==> hollight.dist (dest_nadd x n * nadd_rinv x n, n * n) <= dest_nadd x n"
- by (import hollight NADD_MUL_LINV_LEMMA1)
-
-lemma NADD_MUL_LINV_LEMMA2: "~ nadd_eq x (nadd_of_num 0)
-==> EX N. ALL n>=N.
- hollight.dist (dest_nadd x n * nadd_rinv x n, n * n)
- <= dest_nadd x n"
- by (import hollight NADD_MUL_LINV_LEMMA2)
-
-lemma NADD_MUL_LINV_LEMMA3: "~ nadd_eq x (nadd_of_num 0)
-==> EX N. ALL m n.
- N <= n -->
- hollight.dist
- (m * (dest_nadd x m * (dest_nadd x n * nadd_rinv x n)),
- m * (dest_nadd x m * (n * n)))
- <= m * (dest_nadd x m * dest_nadd x n)"
- by (import hollight NADD_MUL_LINV_LEMMA3)
-
-lemma NADD_MUL_LINV_LEMMA4: "~ nadd_eq x (nadd_of_num 0)
-==> EX N. ALL m n.
- N <= m & N <= n -->
- dest_nadd x m * dest_nadd x n *
- hollight.dist (m * nadd_rinv x n, n * nadd_rinv x m)
- <= m * n *
- hollight.dist (m * dest_nadd x n, n * dest_nadd x m) +
- dest_nadd x m * dest_nadd x n * (m + n)"
- by (import hollight NADD_MUL_LINV_LEMMA4)
-
-lemma NADD_MUL_LINV_LEMMA5: "~ nadd_eq x (nadd_of_num 0)
-==> EX B N.
- ALL m n.
- N <= m & N <= n -->
- dest_nadd x m * dest_nadd x n *
- hollight.dist (m * nadd_rinv x n, n * nadd_rinv x m)
- <= B * (m * n * (m + n))"
- by (import hollight NADD_MUL_LINV_LEMMA5)
-
-lemma NADD_MUL_LINV_LEMMA6: "~ nadd_eq x (nadd_of_num 0)
-==> EX B N.
- ALL m n.
- N <= m & N <= n -->
- m * n * hollight.dist (m * nadd_rinv x n, n * nadd_rinv x m)
- <= B * (m * n * (m + n))"
- by (import hollight NADD_MUL_LINV_LEMMA6)
-
-lemma NADD_MUL_LINV_LEMMA7: "~ nadd_eq x (nadd_of_num 0)
-==> EX B N.
- ALL m n.
- N <= m & N <= n -->
- hollight.dist (m * nadd_rinv x n, n * nadd_rinv x m)
- <= B * (m + n)"
- by (import hollight NADD_MUL_LINV_LEMMA7)
-
-lemma NADD_MUL_LINV_LEMMA7a: "~ nadd_eq x (nadd_of_num 0)
-==> EX A B.
- ALL m n.
- m <= N -->
- hollight.dist (m * nadd_rinv x n, n * nadd_rinv x m) <= A * n + B"
- by (import hollight NADD_MUL_LINV_LEMMA7a)
-
-lemma NADD_MUL_LINV_LEMMA8: "~ nadd_eq x (nadd_of_num 0)
-==> EX B. ALL m n.
- hollight.dist (m * nadd_rinv x n, n * nadd_rinv x m)
- <= B * (m + n)"
- by (import hollight NADD_MUL_LINV_LEMMA8)
-
-definition
- nadd_inv :: "nadd => nadd" where
- "nadd_inv ==
-%u. if nadd_eq u (nadd_of_num 0) then nadd_of_num 0
- else mk_nadd (nadd_rinv u)"
-
-lemma DEF_nadd_inv: "nadd_inv =
-(%u. if nadd_eq u (nadd_of_num 0) then nadd_of_num 0
- else mk_nadd (nadd_rinv u))"
- by (import hollight DEF_nadd_inv)
-
-lemma NADD_INV: "dest_nadd (nadd_inv x) =
-(if nadd_eq x (nadd_of_num 0) then %n. 0 else nadd_rinv x)"
- by (import hollight NADD_INV)
-
-lemma NADD_MUL_LINV: "~ nadd_eq x (nadd_of_num 0)
-==> nadd_eq (nadd_mul (nadd_inv x) x) (nadd_of_num 1)"
- by (import hollight NADD_MUL_LINV)
-
-lemma NADD_INV_0: "nadd_eq (nadd_inv (nadd_of_num 0)) (nadd_of_num 0)"
- by (import hollight NADD_INV_0)
-
-lemma NADD_INV_WELLDEF: "nadd_eq x y ==> nadd_eq (nadd_inv x) (nadd_inv y)"
- by (import hollight NADD_INV_WELLDEF)
-
-typedef (open) hreal = "{s. EX x. s = nadd_eq x}" morphisms "dest_hreal" "mk_hreal"
- apply (rule light_ex_imp_nonempty[where t="nadd_eq x"])
- by (import hollight TYDEF_hreal)
-
-syntax
- dest_hreal :: _
-
-syntax
- mk_hreal :: _
-
-lemmas "TYDEF_hreal_@intern" = typedef_hol2hollight
- [where a="a :: hreal" and r=r ,
- OF type_definition_hreal]
-
-definition
- hreal_of_num :: "nat => hreal" where
- "hreal_of_num == %m. mk_hreal (nadd_eq (nadd_of_num m))"
-
-lemma DEF_hreal_of_num: "hreal_of_num = (%m. mk_hreal (nadd_eq (nadd_of_num m)))"
- by (import hollight DEF_hreal_of_num)
-
-definition
- hreal_add :: "hreal => hreal => hreal" where
- "hreal_add ==
-%x y. mk_hreal
- (%u. EX xa ya.
- nadd_eq (nadd_add xa ya) u &
- dest_hreal x xa & dest_hreal y ya)"
-
-lemma DEF_hreal_add: "hreal_add =
-(%x y. mk_hreal
- (%u. EX xa ya.
- nadd_eq (nadd_add xa ya) u &
- dest_hreal x xa & dest_hreal y ya))"
- by (import hollight DEF_hreal_add)
-
-definition
- hreal_mul :: "hreal => hreal => hreal" where
- "hreal_mul ==
-%x y. mk_hreal
- (%u. EX xa ya.
- nadd_eq (nadd_mul xa ya) u &
- dest_hreal x xa & dest_hreal y ya)"
-
-lemma DEF_hreal_mul: "hreal_mul =
-(%x y. mk_hreal
- (%u. EX xa ya.
- nadd_eq (nadd_mul xa ya) u &
- dest_hreal x xa & dest_hreal y ya))"
- by (import hollight DEF_hreal_mul)
-
-definition
- hreal_le :: "hreal => hreal => bool" where
- "hreal_le ==
-%x y. SOME u.
- EX xa ya. nadd_le xa ya = u & dest_hreal x xa & dest_hreal y ya"
-
-lemma DEF_hreal_le: "hreal_le =
-(%x y. SOME u.
- EX xa ya. nadd_le xa ya = u & dest_hreal x xa & dest_hreal y ya)"
- by (import hollight DEF_hreal_le)
-
-definition
- hreal_inv :: "hreal => hreal" where
- "hreal_inv ==
-%x. mk_hreal (%u. EX xa. nadd_eq (nadd_inv xa) u & dest_hreal x xa)"
-
-lemma DEF_hreal_inv: "hreal_inv =
-(%x. mk_hreal (%u. EX xa. nadd_eq (nadd_inv xa) u & dest_hreal x xa))"
- by (import hollight DEF_hreal_inv)
-
-lemma HREAL_LE_EXISTS_DEF: "hreal_le m n = (EX d. n = hreal_add m d)"
- by (import hollight HREAL_LE_EXISTS_DEF)
-
-lemma HREAL_EQ_ADD_LCANCEL: "(hreal_add m n = hreal_add m p) = (n = p)"
- by (import hollight HREAL_EQ_ADD_LCANCEL)
-
-lemma HREAL_EQ_ADD_RCANCEL: "(hreal_add x xb = hreal_add xa xb) = (x = xa)"
- by (import hollight HREAL_EQ_ADD_RCANCEL)
-
-lemma HREAL_LE_ADD_LCANCEL: "hreal_le (hreal_add x xa) (hreal_add x xb) = hreal_le xa xb"
- by (import hollight HREAL_LE_ADD_LCANCEL)
-
-lemma HREAL_LE_ADD_RCANCEL: "hreal_le (hreal_add x xb) (hreal_add xa xb) = hreal_le x xa"
- by (import hollight HREAL_LE_ADD_RCANCEL)
-
-lemma HREAL_ADD_RID: "hreal_add x (hreal_of_num 0) = x"
- by (import hollight HREAL_ADD_RID)
-
-lemma HREAL_ADD_RDISTRIB: "hreal_mul (hreal_add x xa) xb = hreal_add (hreal_mul x xb) (hreal_mul xa xb)"
- by (import hollight HREAL_ADD_RDISTRIB)
-
-lemma HREAL_MUL_LZERO: "hreal_mul (hreal_of_num 0) m = hreal_of_num 0"
- by (import hollight HREAL_MUL_LZERO)
-
-lemma HREAL_MUL_RZERO: "hreal_mul x (hreal_of_num 0) = hreal_of_num 0"
- by (import hollight HREAL_MUL_RZERO)
-
-lemma HREAL_ADD_AC: "hreal_add m n = hreal_add n m &
-hreal_add (hreal_add m n) p = hreal_add m (hreal_add n p) &
-hreal_add m (hreal_add n p) = hreal_add n (hreal_add m p)"
- by (import hollight HREAL_ADD_AC)
-
-lemma HREAL_LE_ADD2: "hreal_le a b & hreal_le c d ==> hreal_le (hreal_add a c) (hreal_add b d)"
- by (import hollight HREAL_LE_ADD2)
-
-lemma HREAL_LE_MUL_RCANCEL_IMP: "hreal_le a b ==> hreal_le (hreal_mul a c) (hreal_mul b c)"
- by (import hollight HREAL_LE_MUL_RCANCEL_IMP)
-
-definition
- treal_of_num :: "nat => hreal * hreal" where
- "treal_of_num == %u. (hreal_of_num u, hreal_of_num 0)"
-
-lemma DEF_treal_of_num: "treal_of_num = (%u. (hreal_of_num u, hreal_of_num 0))"
- by (import hollight DEF_treal_of_num)
-
-definition
- treal_neg :: "hreal * hreal => hreal * hreal" where
- "treal_neg == %u. (snd u, fst u)"
-
-lemma DEF_treal_neg: "treal_neg = (%u. (snd u, fst u))"
- by (import hollight DEF_treal_neg)
-
-definition
- treal_add :: "hreal * hreal => hreal * hreal => hreal * hreal" where
- "treal_add == %u ua. (hreal_add (fst u) (fst ua), hreal_add (snd u) (snd ua))"
-
-lemma DEF_treal_add: "treal_add =
-(%u ua. (hreal_add (fst u) (fst ua), hreal_add (snd u) (snd ua)))"
- by (import hollight DEF_treal_add)
-
-definition
- treal_mul :: "hreal * hreal => hreal * hreal => hreal * hreal" where
- "treal_mul ==
-%u ua.
- (hreal_add (hreal_mul (fst u) (fst ua)) (hreal_mul (snd u) (snd ua)),
- hreal_add (hreal_mul (fst u) (snd ua)) (hreal_mul (snd u) (fst ua)))"
-
-lemma DEF_treal_mul: "treal_mul =
-(%u ua.
- (hreal_add (hreal_mul (fst u) (fst ua)) (hreal_mul (snd u) (snd ua)),
- hreal_add (hreal_mul (fst u) (snd ua)) (hreal_mul (snd u) (fst ua))))"
- by (import hollight DEF_treal_mul)
-
-definition
- treal_le :: "hreal * hreal => hreal * hreal => bool" where
- "treal_le ==
-%u ua. hreal_le (hreal_add (fst u) (snd ua)) (hreal_add (fst ua) (snd u))"
-
-lemma DEF_treal_le: "treal_le =
-(%u ua. hreal_le (hreal_add (fst u) (snd ua)) (hreal_add (fst ua) (snd u)))"
- by (import hollight DEF_treal_le)
-
-definition
- treal_inv :: "hreal * hreal => hreal * hreal" where
- "treal_inv ==
-%u. if fst u = snd u then (hreal_of_num 0, hreal_of_num 0)
- else if hreal_le (snd u) (fst u)
- then (hreal_inv (SOME d. fst u = hreal_add (snd u) d),
- hreal_of_num 0)
- else (hreal_of_num 0,
- hreal_inv (SOME d. snd u = hreal_add (fst u) d))"
-
-lemma DEF_treal_inv: "treal_inv =
-(%u. if fst u = snd u then (hreal_of_num 0, hreal_of_num 0)
- else if hreal_le (snd u) (fst u)
- then (hreal_inv (SOME d. fst u = hreal_add (snd u) d),
- hreal_of_num 0)
- else (hreal_of_num 0,
- hreal_inv (SOME d. snd u = hreal_add (fst u) d)))"
- by (import hollight DEF_treal_inv)
-
-definition
- treal_eq :: "hreal * hreal => hreal * hreal => bool" where
- "treal_eq == %u ua. hreal_add (fst u) (snd ua) = hreal_add (fst ua) (snd u)"
-
-lemma DEF_treal_eq: "treal_eq = (%u ua. hreal_add (fst u) (snd ua) = hreal_add (fst ua) (snd u))"
- by (import hollight DEF_treal_eq)
-
-lemma TREAL_EQ_REFL: "treal_eq x x"
- by (import hollight TREAL_EQ_REFL)
-
-lemma TREAL_EQ_SYM: "treal_eq x y = treal_eq y x"
- by (import hollight TREAL_EQ_SYM)
-
-lemma TREAL_EQ_TRANS: "treal_eq x y & treal_eq y z ==> treal_eq x z"
- by (import hollight TREAL_EQ_TRANS)
-
-lemma TREAL_EQ_AP: "x = xa ==> treal_eq x xa"
- by (import hollight TREAL_EQ_AP)
-
-lemma TREAL_OF_NUM_EQ: "treal_eq (treal_of_num x) (treal_of_num xa) = (x = xa)"
- by (import hollight TREAL_OF_NUM_EQ)
-
-lemma TREAL_OF_NUM_LE: "treal_le (treal_of_num x) (treal_of_num xa) = (x <= xa)"
- by (import hollight TREAL_OF_NUM_LE)
-
-lemma TREAL_OF_NUM_ADD: "treal_eq (treal_add (treal_of_num x) (treal_of_num xa))
- (treal_of_num (x + xa))"
- by (import hollight TREAL_OF_NUM_ADD)
-
-lemma TREAL_OF_NUM_MUL: "treal_eq (treal_mul (treal_of_num x) (treal_of_num xa))
- (treal_of_num (x * xa))"
- by (import hollight TREAL_OF_NUM_MUL)
-
-lemma TREAL_ADD_SYM_EQ: "treal_add x y = treal_add y x"
- by (import hollight TREAL_ADD_SYM_EQ)
-
-lemma TREAL_MUL_SYM_EQ: "treal_mul x y = treal_mul y x"
- by (import hollight TREAL_MUL_SYM_EQ)
-
-lemma TREAL_ADD_SYM: "treal_eq (treal_add x y) (treal_add y x)"
- by (import hollight TREAL_ADD_SYM)
-
-lemma TREAL_ADD_ASSOC: "treal_eq (treal_add x (treal_add y z)) (treal_add (treal_add x y) z)"
- by (import hollight TREAL_ADD_ASSOC)
-
-lemma TREAL_ADD_LID: "treal_eq (treal_add (treal_of_num 0) x) x"
- by (import hollight TREAL_ADD_LID)
-
-lemma TREAL_ADD_LINV: "treal_eq (treal_add (treal_neg x) x) (treal_of_num 0)"
- by (import hollight TREAL_ADD_LINV)
-
-lemma TREAL_MUL_SYM: "treal_eq (treal_mul x xa) (treal_mul xa x)"
- by (import hollight TREAL_MUL_SYM)
-
-lemma TREAL_MUL_ASSOC: "treal_eq (treal_mul x (treal_mul y z)) (treal_mul (treal_mul x y) z)"
- by (import hollight TREAL_MUL_ASSOC)
-
-lemma TREAL_MUL_LID: "treal_eq (treal_mul (treal_of_num 1) x) x"
- by (import hollight TREAL_MUL_LID)
-
-lemma TREAL_ADD_LDISTRIB: "treal_eq (treal_mul x (treal_add y z))
- (treal_add (treal_mul x y) (treal_mul x z))"
- by (import hollight TREAL_ADD_LDISTRIB)
-
-lemma TREAL_LE_REFL: "treal_le x x"
- by (import hollight TREAL_LE_REFL)
-
-lemma TREAL_LE_ANTISYM: "(treal_le x y & treal_le y x) = treal_eq x y"
- by (import hollight TREAL_LE_ANTISYM)
-
-lemma TREAL_LE_TRANS: "treal_le x y & treal_le y z ==> treal_le x z"
- by (import hollight TREAL_LE_TRANS)
-
-lemma TREAL_LE_TOTAL: "treal_le x y | treal_le y x"
- by (import hollight TREAL_LE_TOTAL)
-
-lemma TREAL_LE_LADD_IMP: "treal_le y z ==> treal_le (treal_add x y) (treal_add x z)"
- by (import hollight TREAL_LE_LADD_IMP)
-
-lemma TREAL_LE_MUL: "treal_le (treal_of_num 0) x & treal_le (treal_of_num 0) y
-==> treal_le (treal_of_num 0) (treal_mul x y)"
- by (import hollight TREAL_LE_MUL)
-
-lemma TREAL_INV_0: "treal_eq (treal_inv (treal_of_num 0)) (treal_of_num 0)"
- by (import hollight TREAL_INV_0)
-
-lemma TREAL_MUL_LINV: "~ treal_eq x (treal_of_num 0)
-==> treal_eq (treal_mul (treal_inv x) x) (treal_of_num 1)"
- by (import hollight TREAL_MUL_LINV)
-
-lemma TREAL_OF_NUM_WELLDEF: "m = n ==> treal_eq (treal_of_num m) (treal_of_num n)"
- by (import hollight TREAL_OF_NUM_WELLDEF)
-
-lemma TREAL_NEG_WELLDEF: "treal_eq x1 x2 ==> treal_eq (treal_neg x1) (treal_neg x2)"
- by (import hollight TREAL_NEG_WELLDEF)
-
-lemma TREAL_ADD_WELLDEFR: "treal_eq x1 x2 ==> treal_eq (treal_add x1 y) (treal_add x2 y)"
- by (import hollight TREAL_ADD_WELLDEFR)
-
-lemma TREAL_ADD_WELLDEF: "treal_eq x1 x2 & treal_eq y1 y2
-==> treal_eq (treal_add x1 y1) (treal_add x2 y2)"
- by (import hollight TREAL_ADD_WELLDEF)
-
-lemma TREAL_MUL_WELLDEFR: "treal_eq x1 x2 ==> treal_eq (treal_mul x1 y) (treal_mul x2 y)"
- by (import hollight TREAL_MUL_WELLDEFR)
-
-lemma TREAL_MUL_WELLDEF: "treal_eq x1 x2 & treal_eq y1 y2
-==> treal_eq (treal_mul x1 y1) (treal_mul x2 y2)"
- by (import hollight TREAL_MUL_WELLDEF)
-
-lemma TREAL_EQ_IMP_LE: "treal_eq x y ==> treal_le x y"
- by (import hollight TREAL_EQ_IMP_LE)
-
-lemma TREAL_LE_WELLDEF: "treal_eq x1 x2 & treal_eq y1 y2 ==> treal_le x1 y1 = treal_le x2 y2"
- by (import hollight TREAL_LE_WELLDEF)
-
-lemma TREAL_INV_WELLDEF: "treal_eq x y ==> treal_eq (treal_inv x) (treal_inv y)"
- by (import hollight TREAL_INV_WELLDEF)
-
-typedef (open) real = "{s. EX x. s = treal_eq x}" morphisms "dest_real" "mk_real"
- apply (rule light_ex_imp_nonempty[where t="treal_eq x"])
- by (import hollight TYDEF_real)
-
-syntax
- dest_real :: _
-
-syntax
- mk_real :: _
-
-lemmas "TYDEF_real_@intern" = typedef_hol2hollight
- [where a="a :: hollight.real" and r=r ,
- OF type_definition_real]
-
-definition
- real_of_num :: "nat => hollight.real" where
- "real_of_num == %m. mk_real (treal_eq (treal_of_num m))"
-
-lemma DEF_real_of_num: "real_of_num = (%m. mk_real (treal_eq (treal_of_num m)))"
- by (import hollight DEF_real_of_num)
-
-definition
- real_neg :: "hollight.real => hollight.real" where
- "real_neg ==
-%x1. mk_real (%u. EX x1a. treal_eq (treal_neg x1a) u & dest_real x1 x1a)"
-
-lemma DEF_real_neg: "real_neg =
-(%x1. mk_real (%u. EX x1a. treal_eq (treal_neg x1a) u & dest_real x1 x1a))"
- by (import hollight DEF_real_neg)
-
-definition
- real_add :: "hollight.real => hollight.real => hollight.real" where
- "real_add ==
-%x1 y1.
- mk_real
- (%u. EX x1a y1a.
- treal_eq (treal_add x1a y1a) u &
- dest_real x1 x1a & dest_real y1 y1a)"
-
-lemma DEF_real_add: "real_add =
-(%x1 y1.
- mk_real
- (%u. EX x1a y1a.
- treal_eq (treal_add x1a y1a) u &
- dest_real x1 x1a & dest_real y1 y1a))"
- by (import hollight DEF_real_add)
-
-definition
- real_mul :: "hollight.real => hollight.real => hollight.real" where
- "real_mul ==
-%x1 y1.
- mk_real
- (%u. EX x1a y1a.
- treal_eq (treal_mul x1a y1a) u &
- dest_real x1 x1a & dest_real y1 y1a)"
-
-lemma DEF_real_mul: "real_mul =
-(%x1 y1.
- mk_real
- (%u. EX x1a y1a.
- treal_eq (treal_mul x1a y1a) u &
- dest_real x1 x1a & dest_real y1 y1a))"
- by (import hollight DEF_real_mul)
-
-definition
- real_le :: "hollight.real => hollight.real => bool" where
- "real_le ==
-%x1 y1.
- SOME u.
- EX x1a y1a. treal_le x1a y1a = u & dest_real x1 x1a & dest_real y1 y1a"
-
-lemma DEF_real_le: "real_le =
-(%x1 y1.
- SOME u.
- EX x1a y1a.
- treal_le x1a y1a = u & dest_real x1 x1a & dest_real y1 y1a)"
- by (import hollight DEF_real_le)
-
-definition
- real_inv :: "hollight.real => hollight.real" where
- "real_inv ==
-%x. mk_real (%u. EX xa. treal_eq (treal_inv xa) u & dest_real x xa)"
-
-lemma DEF_real_inv: "real_inv =
-(%x. mk_real (%u. EX xa. treal_eq (treal_inv xa) u & dest_real x xa))"
- by (import hollight DEF_real_inv)
-
-definition
- real_sub :: "hollight.real => hollight.real => hollight.real" where
- "real_sub == %u ua. real_add u (real_neg ua)"
-
-lemma DEF_real_sub: "real_sub = (%u ua. real_add u (real_neg ua))"
- by (import hollight DEF_real_sub)
-
-definition
- real_lt :: "hollight.real => hollight.real => bool" where
- "real_lt == %u ua. ~ real_le ua u"
-
-lemma DEF_real_lt: "real_lt = (%u ua. ~ real_le ua u)"
- by (import hollight DEF_real_lt)
-
-definition
- real_ge :: "hollight.real => hollight.real => bool" where
- "real_ge == %u ua. real_le ua u"
-
-lemma DEF_real_ge: "real_ge = (%u ua. real_le ua u)"
- by (import hollight DEF_real_ge)
-
-definition
- real_gt :: "hollight.real => hollight.real => bool" where
- "real_gt == %u ua. real_lt ua u"
-
-lemma DEF_real_gt: "real_gt = (%u ua. real_lt ua u)"
- by (import hollight DEF_real_gt)
-
-definition
- real_abs :: "hollight.real => hollight.real" where
- "real_abs == %u. if real_le (real_of_num 0) u then u else real_neg u"
-
-lemma DEF_real_abs: "real_abs = (%u. if real_le (real_of_num 0) u then u else real_neg u)"
- by (import hollight DEF_real_abs)
-
-definition
- real_pow :: "hollight.real => nat => hollight.real" where
- "real_pow ==
-SOME real_pow.
- (ALL x. real_pow x 0 = real_of_num 1) &
- (ALL x n. real_pow x (Suc n) = real_mul x (real_pow x n))"
-
-lemma DEF_real_pow: "real_pow =
-(SOME real_pow.
- (ALL x. real_pow x 0 = real_of_num 1) &
- (ALL x n. real_pow x (Suc n) = real_mul x (real_pow x n)))"
- by (import hollight DEF_real_pow)
-
-definition
- real_div :: "hollight.real => hollight.real => hollight.real" where
- "real_div == %u ua. real_mul u (real_inv ua)"
-
-lemma DEF_real_div: "real_div = (%u ua. real_mul u (real_inv ua))"
- by (import hollight DEF_real_div)
-
-definition
- real_max :: "hollight.real => hollight.real => hollight.real" where
- "real_max == %u ua. if real_le u ua then ua else u"
-
-lemma DEF_real_max: "real_max = (%u ua. if real_le u ua then ua else u)"
- by (import hollight DEF_real_max)
-
-definition
- real_min :: "hollight.real => hollight.real => hollight.real" where
- "real_min == %u ua. if real_le u ua then u else ua"
-
-lemma DEF_real_min: "real_min = (%u ua. if real_le u ua then u else ua)"
- by (import hollight DEF_real_min)
-
-lemma REAL_HREAL_LEMMA1: "EX x. (ALL xa. real_le (real_of_num 0) xa = (EX y. xa = x y)) &
- (ALL y z. hreal_le y z = real_le (x y) (x z))"
- by (import hollight REAL_HREAL_LEMMA1)
-
-lemma REAL_HREAL_LEMMA2: "EX x r.
- (ALL xa. x (r xa) = xa) &
- (ALL xa. real_le (real_of_num 0) xa --> r (x xa) = xa) &
- (ALL x. real_le (real_of_num 0) (r x)) &
- (ALL x y. hreal_le x y = real_le (r x) (r y))"
- by (import hollight REAL_HREAL_LEMMA2)
-
-lemma REAL_COMPLETE_SOMEPOS: "(EX x. P x & real_le (real_of_num 0) x) & (EX M. ALL x. P x --> real_le x M)
-==> EX M. (ALL x. P x --> real_le x M) &
- (ALL M'. (ALL x. P x --> real_le x M') --> real_le M M')"
- by (import hollight REAL_COMPLETE_SOMEPOS)
-
-lemma REAL_COMPLETE: "Ex P & (EX M. ALL x. P x --> real_le x M)
-==> EX M. (ALL x. P x --> real_le x M) &
- (ALL M'. (ALL x. P x --> real_le x M') --> real_le M M')"
- by (import hollight REAL_COMPLETE)
-
-lemma REAL_ADD_AC: "real_add m n = real_add n m &
-real_add (real_add m n) p = real_add m (real_add n p) &
-real_add m (real_add n p) = real_add n (real_add m p)"
- by (import hollight REAL_ADD_AC)
-
-lemma REAL_ADD_RINV: "real_add x (real_neg x) = real_of_num 0"
- by (import hollight REAL_ADD_RINV)
-
-lemma REAL_EQ_ADD_LCANCEL: "(real_add x y = real_add x z) = (y = z)"
- by (import hollight REAL_EQ_ADD_LCANCEL)
-
-lemma REAL_EQ_ADD_RCANCEL: "(real_add x z = real_add y z) = (x = y)"
- by (import hollight REAL_EQ_ADD_RCANCEL)
-
-lemma REAL_MUL_RZERO: "real_mul x (real_of_num 0) = real_of_num 0"
- by (import hollight REAL_MUL_RZERO)
-
-lemma REAL_MUL_LZERO: "real_mul (real_of_num 0) x = real_of_num 0"
- by (import hollight REAL_MUL_LZERO)
-
-lemma REAL_NEG_NEG: "real_neg (real_neg x) = x"
- by (import hollight REAL_NEG_NEG)
-
-lemma REAL_MUL_RNEG: "real_mul x (real_neg y) = real_neg (real_mul x y)"
- by (import hollight REAL_MUL_RNEG)
-
-lemma REAL_MUL_LNEG: "real_mul (real_neg x) y = real_neg (real_mul x y)"
- by (import hollight REAL_MUL_LNEG)
-
-lemma REAL_NEG_ADD: "real_neg (real_add x y) = real_add (real_neg x) (real_neg y)"
- by (import hollight REAL_NEG_ADD)
-
-lemma REAL_ADD_RID: "real_add x (real_of_num 0) = x"
- by (import hollight REAL_ADD_RID)
-
-lemma REAL_NEG_0: "real_neg (real_of_num 0) = real_of_num 0"
- by (import hollight REAL_NEG_0)
-
-lemma REAL_LE_LNEG: "real_le (real_neg x) y = real_le (real_of_num 0) (real_add x y)"
- by (import hollight REAL_LE_LNEG)
-
-lemma REAL_LE_NEG2: "real_le (real_neg x) (real_neg y) = real_le y x"
- by (import hollight REAL_LE_NEG2)
-
-lemma REAL_LE_RNEG: "real_le x (real_neg y) = real_le (real_add x y) (real_of_num 0)"
- by (import hollight REAL_LE_RNEG)
-
-lemma REAL_OF_NUM_POW: "real_pow (real_of_num x) n = real_of_num (x ^ n)"
- by (import hollight REAL_OF_NUM_POW)
-
-lemma REAL_POW_NEG: "real_pow (real_neg x) n =
-(if even n then real_pow x n else real_neg (real_pow x n))"
- by (import hollight REAL_POW_NEG)
-
-lemma REAL_ABS_NUM: "real_abs (real_of_num x) = real_of_num x"
- by (import hollight REAL_ABS_NUM)
-
-lemma REAL_ABS_NEG: "real_abs (real_neg x) = real_abs x"
- by (import hollight REAL_ABS_NEG)
-
-lemma REAL_LTE_TOTAL: "real_lt x xa | real_le xa x"
- by (import hollight REAL_LTE_TOTAL)
-
-lemma REAL_LET_TOTAL: "real_le x xa | real_lt xa x"
- by (import hollight REAL_LET_TOTAL)
-
-lemma REAL_LT_IMP_LE: "real_lt x y ==> real_le x y"
- by (import hollight REAL_LT_IMP_LE)
-
-lemma REAL_LTE_TRANS: "real_lt x y & real_le y z ==> real_lt x z"
- by (import hollight REAL_LTE_TRANS)
-
-lemma REAL_LET_TRANS: "real_le x y & real_lt y z ==> real_lt x z"
- by (import hollight REAL_LET_TRANS)
-
-lemma REAL_LT_TRANS: "real_lt x y & real_lt y z ==> real_lt x z"
- by (import hollight REAL_LT_TRANS)
-
-lemma REAL_LE_ADD: "real_le (real_of_num 0) x & real_le (real_of_num 0) y
-==> real_le (real_of_num 0) (real_add x y)"
- by (import hollight REAL_LE_ADD)
-
-lemma REAL_LTE_ANTISYM: "~ (real_lt x y & real_le y x)"
- by (import hollight REAL_LTE_ANTISYM)
-
-lemma REAL_SUB_LE: "real_le (real_of_num 0) (real_sub x xa) = real_le xa x"
- by (import hollight REAL_SUB_LE)
-
-lemma REAL_NEG_SUB: "real_neg (real_sub x xa) = real_sub xa x"
- by (import hollight REAL_NEG_SUB)
-
-lemma REAL_LE_LT: "real_le x xa = (real_lt x xa | x = xa)"
- by (import hollight REAL_LE_LT)
-
-lemma REAL_SUB_LT: "real_lt (real_of_num 0) (real_sub x xa) = real_lt xa x"
- by (import hollight REAL_SUB_LT)
-
-lemma REAL_NOT_LT: "(~ real_lt x xa) = real_le xa x"
- by (import hollight REAL_NOT_LT)
-
-lemma REAL_SUB_0: "(real_sub x y = real_of_num 0) = (x = y)"
- by (import hollight REAL_SUB_0)
-
-lemma REAL_LT_LE: "real_lt x y = (real_le x y & x ~= y)"
- by (import hollight REAL_LT_LE)
-
-lemma REAL_LT_REFL: "~ real_lt x x"
- by (import hollight REAL_LT_REFL)
-
-lemma REAL_LTE_ADD: "real_lt (real_of_num 0) x & real_le (real_of_num 0) y
-==> real_lt (real_of_num 0) (real_add x y)"
- by (import hollight REAL_LTE_ADD)
-
-lemma REAL_LET_ADD: "real_le (real_of_num 0) x & real_lt (real_of_num 0) y
-==> real_lt (real_of_num 0) (real_add x y)"
- by (import hollight REAL_LET_ADD)
-
-lemma REAL_LT_ADD: "real_lt (real_of_num 0) x & real_lt (real_of_num 0) y
-==> real_lt (real_of_num 0) (real_add x y)"
- by (import hollight REAL_LT_ADD)
-
-lemma REAL_ENTIRE: "(real_mul x y = real_of_num 0) = (x = real_of_num 0 | y = real_of_num 0)"
- by (import hollight REAL_ENTIRE)
-
-lemma REAL_LE_NEGTOTAL: "real_le (real_of_num 0) x | real_le (real_of_num 0) (real_neg x)"
- by (import hollight REAL_LE_NEGTOTAL)
-
-lemma REAL_LE_SQUARE: "real_le (real_of_num 0) (real_mul x x)"
- by (import hollight REAL_LE_SQUARE)
-
-lemma REAL_MUL_RID: "real_mul x (real_of_num 1) = x"
- by (import hollight REAL_MUL_RID)
-
-lemma REAL_POW_2: "real_pow x 2 = real_mul x x"
- by (import hollight REAL_POW_2)
-
-lemma REAL_POLY_CLAUSES: "(ALL x y z. real_add x (real_add y z) = real_add (real_add x y) z) &
-(ALL x y. real_add x y = real_add y x) &
-(ALL x. real_add (real_of_num 0) x = x) &
-(ALL x y z. real_mul x (real_mul y z) = real_mul (real_mul x y) z) &
-(ALL x y. real_mul x y = real_mul y x) &
-(ALL x. real_mul (real_of_num 1) x = x) &
-(ALL x. real_mul (real_of_num 0) x = real_of_num 0) &
-(ALL x xa xb.
- real_mul x (real_add xa xb) =
- real_add (real_mul x xa) (real_mul x xb)) &
-(ALL x. real_pow x 0 = real_of_num 1) &
-(ALL x xa. real_pow x (Suc xa) = real_mul x (real_pow x xa))"
- by (import hollight REAL_POLY_CLAUSES)
-
-lemma REAL_POLY_NEG_CLAUSES: "(ALL x. real_neg x = real_mul (real_neg (real_of_num 1)) x) &
-(ALL x xa.
- real_sub x xa = real_add x (real_mul (real_neg (real_of_num 1)) xa))"
- by (import hollight REAL_POLY_NEG_CLAUSES)
-
-lemma REAL_POS: "real_le (real_of_num 0) (real_of_num x)"
- by (import hollight REAL_POS)
-
-lemma REAL_OF_NUM_LT: "real_lt (real_of_num x) (real_of_num xa) = (x < xa)"
- by (import hollight REAL_OF_NUM_LT)
-
-lemma REAL_OF_NUM_GE: "real_ge (real_of_num x) (real_of_num xa) = (xa <= x)"
- by (import hollight REAL_OF_NUM_GE)
-
-lemma REAL_OF_NUM_GT: "real_gt (real_of_num x) (real_of_num xa) = (xa < x)"
- by (import hollight REAL_OF_NUM_GT)
-
-lemma REAL_OF_NUM_MAX: "real_max (real_of_num x) (real_of_num xa) = real_of_num (max x xa)"
- by (import hollight REAL_OF_NUM_MAX)
-
-lemma REAL_OF_NUM_MIN: "real_min (real_of_num x) (real_of_num xa) = real_of_num (min x xa)"
- by (import hollight REAL_OF_NUM_MIN)
-
-lemma REAL_OF_NUM_SUC: "real_add (real_of_num x) (real_of_num 1) = real_of_num (Suc x)"
- by (import hollight REAL_OF_NUM_SUC)
-
-lemma REAL_OF_NUM_SUB: "m <= n ==> real_sub (real_of_num n) (real_of_num m) = real_of_num (n - m)"
- by (import hollight REAL_OF_NUM_SUB)
-
-lemma REAL_MUL_AC: "real_mul m n = real_mul n m &
-real_mul (real_mul m n) p = real_mul m (real_mul n p) &
-real_mul m (real_mul n p) = real_mul n (real_mul m p)"
- by (import hollight REAL_MUL_AC)
-
-lemma REAL_ADD_RDISTRIB: "real_mul (real_add x y) z = real_add (real_mul x z) (real_mul y z)"
- by (import hollight REAL_ADD_RDISTRIB)
-
-lemma REAL_LT_LADD_IMP: "real_lt y z ==> real_lt (real_add x y) (real_add x z)"
- by (import hollight REAL_LT_LADD_IMP)
-
-lemma REAL_LT_MUL: "real_lt (real_of_num 0) x & real_lt (real_of_num 0) y
-==> real_lt (real_of_num 0) (real_mul x y)"
- by (import hollight REAL_LT_MUL)
-
-lemma REAL_EQ_ADD_LCANCEL_0: "(real_add x y = x) = (y = real_of_num 0)"
- by (import hollight REAL_EQ_ADD_LCANCEL_0)
-
-lemma REAL_EQ_ADD_RCANCEL_0: "(real_add x y = y) = (x = real_of_num 0)"
- by (import hollight REAL_EQ_ADD_RCANCEL_0)
-
-lemma REAL_LNEG_UNIQ: "(real_add x y = real_of_num 0) = (x = real_neg y)"
- by (import hollight REAL_LNEG_UNIQ)
-
-lemma REAL_RNEG_UNIQ: "(real_add x y = real_of_num 0) = (y = real_neg x)"
- by (import hollight REAL_RNEG_UNIQ)
-
-lemma REAL_NEG_LMUL: "real_neg (real_mul x y) = real_mul (real_neg x) y"
- by (import hollight REAL_NEG_LMUL)
-
-lemma REAL_NEG_RMUL: "real_neg (real_mul x y) = real_mul x (real_neg y)"
- by (import hollight REAL_NEG_RMUL)
-
-lemma REAL_NEGNEG: "real_neg (real_neg x) = x"
- by (import hollight REAL_NEGNEG)
-
-lemma REAL_NEG_MUL2: "real_mul (real_neg x) (real_neg y) = real_mul x y"
- by (import hollight REAL_NEG_MUL2)
-
-lemma REAL_LT_LADD: "real_lt (real_add x y) (real_add x z) = real_lt y z"
- by (import hollight REAL_LT_LADD)
-
-lemma REAL_LT_RADD: "real_lt (real_add x z) (real_add y z) = real_lt x y"
- by (import hollight REAL_LT_RADD)
-
-lemma REAL_LT_ANTISYM: "~ (real_lt x y & real_lt y x)"
- by (import hollight REAL_LT_ANTISYM)
-
-lemma REAL_LT_GT: "real_lt x y ==> ~ real_lt y x"
- by (import hollight REAL_LT_GT)
-
-lemma REAL_NOT_EQ: "(x ~= y) = (real_lt x y | real_lt y x)"
- by (import hollight REAL_NOT_EQ)
-
-lemma REAL_LET_ANTISYM: "~ (real_le x y & real_lt y x)"
- by (import hollight REAL_LET_ANTISYM)
-
-lemma REAL_NEG_LT0: "real_lt (real_neg x) (real_of_num 0) = real_lt (real_of_num 0) x"
- by (import hollight REAL_NEG_LT0)
-
-lemma REAL_NEG_GT0: "real_lt (real_of_num 0) (real_neg x) = real_lt x (real_of_num 0)"
- by (import hollight REAL_NEG_GT0)
-
-lemma REAL_NEG_LE0: "real_le (real_neg x) (real_of_num 0) = real_le (real_of_num 0) x"
- by (import hollight REAL_NEG_LE0)
-
-lemma REAL_NEG_GE0: "real_le (real_of_num 0) (real_neg x) = real_le x (real_of_num 0)"
- by (import hollight REAL_NEG_GE0)
-
-lemma REAL_LT_TOTAL: "x = y | real_lt x y | real_lt y x"
- by (import hollight REAL_LT_TOTAL)
-
-lemma REAL_LT_NEGTOTAL: "x = real_of_num 0 |
-real_lt (real_of_num 0) x | real_lt (real_of_num 0) (real_neg x)"
- by (import hollight REAL_LT_NEGTOTAL)
-
-lemma REAL_LE_01: "real_le (real_of_num 0) (real_of_num 1)"
- by (import hollight REAL_LE_01)
-
-lemma REAL_LT_01: "real_lt (real_of_num 0) (real_of_num 1)"
- by (import hollight REAL_LT_01)
-
-lemma REAL_LE_LADD: "real_le (real_add x y) (real_add x z) = real_le y z"
- by (import hollight REAL_LE_LADD)
-
-lemma REAL_LE_RADD: "real_le (real_add x z) (real_add y z) = real_le x y"
- by (import hollight REAL_LE_RADD)
-
-lemma REAL_LT_ADD2: "real_lt w x & real_lt y z ==> real_lt (real_add w y) (real_add x z)"
- by (import hollight REAL_LT_ADD2)
-
-lemma REAL_LE_ADD2: "real_le w x & real_le y z ==> real_le (real_add w y) (real_add x z)"
- by (import hollight REAL_LE_ADD2)
-
-lemma REAL_LT_LNEG: "real_lt (real_neg x) xa = real_lt (real_of_num 0) (real_add x xa)"
- by (import hollight REAL_LT_LNEG)
-
-lemma REAL_LT_RNEG: "real_lt x (real_neg xa) = real_lt (real_add x xa) (real_of_num 0)"
- by (import hollight REAL_LT_RNEG)
-
-lemma REAL_LT_ADDNEG: "real_lt y (real_add x (real_neg z)) = real_lt (real_add y z) x"
- by (import hollight REAL_LT_ADDNEG)
-
-lemma REAL_LT_ADDNEG2: "real_lt (real_add x (real_neg y)) z = real_lt x (real_add z y)"
- by (import hollight REAL_LT_ADDNEG2)
-
-lemma REAL_LT_ADD1: "real_le x y ==> real_lt x (real_add y (real_of_num 1))"
- by (import hollight REAL_LT_ADD1)
-
-lemma REAL_SUB_ADD: "real_add (real_sub x y) y = x"
- by (import hollight REAL_SUB_ADD)
-
-lemma REAL_SUB_ADD2: "real_add y (real_sub x y) = x"
- by (import hollight REAL_SUB_ADD2)
-
-lemma REAL_SUB_REFL: "real_sub x x = real_of_num 0"
- by (import hollight REAL_SUB_REFL)
-
-lemma REAL_LE_DOUBLE: "real_le (real_of_num 0) (real_add x x) = real_le (real_of_num 0) x"
- by (import hollight REAL_LE_DOUBLE)
-
-lemma REAL_LE_NEGL: "real_le (real_neg x) x = real_le (real_of_num 0) x"
- by (import hollight REAL_LE_NEGL)
-
-lemma REAL_LE_NEGR: "real_le x (real_neg x) = real_le x (real_of_num 0)"
- by (import hollight REAL_LE_NEGR)
-
-lemma REAL_NEG_EQ_0: "(real_neg x = real_of_num 0) = (x = real_of_num 0)"
- by (import hollight REAL_NEG_EQ_0)
-
-lemma REAL_ADD_SUB: "real_sub (real_add x y) x = y"
- by (import hollight REAL_ADD_SUB)
-
-lemma REAL_NEG_EQ: "(real_neg x = y) = (x = real_neg y)"
- by (import hollight REAL_NEG_EQ)
-
-lemma REAL_NEG_MINUS1: "real_neg x = real_mul (real_neg (real_of_num 1)) x"
- by (import hollight REAL_NEG_MINUS1)
-
-lemma REAL_LT_IMP_NE: "real_lt x y ==> x ~= y"
- by (import hollight REAL_LT_IMP_NE)
-
-lemma REAL_LE_ADDR: "real_le x (real_add x y) = real_le (real_of_num 0) y"
- by (import hollight REAL_LE_ADDR)
-
-lemma REAL_LE_ADDL: "real_le y (real_add x y) = real_le (real_of_num 0) x"
- by (import hollight REAL_LE_ADDL)
-
-lemma REAL_LT_ADDR: "real_lt x (real_add x y) = real_lt (real_of_num 0) y"
- by (import hollight REAL_LT_ADDR)
-
-lemma REAL_LT_ADDL: "real_lt y (real_add x y) = real_lt (real_of_num 0) x"
- by (import hollight REAL_LT_ADDL)
-
-lemma REAL_SUB_SUB: "real_sub (real_sub x y) x = real_neg y"
- by (import hollight REAL_SUB_SUB)
-
-lemma REAL_LT_ADD_SUB: "real_lt (real_add x y) z = real_lt x (real_sub z y)"
- by (import hollight REAL_LT_ADD_SUB)
-
-lemma REAL_LT_SUB_RADD: "real_lt (real_sub x y) z = real_lt x (real_add z y)"
- by (import hollight REAL_LT_SUB_RADD)
-
-lemma REAL_LT_SUB_LADD: "real_lt x (real_sub y z) = real_lt (real_add x z) y"
- by (import hollight REAL_LT_SUB_LADD)
-
-lemma REAL_LE_SUB_LADD: "real_le x (real_sub y z) = real_le (real_add x z) y"
- by (import hollight REAL_LE_SUB_LADD)
-
-lemma REAL_LE_SUB_RADD: "real_le (real_sub x y) z = real_le x (real_add z y)"
- by (import hollight REAL_LE_SUB_RADD)
-
-lemma REAL_LT_NEG: "real_lt (real_neg x) (real_neg y) = real_lt y x"
- by (import hollight REAL_LT_NEG)
-
-lemma REAL_LE_NEG: "real_le (real_neg x) (real_neg y) = real_le y x"
- by (import hollight REAL_LE_NEG)
-
-lemma REAL_ADD2_SUB2: "real_sub (real_add a b) (real_add c d) =
-real_add (real_sub a c) (real_sub b d)"
- by (import hollight REAL_ADD2_SUB2)
-
-lemma REAL_SUB_LZERO: "real_sub (real_of_num 0) x = real_neg x"
- by (import hollight REAL_SUB_LZERO)
-
-lemma REAL_SUB_RZERO: "real_sub x (real_of_num 0) = x"
- by (import hollight REAL_SUB_RZERO)
-
-lemma REAL_LET_ADD2: "real_le w x & real_lt y z ==> real_lt (real_add w y) (real_add x z)"
- by (import hollight REAL_LET_ADD2)
-
-lemma REAL_LTE_ADD2: "real_lt w x & real_le y z ==> real_lt (real_add w y) (real_add x z)"
- by (import hollight REAL_LTE_ADD2)
-
-lemma REAL_SUB_LNEG: "real_sub (real_neg x) y = real_neg (real_add x y)"
- by (import hollight REAL_SUB_LNEG)
-
-lemma REAL_SUB_RNEG: "real_sub x (real_neg y) = real_add x y"
- by (import hollight REAL_SUB_RNEG)
-
-lemma REAL_SUB_NEG2: "real_sub (real_neg x) (real_neg y) = real_sub y x"
- by (import hollight REAL_SUB_NEG2)
-
-lemma REAL_SUB_TRIANGLE: "real_add (real_sub a b) (real_sub b c) = real_sub a c"
- by (import hollight REAL_SUB_TRIANGLE)
-
-lemma REAL_EQ_SUB_LADD: "(x = real_sub y z) = (real_add x z = y)"
- by (import hollight REAL_EQ_SUB_LADD)
-
-lemma REAL_EQ_SUB_RADD: "(real_sub x y = z) = (x = real_add z y)"
- by (import hollight REAL_EQ_SUB_RADD)
-
-lemma REAL_SUB_SUB2: "real_sub x (real_sub x y) = y"
- by (import hollight REAL_SUB_SUB2)
-
-lemma REAL_ADD_SUB2: "real_sub x (real_add x y) = real_neg y"
- by (import hollight REAL_ADD_SUB2)
-
-lemma REAL_EQ_IMP_LE: "x = y ==> real_le x y"
- by (import hollight REAL_EQ_IMP_LE)
-
-lemma REAL_POS_NZ: "real_lt (real_of_num 0) x ==> x ~= real_of_num 0"
- by (import hollight REAL_POS_NZ)
-
-lemma REAL_DIFFSQ: "real_mul (real_add x y) (real_sub x y) =
-real_sub (real_mul x x) (real_mul y y)"
- by (import hollight REAL_DIFFSQ)
-
-lemma REAL_EQ_NEG2: "(real_neg x = real_neg y) = (x = y)"
- by (import hollight REAL_EQ_NEG2)
-
-lemma REAL_LT_NEG2: "real_lt (real_neg x) (real_neg y) = real_lt y x"
- by (import hollight REAL_LT_NEG2)
-
-lemma REAL_SUB_LDISTRIB: "real_mul x (real_sub y z) = real_sub (real_mul x y) (real_mul x z)"
- by (import hollight REAL_SUB_LDISTRIB)
-
-lemma REAL_SUB_RDISTRIB: "real_mul (real_sub x y) z = real_sub (real_mul x z) (real_mul y z)"
- by (import hollight REAL_SUB_RDISTRIB)
-
-lemma REAL_ABS_ZERO: "(real_abs x = real_of_num 0) = (x = real_of_num 0)"
- by (import hollight REAL_ABS_ZERO)
-
-lemma REAL_ABS_0: "real_abs (real_of_num 0) = real_of_num 0"
- by (import hollight REAL_ABS_0)
-
-lemma REAL_ABS_1: "real_abs (real_of_num 1) = real_of_num 1"
- by (import hollight REAL_ABS_1)
-
-lemma REAL_ABS_TRIANGLE: "real_le (real_abs (real_add x y)) (real_add (real_abs x) (real_abs y))"
- by (import hollight REAL_ABS_TRIANGLE)
-
-lemma REAL_ABS_TRIANGLE_LE: "real_le (real_add (real_abs x) (real_abs (real_sub y x))) z
-==> real_le (real_abs y) z"
- by (import hollight REAL_ABS_TRIANGLE_LE)
-
-lemma REAL_ABS_TRIANGLE_LT: "real_lt (real_add (real_abs x) (real_abs (real_sub y x))) z
-==> real_lt (real_abs y) z"
- by (import hollight REAL_ABS_TRIANGLE_LT)
-
-lemma REAL_ABS_POS: "real_le (real_of_num 0) (real_abs x)"
- by (import hollight REAL_ABS_POS)
-
-lemma REAL_ABS_SUB: "real_abs (real_sub x y) = real_abs (real_sub y x)"
- by (import hollight REAL_ABS_SUB)
-
-lemma REAL_ABS_NZ: "(x ~= real_of_num 0) = real_lt (real_of_num 0) (real_abs x)"
- by (import hollight REAL_ABS_NZ)
-
-lemma REAL_ABS_ABS: "real_abs (real_abs x) = real_abs x"
- by (import hollight REAL_ABS_ABS)
-
-lemma REAL_ABS_LE: "real_le x (real_abs x)"
- by (import hollight REAL_ABS_LE)
-
-lemma REAL_ABS_REFL: "(real_abs x = x) = real_le (real_of_num 0) x"
- by (import hollight REAL_ABS_REFL)
-
-lemma REAL_ABS_BETWEEN: "(real_lt (real_of_num 0) d &
- real_lt (real_sub x d) y & real_lt y (real_add x d)) =
-real_lt (real_abs (real_sub y x)) d"
- by (import hollight REAL_ABS_BETWEEN)
-
-lemma REAL_ABS_BOUND: "real_lt (real_abs (real_sub x y)) d ==> real_lt y (real_add x d)"
- by (import hollight REAL_ABS_BOUND)
-
-lemma REAL_ABS_STILLNZ: "real_lt (real_abs (real_sub x y)) (real_abs y) ==> x ~= real_of_num 0"
- by (import hollight REAL_ABS_STILLNZ)
-
-lemma REAL_ABS_CASES: "x = real_of_num 0 | real_lt (real_of_num 0) (real_abs x)"
- by (import hollight REAL_ABS_CASES)
-
-lemma REAL_ABS_BETWEEN1: "real_lt x z & real_lt (real_abs (real_sub y x)) (real_sub z x)
-==> real_lt y z"
- by (import hollight REAL_ABS_BETWEEN1)
-
-lemma REAL_ABS_SIGN: "real_lt (real_abs (real_sub x y)) y ==> real_lt (real_of_num 0) x"
- by (import hollight REAL_ABS_SIGN)
-
-lemma REAL_ABS_SIGN2: "real_lt (real_abs (real_sub x y)) (real_neg y) ==> real_lt x (real_of_num 0)"
- by (import hollight REAL_ABS_SIGN2)
-
-lemma REAL_ABS_CIRCLE: "real_lt (real_abs h) (real_sub (real_abs y) (real_abs x))
-==> real_lt (real_abs (real_add x h)) (real_abs y)"
- by (import hollight REAL_ABS_CIRCLE)
-
-lemma REAL_SUB_ABS: "real_le (real_sub (real_abs x) (real_abs y)) (real_abs (real_sub x y))"
- by (import hollight REAL_SUB_ABS)
-
-lemma REAL_ABS_SUB_ABS: "real_le (real_abs (real_sub (real_abs x) (real_abs y)))
- (real_abs (real_sub x y))"
- by (import hollight REAL_ABS_SUB_ABS)
-
-lemma REAL_ABS_BETWEEN2: "real_lt x0 y0 &
-real_lt (real_mul (real_of_num 2) (real_abs (real_sub x x0)))
- (real_sub y0 x0) &
-real_lt (real_mul (real_of_num 2) (real_abs (real_sub y y0)))
- (real_sub y0 x0)
-==> real_lt x y"
- by (import hollight REAL_ABS_BETWEEN2)
-
-lemma REAL_ABS_BOUNDS: "real_le (real_abs x) k = (real_le (real_neg k) x & real_le x k)"
- by (import hollight REAL_ABS_BOUNDS)
-
-lemma REAL_BOUNDS_LE: "(real_le (real_neg k) x & real_le x k) = real_le (real_abs x) k"
- by (import hollight REAL_BOUNDS_LE)
-
-lemma REAL_BOUNDS_LT: "(real_lt (real_neg k) x & real_lt x k) = real_lt (real_abs x) k"
- by (import hollight REAL_BOUNDS_LT)
-
-lemma REAL_MIN_MAX: "real_min x y = real_neg (real_max (real_neg x) (real_neg y))"
- by (import hollight REAL_MIN_MAX)
-
-lemma REAL_MAX_MIN: "real_max x y = real_neg (real_min (real_neg x) (real_neg y))"
- by (import hollight REAL_MAX_MIN)
-
-lemma REAL_MAX_MAX: "real_le x (real_max x y) & real_le y (real_max x y)"
- by (import hollight REAL_MAX_MAX)
-
-lemma REAL_MIN_MIN: "real_le (real_min x y) x & real_le (real_min x y) y"
- by (import hollight REAL_MIN_MIN)
-
-lemma REAL_MAX_SYM: "real_max x y = real_max y x"
- by (import hollight REAL_MAX_SYM)
-
-lemma REAL_MIN_SYM: "real_min x y = real_min y x"
- by (import hollight REAL_MIN_SYM)
-
-lemma REAL_LE_MAX: "real_le z (real_max x y) = (real_le z x | real_le z y)"
- by (import hollight REAL_LE_MAX)
-
-lemma REAL_LE_MIN: "real_le z (real_min x y) = (real_le z x & real_le z y)"
- by (import hollight REAL_LE_MIN)
-
-lemma REAL_LT_MAX: "real_lt z (real_max x y) = (real_lt z x | real_lt z y)"
- by (import hollight REAL_LT_MAX)
-
-lemma REAL_LT_MIN: "real_lt z (real_min x y) = (real_lt z x & real_lt z y)"
- by (import hollight REAL_LT_MIN)
-
-lemma REAL_MAX_LE: "real_le (real_max x y) z = (real_le x z & real_le y z)"
- by (import hollight REAL_MAX_LE)
-
-lemma REAL_MIN_LE: "real_le (real_min x y) z = (real_le x z | real_le y z)"
- by (import hollight REAL_MIN_LE)
-
-lemma REAL_MAX_LT: "real_lt (real_max x y) z = (real_lt x z & real_lt y z)"
- by (import hollight REAL_MAX_LT)
-
-lemma REAL_MIN_LT: "real_lt (real_min x y) z = (real_lt x z | real_lt y z)"
- by (import hollight REAL_MIN_LT)
-
-lemma REAL_MAX_ASSOC: "real_max x (real_max y z) = real_max (real_max x y) z"
- by (import hollight REAL_MAX_ASSOC)
-
-lemma REAL_MIN_ASSOC: "real_min x (real_min y z) = real_min (real_min x y) z"
- by (import hollight REAL_MIN_ASSOC)
-
-lemma REAL_MAX_ACI: "real_max x y = real_max y x &
-real_max (real_max x y) z = real_max x (real_max y z) &
-real_max x (real_max y z) = real_max y (real_max x z) &
-real_max x x = x & real_max x (real_max x y) = real_max x y"
- by (import hollight REAL_MAX_ACI)
-
-lemma REAL_MIN_ACI: "real_min x y = real_min y x &
-real_min (real_min x y) z = real_min x (real_min y z) &
-real_min x (real_min y z) = real_min y (real_min x z) &
-real_min x x = x & real_min x (real_min x y) = real_min x y"
- by (import hollight REAL_MIN_ACI)
-
-lemma REAL_ABS_MUL: "real_abs (real_mul x y) = real_mul (real_abs x) (real_abs y)"
- by (import hollight REAL_ABS_MUL)
-
-lemma REAL_POW_LE: "real_le (real_of_num 0) x ==> real_le (real_of_num 0) (real_pow x n)"
- by (import hollight REAL_POW_LE)
-
-lemma REAL_POW_LT: "real_lt (real_of_num 0) x ==> real_lt (real_of_num 0) (real_pow x n)"
- by (import hollight REAL_POW_LT)
-
-lemma REAL_ABS_POW: "real_abs (real_pow x n) = real_pow (real_abs x) n"
- by (import hollight REAL_ABS_POW)
-
-lemma REAL_LE_LMUL: "real_le (real_of_num 0) x & real_le xa xb
-==> real_le (real_mul x xa) (real_mul x xb)"
- by (import hollight REAL_LE_LMUL)
-
-lemma REAL_LE_RMUL: "real_le x y & real_le (real_of_num 0) z
-==> real_le (real_mul x z) (real_mul y z)"
- by (import hollight REAL_LE_RMUL)
-
-lemma REAL_LT_LMUL: "real_lt (real_of_num 0) x & real_lt xa xb
-==> real_lt (real_mul x xa) (real_mul x xb)"
- by (import hollight REAL_LT_LMUL)
-
-lemma REAL_LT_RMUL: "real_lt x y & real_lt (real_of_num 0) z
-==> real_lt (real_mul x z) (real_mul y z)"
- by (import hollight REAL_LT_RMUL)
-
-lemma REAL_EQ_MUL_LCANCEL: "(real_mul x y = real_mul x z) = (x = real_of_num 0 | y = z)"
- by (import hollight REAL_EQ_MUL_LCANCEL)
-
-lemma REAL_EQ_MUL_RCANCEL: "(real_mul x xb = real_mul xa xb) = (x = xa | xb = real_of_num 0)"
- by (import hollight REAL_EQ_MUL_RCANCEL)
-
-lemma REAL_MUL_LINV_UNIQ: "real_mul x y = real_of_num 1 ==> real_inv y = x"
- by (import hollight REAL_MUL_LINV_UNIQ)
-
-lemma REAL_MUL_RINV_UNIQ: "real_mul x xa = real_of_num 1 ==> real_inv x = xa"
- by (import hollight REAL_MUL_RINV_UNIQ)
-
-lemma REAL_INV_INV: "real_inv (real_inv x) = x"
- by (import hollight REAL_INV_INV)
-
-lemma REAL_EQ_INV2: "(real_inv x = real_inv y) = (x = y)"
- by (import hollight REAL_EQ_INV2)
-
-lemma REAL_INV_EQ_0: "(real_inv x = real_of_num 0) = (x = real_of_num 0)"
- by (import hollight REAL_INV_EQ_0)
-
-lemma REAL_LT_INV: "real_lt (real_of_num 0) x ==> real_lt (real_of_num 0) (real_inv x)"
- by (import hollight REAL_LT_INV)
-
-lemma REAL_LT_INV_EQ: "real_lt (real_of_num 0) (real_inv x) = real_lt (real_of_num 0) x"
- by (import hollight REAL_LT_INV_EQ)
-
-lemma REAL_INV_NEG: "real_inv (real_neg x) = real_neg (real_inv x)"
- by (import hollight REAL_INV_NEG)
-
-lemma REAL_LE_INV_EQ: "real_le (real_of_num 0) (real_inv x) = real_le (real_of_num 0) x"
- by (import hollight REAL_LE_INV_EQ)
-
-lemma REAL_LE_INV: "real_le (real_of_num 0) x ==> real_le (real_of_num 0) (real_inv x)"
- by (import hollight REAL_LE_INV)
-
-lemma REAL_MUL_RINV: "x ~= real_of_num 0 ==> real_mul x (real_inv x) = real_of_num 1"
- by (import hollight REAL_MUL_RINV)
-
-lemma REAL_INV_1: "real_inv (real_of_num 1) = real_of_num 1"
- by (import hollight REAL_INV_1)
-
-lemma REAL_INV_EQ_1: "(real_inv x = real_of_num 1) = (x = real_of_num 1)"
- by (import hollight REAL_INV_EQ_1)
-
-lemma REAL_DIV_1: "real_div x (real_of_num 1) = x"
- by (import hollight REAL_DIV_1)
-
-lemma REAL_DIV_REFL: "x ~= real_of_num 0 ==> real_div x x = real_of_num 1"
- by (import hollight REAL_DIV_REFL)
-
-lemma REAL_DIV_RMUL: "xa ~= real_of_num 0 ==> real_mul (real_div x xa) xa = x"
- by (import hollight REAL_DIV_RMUL)
-
-lemma REAL_DIV_LMUL: "xa ~= real_of_num 0 ==> real_mul xa (real_div x xa) = x"
- by (import hollight REAL_DIV_LMUL)
-
-lemma REAL_ABS_INV: "real_abs (real_inv x) = real_inv (real_abs x)"
- by (import hollight REAL_ABS_INV)
-
-lemma REAL_ABS_DIV: "real_abs (real_div x xa) = real_div (real_abs x) (real_abs xa)"
- by (import hollight REAL_ABS_DIV)
-
-lemma REAL_INV_MUL: "real_inv (real_mul x y) = real_mul (real_inv x) (real_inv y)"
- by (import hollight REAL_INV_MUL)
-
-lemma REAL_INV_DIV: "real_inv (real_div x xa) = real_div xa x"
- by (import hollight REAL_INV_DIV)
-
-lemma REAL_POW_MUL: "real_pow (real_mul x y) n = real_mul (real_pow x n) (real_pow y n)"
- by (import hollight REAL_POW_MUL)
-
-lemma REAL_POW_INV: "real_pow (real_inv x) n = real_inv (real_pow x n)"
- by (import hollight REAL_POW_INV)
-
-lemma REAL_INV_POW: "real_inv (real_pow x xa) = real_pow (real_inv x) xa"
- by (import hollight REAL_INV_POW)
-
-lemma REAL_POW_DIV: "real_pow (real_div x xa) xb = real_div (real_pow x xb) (real_pow xa xb)"
- by (import hollight REAL_POW_DIV)
-
-lemma REAL_POW_ADD: "real_pow x (m + n) = real_mul (real_pow x m) (real_pow x n)"
- by (import hollight REAL_POW_ADD)
-
-lemma REAL_POW_NZ: "x ~= real_of_num 0 ==> real_pow x n ~= real_of_num 0"
- by (import hollight REAL_POW_NZ)
-
-lemma REAL_POW_SUB: "x ~= real_of_num 0 & m <= n
-==> real_pow x (n - m) = real_div (real_pow x n) (real_pow x m)"
- by (import hollight REAL_POW_SUB)
-
-lemma REAL_LT_IMP_NZ: "real_lt (real_of_num 0) x ==> x ~= real_of_num 0"
- by (import hollight REAL_LT_IMP_NZ)
-
-lemma REAL_LT_LCANCEL_IMP: "real_lt (real_of_num 0) x & real_lt (real_mul x y) (real_mul x z)
-==> real_lt y z"
- by (import hollight REAL_LT_LCANCEL_IMP)
-
-lemma REAL_LT_RCANCEL_IMP: "real_lt (real_of_num 0) xb & real_lt (real_mul x xb) (real_mul xa xb)
-==> real_lt x xa"
- by (import hollight REAL_LT_RCANCEL_IMP)
-
-lemma REAL_LE_LCANCEL_IMP: "real_lt (real_of_num 0) x & real_le (real_mul x y) (real_mul x z)
-==> real_le y z"
- by (import hollight REAL_LE_LCANCEL_IMP)
-
-lemma REAL_LE_RCANCEL_IMP: "real_lt (real_of_num 0) xb & real_le (real_mul x xb) (real_mul xa xb)
-==> real_le x xa"
- by (import hollight REAL_LE_RCANCEL_IMP)
-
-lemma REAL_LE_RMUL_EQ: "real_lt (real_of_num 0) z
-==> real_le (real_mul x z) (real_mul y z) = real_le x y"
- by (import hollight REAL_LE_RMUL_EQ)
-
-lemma REAL_LE_LMUL_EQ: "real_lt (real_of_num 0) z
-==> real_le (real_mul z x) (real_mul z y) = real_le x y"
- by (import hollight REAL_LE_LMUL_EQ)
-
-lemma REAL_LT_RMUL_EQ: "real_lt (real_of_num 0) xb
-==> real_lt (real_mul x xb) (real_mul xa xb) = real_lt x xa"
- by (import hollight REAL_LT_RMUL_EQ)
-
-lemma REAL_LT_LMUL_EQ: "real_lt (real_of_num 0) xb
-==> real_lt (real_mul xb x) (real_mul xb xa) = real_lt x xa"
- by (import hollight REAL_LT_LMUL_EQ)
-
-lemma REAL_LE_MUL_EQ: "(ALL x y.
- real_lt (real_of_num 0) x -->
- real_le (real_of_num 0) (real_mul x y) = real_le (real_of_num 0) y) &
-(ALL x y.
- real_lt (real_of_num 0) y -->
- real_le (real_of_num 0) (real_mul x y) = real_le (real_of_num 0) x)"
- by (import hollight REAL_LE_MUL_EQ)
-
-lemma REAL_LT_MUL_EQ: "(ALL x y.
- real_lt (real_of_num 0) x -->
- real_lt (real_of_num 0) (real_mul x y) = real_lt (real_of_num 0) y) &
-(ALL x y.
- real_lt (real_of_num 0) y -->
- real_lt (real_of_num 0) (real_mul x y) = real_lt (real_of_num 0) x)"
- by (import hollight REAL_LT_MUL_EQ)
-
-lemma REAL_MUL_POS_LT: "real_lt (real_of_num 0) (real_mul x y) =
-(real_lt (real_of_num 0) x & real_lt (real_of_num 0) y |
- real_lt x (real_of_num 0) & real_lt y (real_of_num 0))"
- by (import hollight REAL_MUL_POS_LT)
-
-lemma REAL_MUL_POS_LE: "real_le (real_of_num 0) (real_mul x xa) =
-(x = real_of_num 0 |
- xa = real_of_num 0 |
- real_lt (real_of_num 0) x & real_lt (real_of_num 0) xa |
- real_lt x (real_of_num 0) & real_lt xa (real_of_num 0))"
- by (import hollight REAL_MUL_POS_LE)
-
-lemma REAL_LE_RDIV_EQ: "real_lt (real_of_num 0) z
-==> real_le x (real_div y z) = real_le (real_mul x z) y"
- by (import hollight REAL_LE_RDIV_EQ)
-
-lemma REAL_LE_LDIV_EQ: "real_lt (real_of_num 0) z
-==> real_le (real_div x z) y = real_le x (real_mul y z)"
- by (import hollight REAL_LE_LDIV_EQ)
-
-lemma REAL_LT_RDIV_EQ: "real_lt (real_of_num 0) xb
-==> real_lt x (real_div xa xb) = real_lt (real_mul x xb) xa"
- by (import hollight REAL_LT_RDIV_EQ)
-
-lemma REAL_LT_LDIV_EQ: "real_lt (real_of_num 0) xb
-==> real_lt (real_div x xb) xa = real_lt x (real_mul xa xb)"
- by (import hollight REAL_LT_LDIV_EQ)
-
-lemma REAL_EQ_RDIV_EQ: "real_lt (real_of_num 0) xb ==> (x = real_div xa xb) = (real_mul x xb = xa)"
- by (import hollight REAL_EQ_RDIV_EQ)
-
-lemma REAL_EQ_LDIV_EQ: "real_lt (real_of_num 0) xb ==> (real_div x xb = xa) = (x = real_mul xa xb)"
- by (import hollight REAL_EQ_LDIV_EQ)
-
-lemma REAL_LT_DIV2_EQ: "real_lt (real_of_num 0) xb
-==> real_lt (real_div x xb) (real_div xa xb) = real_lt x xa"
- by (import hollight REAL_LT_DIV2_EQ)
-
-lemma REAL_LE_DIV2_EQ: "real_lt (real_of_num 0) xb
-==> real_le (real_div x xb) (real_div xa xb) = real_le x xa"
- by (import hollight REAL_LE_DIV2_EQ)
-
-lemma REAL_MUL_2: "real_mul (real_of_num 2) x = real_add x x"
- by (import hollight REAL_MUL_2)
-
-lemma REAL_POW_EQ_0: "(real_pow x n = real_of_num 0) = (x = real_of_num 0 & n ~= 0)"
- by (import hollight REAL_POW_EQ_0)
-
-lemma REAL_LE_MUL2: "real_le (real_of_num 0) w &
-real_le w x & real_le (real_of_num 0) y & real_le y z
-==> real_le (real_mul w y) (real_mul x z)"
- by (import hollight REAL_LE_MUL2)
-
-lemma REAL_LT_MUL2: "real_le (real_of_num 0) w &
-real_lt w x & real_le (real_of_num 0) y & real_lt y z
-==> real_lt (real_mul w y) (real_mul x z)"
- by (import hollight REAL_LT_MUL2)
-
-lemma REAL_LT_SQUARE: "real_lt (real_of_num 0) (real_mul x x) = (x ~= real_of_num 0)"
- by (import hollight REAL_LT_SQUARE)
-
-lemma REAL_POW_1: "real_pow x 1 = x"
- by (import hollight REAL_POW_1)
-
-lemma REAL_POW_ONE: "real_pow (real_of_num 1) n = real_of_num 1"
- by (import hollight REAL_POW_ONE)
-
-lemma REAL_LT_INV2: "real_lt (real_of_num 0) x & real_lt x y
-==> real_lt (real_inv y) (real_inv x)"
- by (import hollight REAL_LT_INV2)
-
-lemma REAL_LE_INV2: "real_lt (real_of_num 0) x & real_le x y
-==> real_le (real_inv y) (real_inv x)"
- by (import hollight REAL_LE_INV2)
-
-lemma REAL_LT_LINV: "real_lt (real_of_num 0) y & real_lt (real_inv y) x
-==> real_lt (real_inv x) y"
- by (import hollight REAL_LT_LINV)
-
-lemma REAL_LT_RINV: "real_lt (real_of_num 0) x & real_lt x (real_inv y)
-==> real_lt y (real_inv x)"
- by (import hollight REAL_LT_RINV)
-
-lemma REAL_LE_LINV: "real_lt (real_of_num 0) y & real_le (real_inv y) x
-==> real_le (real_inv x) y"
- by (import hollight REAL_LE_LINV)
-
-lemma REAL_LE_RINV: "real_lt (real_of_num 0) x & real_le x (real_inv y)
-==> real_le y (real_inv x)"
- by (import hollight REAL_LE_RINV)
-
-lemma REAL_INV_LE_1: "real_le (real_of_num 1) x ==> real_le (real_inv x) (real_of_num 1)"
- by (import hollight REAL_INV_LE_1)
-
-lemma REAL_INV_1_LE: "real_lt (real_of_num 0) x & real_le x (real_of_num 1)
-==> real_le (real_of_num 1) (real_inv x)"
- by (import hollight REAL_INV_1_LE)
-
-lemma REAL_INV_LT_1: "real_lt (real_of_num 1) x ==> real_lt (real_inv x) (real_of_num 1)"
- by (import hollight REAL_INV_LT_1)
-
-lemma REAL_INV_1_LT: "real_lt (real_of_num 0) x & real_lt x (real_of_num 1)
-==> real_lt (real_of_num 1) (real_inv x)"
- by (import hollight REAL_INV_1_LT)
-
-lemma REAL_SUB_INV: "x ~= real_of_num 0 & xa ~= real_of_num 0
-==> real_sub (real_inv x) (real_inv xa) =
- real_div (real_sub xa x) (real_mul x xa)"
- by (import hollight REAL_SUB_INV)
-
-lemma REAL_DOWN: "real_lt (real_of_num 0) d ==> EX x. real_lt (real_of_num 0) x & real_lt x d"
- by (import hollight REAL_DOWN)
-
-lemma REAL_DOWN2: "real_lt (real_of_num 0) d1 & real_lt (real_of_num 0) d2
-==> EX e. real_lt (real_of_num 0) e & real_lt e d1 & real_lt e d2"
- by (import hollight REAL_DOWN2)
-
-lemma REAL_POW_LE2: "real_le (real_of_num 0) x & real_le x y
-==> real_le (real_pow x n) (real_pow y n)"
- by (import hollight REAL_POW_LE2)
-
-lemma REAL_POW_LE_1: "real_le (real_of_num 1) x ==> real_le (real_of_num 1) (real_pow x n)"
- by (import hollight REAL_POW_LE_1)
-
-lemma REAL_POW_1_LE: "real_le (real_of_num 0) x & real_le x (real_of_num 1)
-==> real_le (real_pow x n) (real_of_num 1)"
- by (import hollight REAL_POW_1_LE)
-
-lemma REAL_POW_MONO: "real_le (real_of_num 1) x & m <= n ==> real_le (real_pow x m) (real_pow x n)"
- by (import hollight REAL_POW_MONO)
-
-lemma REAL_POW_LT2: "n ~= 0 & real_le (real_of_num 0) x & real_lt x y
-==> real_lt (real_pow x n) (real_pow y n)"
- by (import hollight REAL_POW_LT2)
-
-lemma REAL_POW_LT_1: "n ~= 0 & real_lt (real_of_num 1) x
-==> real_lt (real_of_num 1) (real_pow x n)"
- by (import hollight REAL_POW_LT_1)
-
-lemma REAL_POW_1_LT: "n ~= 0 & real_le (real_of_num 0) x & real_lt x (real_of_num 1)
-==> real_lt (real_pow x n) (real_of_num 1)"
- by (import hollight REAL_POW_1_LT)
-
-lemma REAL_POW_MONO_LT: "real_lt (real_of_num 1) x & m < n ==> real_lt (real_pow x m) (real_pow x n)"
- by (import hollight REAL_POW_MONO_LT)
-
-lemma REAL_POW_POW: "real_pow (real_pow x m) n = real_pow x (m * n)"
- by (import hollight REAL_POW_POW)
-
-lemma REAL_EQ_RCANCEL_IMP: "z ~= real_of_num 0 & real_mul x z = real_mul y z ==> x = y"
- by (import hollight REAL_EQ_RCANCEL_IMP)
-
-lemma REAL_EQ_LCANCEL_IMP: "xb ~= real_of_num 0 & real_mul xb x = real_mul xb xa ==> x = xa"
- by (import hollight REAL_EQ_LCANCEL_IMP)
-
-lemma REAL_LT_DIV: "real_lt (real_of_num 0) x & real_lt (real_of_num 0) xa
-==> real_lt (real_of_num 0) (real_div x xa)"
- by (import hollight REAL_LT_DIV)
-
-lemma REAL_LE_DIV: "real_le (real_of_num 0) x & real_le (real_of_num 0) xa
-==> real_le (real_of_num 0) (real_div x xa)"
- by (import hollight REAL_LE_DIV)
-
-lemma REAL_DIV_POW2: "x ~= real_of_num 0
-==> real_div (real_pow x m) (real_pow x n) =
- (if n <= m then real_pow x (m - n) else real_inv (real_pow x (n - m)))"
- by (import hollight REAL_DIV_POW2)
-
-lemma REAL_DIV_POW2_ALT: "x ~= real_of_num 0
-==> real_div (real_pow x m) (real_pow x n) =
- (if n < m then real_pow x (m - n) else real_inv (real_pow x (n - m)))"
- by (import hollight REAL_DIV_POW2_ALT)
-
-lemma REAL_LT_POW2: "real_lt (real_of_num 0) (real_pow (real_of_num 2) x)"
- by (import hollight REAL_LT_POW2)
-
-lemma REAL_LE_POW2: "real_le (real_of_num 1) (real_pow (real_of_num 2) n)"
- by (import hollight REAL_LE_POW2)
-
-lemma REAL_POW2_ABS: "real_pow (real_abs x) 2 = real_pow x 2"
- by (import hollight REAL_POW2_ABS)
-
-lemma REAL_LE_SQUARE_ABS: "real_le (real_abs x) (real_abs y) = real_le (real_pow x 2) (real_pow y 2)"
- by (import hollight REAL_LE_SQUARE_ABS)
-
-lemma REAL_LT_SQUARE_ABS: "real_lt (real_abs x) (real_abs xa) = real_lt (real_pow x 2) (real_pow xa 2)"
- by (import hollight REAL_LT_SQUARE_ABS)
-
-lemma REAL_EQ_SQUARE_ABS: "(real_abs x = real_abs xa) = (real_pow x 2 = real_pow xa 2)"
- by (import hollight REAL_EQ_SQUARE_ABS)
-
-lemma REAL_LE_POW_2: "real_le (real_of_num 0) (real_pow x 2)"
- by (import hollight REAL_LE_POW_2)
-
-lemma REAL_SOS_EQ_0: "(real_add (real_pow x 2) (real_pow y 2) = real_of_num 0) =
-(x = real_of_num 0 & y = real_of_num 0)"
- by (import hollight REAL_SOS_EQ_0)
-
-lemma REAL_POW_ZERO: "real_pow (real_of_num 0) n =
-(if n = 0 then real_of_num 1 else real_of_num 0)"
- by (import hollight REAL_POW_ZERO)
-
-lemma REAL_POW_MONO_INV: "real_le (real_of_num 0) x & real_le x (real_of_num 1) & n <= m
-==> real_le (real_pow x m) (real_pow x n)"
- by (import hollight REAL_POW_MONO_INV)
-
-lemma REAL_POW_LE2_REV: "n ~= 0 & real_le (real_of_num 0) y & real_le (real_pow x n) (real_pow y n)
-==> real_le x y"
- by (import hollight REAL_POW_LE2_REV)
-
-lemma REAL_POW_LT2_REV: "real_le (real_of_num 0) y & real_lt (real_pow x n) (real_pow y n)
-==> real_lt x y"
- by (import hollight REAL_POW_LT2_REV)
-
-lemma REAL_POW_EQ: "x ~= 0 &
-real_le (real_of_num 0) xa &
-real_le (real_of_num 0) xb & real_pow xa x = real_pow xb x
-==> xa = xb"
- by (import hollight REAL_POW_EQ)
-
-lemma REAL_POW_EQ_ABS: "n ~= 0 & real_pow x n = real_pow y n ==> real_abs x = real_abs y"
- by (import hollight REAL_POW_EQ_ABS)
-
-lemma REAL_POW_EQ_1_IMP: "n ~= 0 & real_pow x n = real_of_num 1 ==> real_abs x = real_of_num 1"
- by (import hollight REAL_POW_EQ_1_IMP)
-
-lemma REAL_POW_EQ_1: "(real_pow x n = real_of_num 1) =
-(real_abs x = real_of_num 1 & (real_lt x (real_of_num 0) --> even n) |
- n = 0)"
- by (import hollight REAL_POW_EQ_1)
-
-lemma REAL_POW_LT2_ODD: "real_lt x y & odd n ==> real_lt (real_pow x n) (real_pow y n)"
- by (import hollight REAL_POW_LT2_ODD)
-
-lemma REAL_POW_LE2_ODD: "real_le xa xb & odd x ==> real_le (real_pow xa x) (real_pow xb x)"
- by (import hollight REAL_POW_LE2_ODD)
-
-lemma REAL_POW_LT2_ODD_EQ: "odd n ==> real_lt (real_pow x n) (real_pow y n) = real_lt x y"
- by (import hollight REAL_POW_LT2_ODD_EQ)
-
-lemma REAL_POW_LE2_ODD_EQ: "odd n ==> real_le (real_pow x n) (real_pow y n) = real_le x y"
- by (import hollight REAL_POW_LE2_ODD_EQ)
-
-lemma REAL_POW_EQ_ODD_EQ: "odd x ==> (real_pow xa x = real_pow xb x) = (xa = xb)"
- by (import hollight REAL_POW_EQ_ODD_EQ)
-
-lemma REAL_POW_EQ_ODD: "odd n & real_pow x n = real_pow y n ==> x = y"
- by (import hollight REAL_POW_EQ_ODD)
-
-lemma REAL_POW_EQ_EQ: "(real_pow x n = real_pow y n) =
-(if even n then n = 0 | real_abs x = real_abs y else x = y)"
- by (import hollight REAL_POW_EQ_EQ)
-
-definition
- real_sgn :: "hollight.real => hollight.real" where
- "real_sgn ==
-%u. if real_lt (real_of_num 0) u then real_of_num 1
- else if real_lt u (real_of_num 0) then real_neg (real_of_num 1)
- else real_of_num 0"
-
-lemma DEF_real_sgn: "real_sgn =
-(%u. if real_lt (real_of_num 0) u then real_of_num 1
- else if real_lt u (real_of_num 0) then real_neg (real_of_num 1)
- else real_of_num 0)"
- by (import hollight DEF_real_sgn)
-
-lemma REAL_SGN_0: "real_sgn (real_of_num 0) = real_of_num 0"
- by (import hollight REAL_SGN_0)
-
-lemma REAL_SGN_NEG: "real_sgn (real_neg x) = real_neg (real_sgn x)"
- by (import hollight REAL_SGN_NEG)
-
-lemma REAL_SGN_ABS: "real_mul (real_sgn x) (real_abs x) = x"
- by (import hollight REAL_SGN_ABS)
-
-lemma REAL_ABS_SGN: "real_abs (real_sgn x) = real_sgn (real_abs x)"
- by (import hollight REAL_ABS_SGN)
-
-lemma REAL_SGN: "real_sgn x = real_div x (real_abs x)"
- by (import hollight REAL_SGN)
-
-lemma REAL_SGN_MUL: "real_sgn (real_mul x xa) = real_mul (real_sgn x) (real_sgn xa)"
- by (import hollight REAL_SGN_MUL)
-
-lemma REAL_SGN_INV: "real_sgn (real_inv x) = real_sgn x"
- by (import hollight REAL_SGN_INV)
-
-lemma REAL_SGN_DIV: "real_sgn (real_div x xa) = real_div (real_sgn x) (real_sgn xa)"
- by (import hollight REAL_SGN_DIV)
-
-lemma REAL_WLOG_LE: "(ALL x y. P x y = P y x) & (ALL x y. real_le x y --> P x y) ==> P x xa"
- by (import hollight REAL_WLOG_LE)
-
-lemma REAL_WLOG_LT: "(ALL x. P x x) & (ALL x y. P x y = P y x) & (ALL x y. real_lt x y --> P x y)
-==> P x xa"
- by (import hollight REAL_WLOG_LT)
-
-definition
- DECIMAL :: "nat => nat => hollight.real" where
- "DECIMAL == %u ua. real_div (real_of_num u) (real_of_num ua)"
-
-lemma DEF_DECIMAL: "DECIMAL = (%u ua. real_div (real_of_num u) (real_of_num ua))"
- by (import hollight DEF_DECIMAL)
-
-lemma RAT_LEMMA1: "y1 ~= real_of_num 0 & y2 ~= real_of_num 0
-==> real_add (real_div x1 y1) (real_div x2 y2) =
- real_mul (real_add (real_mul x1 y2) (real_mul x2 y1))
- (real_mul (real_inv y1) (real_inv y2))"
- by (import hollight RAT_LEMMA1)
-
-lemma RAT_LEMMA2: "real_lt (real_of_num 0) y1 & real_lt (real_of_num 0) y2
-==> real_add (real_div x1 y1) (real_div x2 y2) =
- real_mul (real_add (real_mul x1 y2) (real_mul x2 y1))
- (real_mul (real_inv y1) (real_inv y2))"
- by (import hollight RAT_LEMMA2)
-
-lemma RAT_LEMMA3: "real_lt (real_of_num 0) y1 & real_lt (real_of_num 0) y2
-==> real_sub (real_div x1 y1) (real_div x2 y2) =
- real_mul (real_sub (real_mul x1 y2) (real_mul x2 y1))
- (real_mul (real_inv y1) (real_inv y2))"
- by (import hollight RAT_LEMMA3)
-
-lemma RAT_LEMMA4: "real_lt (real_of_num 0) y1 & real_lt (real_of_num 0) y2
-==> real_le (real_div x1 y1) (real_div x2 y2) =
- real_le (real_mul x1 y2) (real_mul x2 y1)"
- by (import hollight RAT_LEMMA4)
-
-lemma RAT_LEMMA5: "real_lt (real_of_num 0) y1 & real_lt (real_of_num 0) y2
-==> (real_div x1 y1 = real_div x2 y2) = (real_mul x1 y2 = real_mul x2 y1)"
- by (import hollight RAT_LEMMA5)
-
-lemma REAL_INTEGRAL: "(ALL x. real_mul (real_of_num 0) x = real_of_num 0) &
-(ALL x y z. (real_add x y = real_add x z) = (y = z)) &
-(ALL w x y z.
- (real_add (real_mul w y) (real_mul x z) =
- real_add (real_mul w z) (real_mul x y)) =
- (w = x | y = z))"
- by (import hollight REAL_INTEGRAL)
-
-definition
- integer :: "hollight.real => bool" where
- "integer == %u. EX n. real_abs u = real_of_num n"
-
-lemma DEF_integer: "integer = (%u. EX n. real_abs u = real_of_num n)"
- by (import hollight DEF_integer)
-
-lemma is_int: "integer x = (EX n. x = real_of_num n | x = real_neg (real_of_num n))"
- by (import hollight is_int)
-
-typedef (open) int = "Collect integer" morphisms "real_of_int" "int_of_real"
- apply (rule light_ex_imp_nonempty[where t="Eps integer"])
- by (import hollight TYDEF_int)
-
-syntax
- real_of_int :: _
-
-syntax
- int_of_real :: _
-
-lemmas "TYDEF_int_@intern" = typedef_hol2hollight
- [where a="a :: hollight.int" and r=r ,
- OF type_definition_int]
-
-lemma dest_int_rep: "EX n. hollight.real_of_int x = real_of_num n |
- hollight.real_of_int x = real_neg (real_of_num n)"
- by (import hollight dest_int_rep)
-
-definition
- int_le :: "hollight.int => hollight.int => bool" where
- "int_le == %u ua. real_le (hollight.real_of_int u) (hollight.real_of_int ua)"
-
-lemma DEF_int_le: "int_le = (%u ua. real_le (hollight.real_of_int u) (hollight.real_of_int ua))"
- by (import hollight DEF_int_le)
-
-definition
- int_lt :: "hollight.int => hollight.int => bool" where
- "int_lt == %u ua. real_lt (hollight.real_of_int u) (hollight.real_of_int ua)"
-
-lemma DEF_int_lt: "int_lt = (%u ua. real_lt (hollight.real_of_int u) (hollight.real_of_int ua))"
- by (import hollight DEF_int_lt)
-
-definition
- int_ge :: "hollight.int => hollight.int => bool" where
- "int_ge == %u ua. real_ge (hollight.real_of_int u) (hollight.real_of_int ua)"
-
-lemma DEF_int_ge: "int_ge = (%u ua. real_ge (hollight.real_of_int u) (hollight.real_of_int ua))"
- by (import hollight DEF_int_ge)
-
-definition
- int_gt :: "hollight.int => hollight.int => bool" where
- "int_gt == %u ua. real_gt (hollight.real_of_int u) (hollight.real_of_int ua)"
-
-lemma DEF_int_gt: "int_gt = (%u ua. real_gt (hollight.real_of_int u) (hollight.real_of_int ua))"
- by (import hollight DEF_int_gt)
-
-definition
- int_of_num :: "nat => hollight.int" where
- "int_of_num == %u. int_of_real (real_of_num u)"
-
-lemma DEF_int_of_num: "int_of_num = (%u. int_of_real (real_of_num u))"
- by (import hollight DEF_int_of_num)
-
-lemma int_of_num_th: "hollight.real_of_int (int_of_num x) = real_of_num x"
- by (import hollight int_of_num_th)
-
-definition
- int_neg :: "hollight.int => hollight.int" where
- "int_neg == %u. int_of_real (real_neg (hollight.real_of_int u))"
-
-lemma DEF_int_neg: "int_neg = (%u. int_of_real (real_neg (hollight.real_of_int u)))"
- by (import hollight DEF_int_neg)
-
-lemma int_neg_th: "hollight.real_of_int (int_neg x) = real_neg (hollight.real_of_int x)"
- by (import hollight int_neg_th)
-
-definition
- int_add :: "hollight.int => hollight.int => hollight.int" where
- "int_add ==
-%u ua.
- int_of_real (real_add (hollight.real_of_int u) (hollight.real_of_int ua))"
-
-lemma DEF_int_add: "int_add =
-(%u ua.
- int_of_real
- (real_add (hollight.real_of_int u) (hollight.real_of_int ua)))"
- by (import hollight DEF_int_add)
-
-lemma int_add_th: "hollight.real_of_int (int_add x xa) =
-real_add (hollight.real_of_int x) (hollight.real_of_int xa)"
- by (import hollight int_add_th)
-
-definition
- int_sub :: "hollight.int => hollight.int => hollight.int" where
- "int_sub ==
-%u ua.
- int_of_real (real_sub (hollight.real_of_int u) (hollight.real_of_int ua))"
-
-lemma DEF_int_sub: "int_sub =
-(%u ua.
- int_of_real
- (real_sub (hollight.real_of_int u) (hollight.real_of_int ua)))"
- by (import hollight DEF_int_sub)
-
-lemma int_sub_th: "hollight.real_of_int (int_sub x xa) =
-real_sub (hollight.real_of_int x) (hollight.real_of_int xa)"
- by (import hollight int_sub_th)
-
-definition
- int_mul :: "hollight.int => hollight.int => hollight.int" where
- "int_mul ==
-%u ua.
- int_of_real (real_mul (hollight.real_of_int u) (hollight.real_of_int ua))"
-
-lemma DEF_int_mul: "int_mul =
-(%u ua.
- int_of_real
- (real_mul (hollight.real_of_int u) (hollight.real_of_int ua)))"
- by (import hollight DEF_int_mul)
-
-lemma int_mul_th: "hollight.real_of_int (int_mul x y) =
-real_mul (hollight.real_of_int x) (hollight.real_of_int y)"
- by (import hollight int_mul_th)
-
-definition
- int_abs :: "hollight.int => hollight.int" where
- "int_abs == %u. int_of_real (real_abs (hollight.real_of_int u))"
-
-lemma DEF_int_abs: "int_abs = (%u. int_of_real (real_abs (hollight.real_of_int u)))"
- by (import hollight DEF_int_abs)
-
-lemma int_abs_th: "hollight.real_of_int (int_abs x) = real_abs (hollight.real_of_int x)"
- by (import hollight int_abs_th)
-
-definition
- int_sgn :: "hollight.int => hollight.int" where
- "int_sgn == %u. int_of_real (real_sgn (hollight.real_of_int u))"
-
-lemma DEF_int_sgn: "int_sgn = (%u. int_of_real (real_sgn (hollight.real_of_int u)))"
- by (import hollight DEF_int_sgn)
-
-lemma int_sgn_th: "hollight.real_of_int (int_sgn x) = real_sgn (hollight.real_of_int x)"
- by (import hollight int_sgn_th)
-
-definition
- int_max :: "hollight.int => hollight.int => hollight.int" where
- "int_max ==
-%u ua.
- int_of_real (real_max (hollight.real_of_int u) (hollight.real_of_int ua))"
-
-lemma DEF_int_max: "int_max =
-(%u ua.
- int_of_real
- (real_max (hollight.real_of_int u) (hollight.real_of_int ua)))"
- by (import hollight DEF_int_max)
-
-lemma int_max_th: "hollight.real_of_int (int_max x y) =
-real_max (hollight.real_of_int x) (hollight.real_of_int y)"
- by (import hollight int_max_th)
-
-definition
- int_min :: "hollight.int => hollight.int => hollight.int" where
- "int_min ==
-%u ua.
- int_of_real (real_min (hollight.real_of_int u) (hollight.real_of_int ua))"
-
-lemma DEF_int_min: "int_min =
-(%u ua.
- int_of_real
- (real_min (hollight.real_of_int u) (hollight.real_of_int ua)))"
- by (import hollight DEF_int_min)
-
-lemma int_min_th: "hollight.real_of_int (int_min x y) =
-real_min (hollight.real_of_int x) (hollight.real_of_int y)"
- by (import hollight int_min_th)
-
-definition
- int_pow :: "hollight.int => nat => hollight.int" where
- "int_pow == %u ua. int_of_real (real_pow (hollight.real_of_int u) ua)"
-
-lemma DEF_int_pow: "int_pow = (%u ua. int_of_real (real_pow (hollight.real_of_int u) ua))"
- by (import hollight DEF_int_pow)
-
-lemma int_pow_th: "hollight.real_of_int (int_pow x xa) = real_pow (hollight.real_of_int x) xa"
- by (import hollight int_pow_th)
-
-lemma INT_IMAGE: "(EX n. x = int_of_num n) | (EX n. x = int_neg (int_of_num n))"
- by (import hollight INT_IMAGE)
-
-lemma INT_LT_DISCRETE: "int_lt x y = int_le (int_add x (int_of_num 1)) y"
- by (import hollight INT_LT_DISCRETE)
-
-lemma INT_GT_DISCRETE: "int_gt x xa = int_ge x (int_add xa (int_of_num 1))"
- by (import hollight INT_GT_DISCRETE)
-
-lemma INT_FORALL_POS: "(ALL n. P (int_of_num n)) = (ALL i. int_le (int_of_num 0) i --> P i)"
- by (import hollight INT_FORALL_POS)
-
-lemma INT_EXISTS_POS: "(EX n. P (int_of_num n)) = (EX i. int_le (int_of_num 0) i & P i)"
- by (import hollight INT_EXISTS_POS)
-
-lemma INT_FORALL_ABS: "(ALL n. x (int_of_num n)) = (ALL xa. x (int_abs xa))"
- by (import hollight INT_FORALL_ABS)
-
-lemma INT_EXISTS_ABS: "(EX n. P (int_of_num n)) = (EX x. P (int_abs x))"
- by (import hollight INT_EXISTS_ABS)
-
-lemma INT_ABS_MUL_1: "(int_abs (int_mul x y) = int_of_num 1) =
-(int_abs x = int_of_num 1 & int_abs y = int_of_num 1)"
- by (import hollight INT_ABS_MUL_1)
-
-lemma INT_WOP: "(EX x. int_le (int_of_num 0) x & P x) =
-(EX x. int_le (int_of_num 0) x &
- P x & (ALL y. int_le (int_of_num 0) y & P y --> int_le x y))"
- by (import hollight INT_WOP)
-
-lemma INT_POW: "int_pow x 0 = int_of_num 1 &
-(ALL xa. int_pow x (Suc xa) = int_mul x (int_pow x xa))"
- by (import hollight INT_POW)
-
-lemma INT_ABS: "int_abs x = (if int_le (int_of_num 0) x then x else int_neg x)"
- by (import hollight INT_ABS)
-
-lemma INT_GE: "int_ge x xa = int_le xa x"
- by (import hollight INT_GE)
-
-lemma INT_GT: "int_gt x xa = int_lt xa x"
- by (import hollight INT_GT)
-
-lemma INT_LT: "int_lt x xa = (~ int_le xa x)"
- by (import hollight INT_LT)
-
-lemma INT_ARCH: "d ~= int_of_num 0 ==> EX c. int_lt x (int_mul c d)"
- by (import hollight INT_ARCH)
-
-lemma INT_DIVMOD_EXIST_0: "EX x xa.
- if n = int_of_num 0 then x = int_of_num 0 & xa = m
- else int_le (int_of_num 0) xa &
- int_lt xa (int_abs n) & m = int_add (int_mul x n) xa"
- by (import hollight INT_DIVMOD_EXIST_0)
-
-consts
- div :: "hollight.int => hollight.int => hollight.int" ("div")
-
-defs
- div_def: "div ==
-SOME q.
- EX r. ALL m n.
- if n = int_of_num 0 then q m n = int_of_num 0 & r m n = m
- else int_le (int_of_num 0) (r m n) &
- int_lt (r m n) (int_abs n) &
- m = int_add (int_mul (q m n) n) (r m n)"
-
-lemma DEF_div: "div =
-(SOME q.
- EX r. ALL m n.
- if n = int_of_num 0 then q m n = int_of_num 0 & r m n = m
- else int_le (int_of_num 0) (r m n) &
- int_lt (r m n) (int_abs n) &
- m = int_add (int_mul (q m n) n) (r m n))"
- by (import hollight DEF_div)
-
-definition
- rem :: "hollight.int => hollight.int => hollight.int" where
- "rem ==
-SOME r.
- ALL m n.
- if n = int_of_num 0 then div m n = int_of_num 0 & r m n = m
- else int_le (int_of_num 0) (r m n) &
- int_lt (r m n) (int_abs n) &
- m = int_add (int_mul (div m n) n) (r m n)"
-
-lemma DEF_rem: "rem =
-(SOME r.
- ALL m n.
- if n = int_of_num 0 then div m n = int_of_num 0 & r m n = m
- else int_le (int_of_num 0) (r m n) &
- int_lt (r m n) (int_abs n) &
- m = int_add (int_mul (div m n) n) (r m n))"
- by (import hollight DEF_rem)
-
-lemma INT_DIVISION: "n ~= int_of_num 0
-==> m = int_add (int_mul (div m n) n) (rem m n) &
- int_le (int_of_num 0) (rem m n) & int_lt (rem m n) (int_abs n)"
- by (import hollight INT_DIVISION)
-
-lemma sth: "(ALL x y z. int_add x (int_add y z) = int_add (int_add x y) z) &
-(ALL x y. int_add x y = int_add y x) &
-(ALL x. int_add (int_of_num 0) x = x) &
-(ALL x y z. int_mul x (int_mul y z) = int_mul (int_mul x y) z) &
-(ALL x y. int_mul x y = int_mul y x) &
-(ALL x. int_mul (int_of_num 1) x = x) &
-(ALL x. int_mul (int_of_num 0) x = int_of_num 0) &
-(ALL x y z. int_mul x (int_add y z) = int_add (int_mul x y) (int_mul x z)) &
-(ALL x. int_pow x 0 = int_of_num 1) &
-(ALL x xa. int_pow x (Suc xa) = int_mul x (int_pow x xa))"
- by (import hollight sth)
-
-lemma INT_INTEGRAL: "(ALL x. int_mul (int_of_num 0) x = int_of_num 0) &
-(ALL x y z. (int_add x y = int_add x z) = (y = z)) &
-(ALL w x y z.
- (int_add (int_mul w y) (int_mul x z) =
- int_add (int_mul w z) (int_mul x y)) =
- (w = x | y = z))"
- by (import hollight INT_INTEGRAL)
-
-lemma INT_DIVMOD_UNIQ: "m = int_add (int_mul q n) r & int_le (int_of_num 0) r & int_lt r (int_abs n)
-==> div m n = q & rem m n = r"
- by (import hollight INT_DIVMOD_UNIQ)
-
-consts
- eqeq :: "'A => 'A => ('A => 'A => bool) => bool"
-
-defs
- eqeq_def: "hollight.eqeq == %(u::'A) (ua::'A) ub::'A => 'A => bool. ub u ua"
-
-lemma DEF__equal__equal_: "hollight.eqeq = (%(u::'A) (ua::'A) ub::'A => 'A => bool. ub u ua)"
- by (import hollight DEF__equal__equal_)
-
-definition
- real_mod :: "hollight.real => hollight.real => hollight.real => bool" where
- "real_mod == %u ua ub. EX q. integer q & real_sub ua ub = real_mul q u"
-
-lemma DEF_real_mod: "real_mod = (%u ua ub. EX q. integer q & real_sub ua ub = real_mul q u)"
- by (import hollight DEF_real_mod)
-
-definition
- int_divides :: "hollight.int => hollight.int => bool" where
- "int_divides == %u ua. EX x. ua = int_mul u x"
-
-lemma DEF_int_divides: "int_divides = (%u ua. EX x. ua = int_mul u x)"
- by (import hollight DEF_int_divides)
-
-consts
- int_mod :: "hollight.int => hollight.int => hollight.int => bool"
-
-defs
- int_mod_def: "hollight.int_mod == %u ua ub. int_divides u (int_sub ua ub)"
-
-lemma DEF_int_mod: "hollight.int_mod = (%u ua ub. int_divides u (int_sub ua ub))"
- by (import hollight DEF_int_mod)
-
-lemma int_congruent: "hollight.eqeq x xa (hollight.int_mod xb) =
-(EX d. int_sub x xa = int_mul xb d)"
- by (import hollight int_congruent)
-
-consts
- int_coprime :: "hollight.int * hollight.int => bool"
-
-defs
- int_coprime_def: "hollight.int_coprime ==
-%u. EX x y. int_add (int_mul (fst u) x) (int_mul (snd u) y) = int_of_num 1"
-
-lemma DEF_int_coprime: "hollight.int_coprime =
-(%u. EX x y. int_add (int_mul (fst u) x) (int_mul (snd u) y) = int_of_num 1)"
- by (import hollight DEF_int_coprime)
-
-lemma FORALL_UNCURRY: "All (P::('A => 'B => 'C) => bool) =
-(ALL f::'A * 'B => 'C. P (%(a::'A) b::'B. f (a, b)))"
- by (import hollight FORALL_UNCURRY)
-
-lemma EXISTS_UNCURRY: "Ex (x::('A => 'B => 'C) => bool) =
-(EX f::'A * 'B => 'C. x (%(a::'A) b::'B. f (a, b)))"
- by (import hollight EXISTS_UNCURRY)
-
-lemma WF_INT_MEASURE: "(ALL x::'A. int_le (int_of_num (0::nat)) ((m::'A => hollight.int) x)) &
-(ALL x::'A. (ALL y::'A. int_lt (m y) (m x) --> (P::'A => bool) y) --> P x)
-==> P (x::'A)"
- by (import hollight WF_INT_MEASURE)
-
-lemma WF_INT_MEASURE_2: "(ALL (x::'A) y::'B.
- int_le (int_of_num (0::nat)) ((m::'A => 'B => hollight.int) x y)) &
-(ALL (x::'A) y::'B.
- (ALL (x'::'A) y'::'B.
- int_lt (m x' y') (m x y) --> (P::'A => 'B => bool) x' y') -->
- P x y)
-==> P (x::'A) (xa::'B)"
- by (import hollight WF_INT_MEASURE_2)
-
-lemma INT_GCD_EXISTS: "EX d. int_divides d a &
- int_divides d b & (EX x y. d = int_add (int_mul a x) (int_mul b y))"
- by (import hollight INT_GCD_EXISTS)
-
-lemma INT_GCD_EXISTS_POS: "EX d. int_le (int_of_num 0) d &
- int_divides d a &
- int_divides d b & (EX x y. d = int_add (int_mul a x) (int_mul b y))"
- by (import hollight INT_GCD_EXISTS_POS)
-
-consts
- int_gcd :: "hollight.int * hollight.int => hollight.int"
-
-defs
- int_gcd_def: "hollight.int_gcd ==
-SOME d.
- ALL a b.
- int_le (int_of_num 0) (d (a, b)) &
- int_divides (d (a, b)) a &
- int_divides (d (a, b)) b &
- (EX x y. d (a, b) = int_add (int_mul a x) (int_mul b y))"
-
-lemma DEF_int_gcd: "hollight.int_gcd =
-(SOME d.
- ALL a b.
- int_le (int_of_num 0) (d (a, b)) &
- int_divides (d (a, b)) a &
- int_divides (d (a, b)) b &
- (EX x y. d (a, b) = int_add (int_mul a x) (int_mul b y)))"
- by (import hollight DEF_int_gcd)
-
-definition
- num_of_int :: "hollight.int => nat" where
- "num_of_int == %u. SOME n. int_of_num n = u"
-
-lemma DEF_num_of_int: "num_of_int = (%u. SOME n. int_of_num n = u)"
- by (import hollight DEF_num_of_int)
-
-lemma NUM_OF_INT_OF_NUM: "num_of_int (int_of_num x) = x"
- by (import hollight NUM_OF_INT_OF_NUM)
-
-lemma INT_OF_NUM_OF_INT: "int_le (int_of_num 0) x ==> int_of_num (num_of_int x) = x"
- by (import hollight INT_OF_NUM_OF_INT)
-
-lemma NUM_OF_INT: "int_le (int_of_num 0) x = (int_of_num (num_of_int x) = x)"
- by (import hollight NUM_OF_INT)
-
-definition
- num_divides :: "nat => nat => bool" where
- "num_divides == %u ua. int_divides (int_of_num u) (int_of_num ua)"
-
-lemma DEF_num_divides: "num_divides = (%u ua. int_divides (int_of_num u) (int_of_num ua))"
- by (import hollight DEF_num_divides)
-
-definition
- num_mod :: "nat => nat => nat => bool" where
- "num_mod ==
-%u ua ub. hollight.int_mod (int_of_num u) (int_of_num ua) (int_of_num ub)"
-
-lemma DEF_num_mod: "num_mod =
-(%u ua ub. hollight.int_mod (int_of_num u) (int_of_num ua) (int_of_num ub))"
- by (import hollight DEF_num_mod)
-
-lemma num_congruent: "hollight.eqeq x xa (num_mod xb) =
-hollight.eqeq (int_of_num x) (int_of_num xa)
- (hollight.int_mod (int_of_num xb))"
- by (import hollight num_congruent)
-
-definition
- num_coprime :: "nat * nat => bool" where
- "num_coprime ==
-%u. hollight.int_coprime (int_of_num (fst u), int_of_num (snd u))"
-
-lemma DEF_num_coprime: "num_coprime =
-(%u. hollight.int_coprime (int_of_num (fst u), int_of_num (snd u)))"
- by (import hollight DEF_num_coprime)
-
-definition
- num_gcd :: "nat * nat => nat" where
- "num_gcd ==
-%u. num_of_int (hollight.int_gcd (int_of_num (fst u), int_of_num (snd u)))"
-
-lemma DEF_num_gcd: "num_gcd =
-(%u. num_of_int (hollight.int_gcd (int_of_num (fst u), int_of_num (snd u))))"
- by (import hollight DEF_num_gcd)
-
-lemma NUM_GCD: "int_of_num (num_gcd (x, xa)) =
-hollight.int_gcd (int_of_num x, int_of_num xa)"
- by (import hollight NUM_GCD)
-
-lemma IN_ELIM_THM: "(ALL (P::(bool => 'q_43295 => bool) => bool) x::'q_43295.
- (x : {v::'q_43295. P (SETSPEC v)}) =
- P (%(p::bool) t::'q_43295. p & x = t)) &
-(ALL (p::'q_43326 => bool) x::'q_43326.
- (x : {v::'q_43326. EX y::'q_43326. p y & v = y}) = p x) &
-(ALL (P::(bool => 'q_43354 => bool) => bool) x::'q_43354.
- x \<in> {v::'q_43354. P (SETSPEC v)} =
- P (%(p::bool) t::'q_43354. p & x = t)) &
-(ALL (p::'q_43383 => bool) x::'q_43383.
- x \<in> {v::'q_43383. EX y::'q_43383. p y & v = y} = p x) &
-(ALL (p::'q_43400 => bool) x::'q_43400. p x \<longleftrightarrow> p x)"
- by (import hollight IN_ELIM_THM)
-
-lemma INSERT: "insert (x::'A) (s::'A set) = {u::'A. EX y::'A. (y : s | y = x) & u = y}"
- by (import hollight INSERT)
-
-definition
- SING :: "('A set) => bool" where
- "SING == %u::'A set. EX x::'A. u = {x}"
-
-lemma DEF_SING: "SING = (%u::'A set. EX x::'A. u = {x})"
- by (import hollight DEF_SING)
-
-definition
- INJ :: "('A => 'B) => ('A => bool) => ('B => bool) => bool" where
- "INJ ==
-%(u::'A => 'B) (ua::'A => bool) ub::'B => bool.
- (ALL x::'A. x : ua --> u x : ub) &
- (ALL (x::'A) y::'A. x : ua & y : ua & u x = u y --> x = y)"
-
-lemma DEF_INJ: "INJ =
-(%(u::'A => 'B) (ua::'A => bool) ub::'B => bool.
- (ALL x::'A. x : ua --> u x : ub) &
- (ALL (x::'A) y::'A. x : ua & y : ua & u x = u y --> x = y))"
- by (import hollight DEF_INJ)
-
-definition
- SURJ :: "('A => 'B) => ('A => bool) => ('B => bool) => bool" where
- "SURJ ==
-%(u::'A => 'B) (ua::'A => bool) ub::'B => bool.
- (ALL x::'A. x : ua --> u x : ub) &
- (ALL x::'B. x : ub --> (EX y::'A. y : ua & u y = x))"
-
-lemma DEF_SURJ: "SURJ =
-(%(u::'A => 'B) (ua::'A => bool) ub::'B => bool.
- (ALL x::'A. x : ua --> u x : ub) &
- (ALL x::'B. x : ub --> (EX y::'A. y : ua & u y = x)))"
- by (import hollight DEF_SURJ)
-
-definition
- BIJ :: "('A => 'B) => ('A => bool) => ('B => bool) => bool" where
- "BIJ ==
-%(u::'A => 'B) (ua::'A => bool) ub::'B => bool. INJ u ua ub & SURJ u ua ub"
-
-lemma DEF_BIJ: "BIJ =
-(%(u::'A => 'B) (ua::'A => bool) ub::'B => bool. INJ u ua ub & SURJ u ua ub)"
- by (import hollight DEF_BIJ)
-
-definition
- REST :: "('A => bool) => 'A => bool" where
- "REST == %u::'A => bool. u - {Eps u}"
-
-lemma DEF_REST: "REST = (%u::'A => bool. u - {Eps u})"
- by (import hollight DEF_REST)
-
-lemma NOT_IN_EMPTY: "(x::'A) ~: {}"
- by (import hollight NOT_IN_EMPTY)
-
-lemma IN_UNIONS: "((xa::'A) : Union (x::('A => bool) => bool)) =
-(EX t::'A => bool. t : x & xa : t)"
- by (import hollight IN_UNIONS)
-
-lemma IN_INTERS: "((xa::'A) : Inter (x::('A => bool) => bool)) =
-(ALL t::'A => bool. t : x --> xa : t)"
- by (import hollight IN_INTERS)
-
-lemma IN_DELETE: "((xa::'A) : (x::'A => bool) - {xb::'A}) = (xa : x & xa ~= xb)"
- by (import hollight IN_DELETE)
-
-lemma IN_IMAGE: "((x::'B) : (xb::'A => 'B) ` (xa::'A => bool)) =
-(EX xc::'A. x = xb xc & xc : xa)"
- by (import hollight IN_IMAGE)
-
-lemma IN_REST: "((x::'A) : REST (xa::'A => bool)) = (x : xa & x ~= Eps xa)"
- by (import hollight IN_REST)
-
-lemma FORALL_IN_INSERT: "(ALL xc::'q_44214.
- xc : insert (xa::'q_44214) (xb::'q_44214 => bool) -->
- (x::'q_44214 => bool) xc) =
-(x xa & (ALL xa::'q_44214. xa : xb --> x xa))"
- by (import hollight FORALL_IN_INSERT)
-
-lemma EXISTS_IN_INSERT: "(EX xc::'q_44255.
- xc : insert (xa::'q_44255) (xb::'q_44255 => bool) &
- (x::'q_44255 => bool) xc) =
-(x xa | (EX xa::'q_44255. xa : xb & x xa))"
- by (import hollight EXISTS_IN_INSERT)
-
-lemma CHOICE_DEF: "(x::'A => bool) ~= {} ==> Eps x : x"
- by (import hollight CHOICE_DEF)
-
-lemma NOT_EQUAL_SETS: "((x::'A => bool) ~= (xa::'A => bool)) = (EX xb::'A. (xb : xa) = (xb ~: x))"
- by (import hollight NOT_EQUAL_SETS)
-
-lemma EMPTY_NOT_UNIV: "(op ~=::('A::type => bool) => ('A::type => bool) => bool)
- ({}::'A::type => bool) (UNIV::'A::type => bool)"
- by (import hollight EMPTY_NOT_UNIV)
-
-lemma EQ_UNIV: "(ALL x::'A. x : (s::'A => bool)) = (s = UNIV)"
- by (import hollight EQ_UNIV)
-
-lemma SING_SUBSET: "({xa::'q_44493} <= (x::'q_44493 => bool)) = (xa : x)"
- by (import hollight SING_SUBSET)
-
-lemma PSUBSET_UNIV: "((x::'A => bool) < UNIV) = (EX xa::'A. xa ~: x)"
- by (import hollight PSUBSET_UNIV)
-
-lemma PSUBSET_ALT: "((x::'A => bool) < (xa::'A => bool)) =
-(x <= xa & (EX a::'A. a : xa & a ~: x))"
- by (import hollight PSUBSET_ALT)
-
-lemma SUBSET_UNION: "(ALL (x::'A => bool) xa::'A => bool. x <= x Un xa) &
-(ALL (x::'A => bool) xa::'A => bool. x <= xa Un x)"
- by (import hollight SUBSET_UNION)
-
-lemma UNION_EMPTY: "(ALL x::'A => bool. {} Un x = x) & (ALL x::'A => bool. x Un {} = x)"
- by (import hollight UNION_EMPTY)
-
-lemma UNION_UNIV: "(ALL x::'A => bool. UNIV Un x = UNIV) &
-(ALL x::'A => bool. x Un UNIV = UNIV)"
- by (import hollight UNION_UNIV)
-
-lemma INTER_SUBSET: "(ALL (x::'A => bool) xa::'A => bool. x Int xa <= x) &
-(ALL (x::'A => bool) xa::'A => bool. xa Int x <= x)"
- by (import hollight INTER_SUBSET)
-
-lemma INTER_EMPTY: "(ALL x::'A => bool. {} Int x = {}) & (ALL x::'A => bool. x Int {} = {})"
- by (import hollight INTER_EMPTY)
-
-lemma INTER_UNIV: "(ALL x::'A => bool. UNIV Int x = x) & (ALL x::'A => bool. x Int UNIV = x)"
- by (import hollight INTER_UNIV)
-
-lemma IN_DISJOINT: "((x::'A => bool) Int (xa::'A => bool) = {}) =
-(~ (EX xb::'A. xb : x & xb : xa))"
- by (import hollight IN_DISJOINT)
-
-lemma DISJOINT_SYM: "((x::'A => bool) Int (xa::'A => bool) = {}) = (xa Int x = {})"
- by (import hollight DISJOINT_SYM)
-
-lemma DISJOINT_EMPTY: "{} Int (x::'A => bool) = {} & x Int {} = {}"
- by (import hollight DISJOINT_EMPTY)
-
-lemma DISJOINT_EMPTY_REFL: "((x::'A => bool) = {}) = (x Int x = {})"
- by (import hollight DISJOINT_EMPTY_REFL)
-
-lemma DISJOINT_UNION: "(((x::'A => bool) Un (xa::'A => bool)) Int (xb::'A => bool) = {}) =
-(x Int xb = {} & xa Int xb = {})"
- by (import hollight DISJOINT_UNION)
-
-lemma DECOMPOSITION: "((x::'A) : (s::'A => bool)) = (EX t::'A => bool. s = insert x t & x ~: t)"
- by (import hollight DECOMPOSITION)
-
-lemma SET_CASES: "(s::'A => bool) = {} | (EX (x::'A) t::'A => bool. s = insert x t & x ~: t)"
- by (import hollight SET_CASES)
-
-lemma ABSORPTION: "((x::'A) : (xa::'A => bool)) = (insert x xa = xa)"
- by (import hollight ABSORPTION)
-
-lemma INSERT_UNIV: "insert (x::'A) UNIV = UNIV"
- by (import hollight INSERT_UNIV)
-
-lemma INSERT_UNION: "insert (x::'A) (s::'A => bool) Un (t::'A => bool) =
-(if x : t then s Un t else insert x (s Un t))"
- by (import hollight INSERT_UNION)
-
-lemma DISJOINT_INSERT: "(insert (x::'A) (xa::'A => bool) Int (xb::'A => bool) = {}) =
-(xa Int xb = {} & x ~: xb)"
- by (import hollight DISJOINT_INSERT)
-
-lemma INSERT_AC: "insert (x::'q_45764) (insert (y::'q_45764) (s::'q_45764 => bool)) =
-insert y (insert x s) &
-insert x (insert x s) = insert x s"
- by (import hollight INSERT_AC)
-
-lemma INTER_ACI: "(p::'q_45831 => bool) Int (q::'q_45831 => bool) = q Int p &
-p Int q Int (r::'q_45831 => bool) = p Int (q Int r) &
-p Int (q Int r) = q Int (p Int r) & p Int p = p & p Int (p Int q) = p Int q"
- by (import hollight INTER_ACI)
-
-lemma UNION_ACI: "(p::'q_45897 => bool) Un (q::'q_45897 => bool) = q Un p &
-p Un q Un (r::'q_45897 => bool) = p Un (q Un r) &
-p Un (q Un r) = q Un (p Un r) & p Un p = p & p Un (p Un q) = p Un q"
- by (import hollight UNION_ACI)
-
-lemma DELETE_NON_ELEMENT: "((x::'A) ~: (xa::'A => bool)) = (xa - {x} = xa)"
- by (import hollight DELETE_NON_ELEMENT)
-
-lemma IN_DELETE_EQ: "(((x::'A) : (s::'A => bool)) = ((x'::'A) : s)) =
-((x : s - {x'}) = (x' : s - {x}))"
- by (import hollight IN_DELETE_EQ)
-
-lemma EMPTY_DELETE: "{} - {x::'A} = {}"
- by (import hollight EMPTY_DELETE)
-
-lemma DELETE_DELETE: "(xa::'A => bool) - {x::'A} - {x} = xa - {x}"
- by (import hollight DELETE_DELETE)
-
-lemma DELETE_COMM: "(xb::'A => bool) - {x::'A} - {xa::'A} = xb - {xa} - {x}"
- by (import hollight DELETE_COMM)
-
-lemma DELETE_SUBSET: "(xa::'A => bool) - {x::'A} <= xa"
- by (import hollight DELETE_SUBSET)
-
-lemma SUBSET_DELETE: "((xa::'A => bool) <= (xb::'A => bool) - {x::'A}) = (x ~: xa & xa <= xb)"
- by (import hollight SUBSET_DELETE)
-
-lemma SUBSET_INSERT_DELETE: "((xa::'A => bool) <= insert (x::'A) (xb::'A => bool)) = (xa - {x} <= xb)"
- by (import hollight SUBSET_INSERT_DELETE)
-
-lemma PSUBSET_INSERT_SUBSET: "((x::'A => bool) < (xa::'A => bool)) =
-(EX xb::'A. xb ~: x & insert xb x <= xa)"
- by (import hollight PSUBSET_INSERT_SUBSET)
-
-lemma PSUBSET_MEMBER: "((x::'A => bool) < (xa::'A => bool)) =
-(x <= xa & (EX y::'A. y : xa & y ~: x))"
- by (import hollight PSUBSET_MEMBER)
-
-lemma DELETE_INSERT: "insert (x::'A) (s::'A => bool) - {y::'A} =
-(if x = y then s - {y} else insert x (s - {y}))"
- by (import hollight DELETE_INSERT)
-
-lemma DELETE_INTER: "((x::'A => bool) - {xb::'A}) Int (xa::'A => bool) = x Int xa - {xb}"
- by (import hollight DELETE_INTER)
-
-lemma DISJOINT_DELETE_SYM: "(((x::'A => bool) - {xb::'A}) Int (xa::'A => bool) = {}) =
-((xa - {xb}) Int x = {})"
- by (import hollight DISJOINT_DELETE_SYM)
-
-lemma FORALL_IN_UNIONS: "(ALL x::'q_46386.
- x : Union (s::('q_46386 => bool) => bool) --> (P::'q_46386 => bool) x) =
-(ALL (t::'q_46386 => bool) x::'q_46386. t : s & x : t --> P x)"
- by (import hollight FORALL_IN_UNIONS)
-
-lemma EXISTS_IN_UNIONS: "(EX x::'q_46428.
- x : Union (s::('q_46428 => bool) => bool) & (P::'q_46428 => bool) x) =
-(EX (t::'q_46428 => bool) x::'q_46428. t : s & x : t & P x)"
- by (import hollight EXISTS_IN_UNIONS)
-
-lemma EMPTY_UNIONS: "(Union (x::('q_46454 => bool) => bool) = {}) =
-(ALL xa::'q_46454 => bool. xa : x --> xa = {})"
- by (import hollight EMPTY_UNIONS)
-
-lemma INTER_UNIONS: "(ALL (x::('q_46493 => bool) => bool) xa::'q_46493 => bool.
- Union x Int xa =
- Union
- {u::'q_46493 => bool.
- EX xb::'q_46493 => bool. xb : x & u = xb Int xa}) &
-(ALL (x::('q_46529 => bool) => bool) xa::'q_46529 => bool.
- xa Int Union x =
- Union
- {u::'q_46529 => bool. EX xb::'q_46529 => bool. xb : x & u = xa Int xb})"
- by (import hollight INTER_UNIONS)
-
-lemma UNIONS_SUBSET: "(Union (x::('q_46545 => bool) => bool) <= (xa::'q_46545 => bool)) =
-(ALL xb::'q_46545 => bool. xb : x --> xb <= xa)"
- by (import hollight UNIONS_SUBSET)
-
-lemma IMAGE_CLAUSES: "(f::'q_46676 => 'q_46680) ` {} = {} &
-f ` insert (x::'q_46676) (s::'q_46676 => bool) = insert (f x) (f ` s)"
- by (import hollight IMAGE_CLAUSES)
-
-lemma IMAGE_INTER_INJ: "(!!(xa::'q_46846) y::'q_46846.
- (x::'q_46846 => 'q_46857) xa = x y ==> xa = y)
-==> x ` ((xa::'q_46846 => bool) Int (xb::'q_46846 => bool)) =
- x ` xa Int x ` xb"
- by (import hollight IMAGE_INTER_INJ)
-
-lemma IMAGE_DIFF_INJ: "(!!(xa::'q_46900) y::'q_46900.
- (x::'q_46900 => 'q_46911) xa = x y ==> xa = y)
-==> x ` ((xa::'q_46900 => bool) - (xb::'q_46900 => bool)) = x ` xa - x ` xb"
- by (import hollight IMAGE_DIFF_INJ)
-
-lemma IMAGE_DELETE_INJ: "(!!xa::'q_46958.
- (x::'q_46958 => 'q_46957) xa = x (xb::'q_46958) ==> xa = xb)
-==> x ` ((xa::'q_46958 => bool) - {xb}) = x ` xa - {x xb}"
- by (import hollight IMAGE_DELETE_INJ)
-
-lemma FORALL_IN_IMAGE: "(ALL xb::'q_47016.
- xb : (x::'q_47017 => 'q_47016) ` (xa::'q_47017 => bool) -->
- (P::'q_47016 => bool) xb) =
-(ALL xb::'q_47017. xb : xa --> P (x xb))"
- by (import hollight FORALL_IN_IMAGE)
-
-lemma EXISTS_IN_IMAGE: "(EX xb::'q_47052.
- xb : (x::'q_47053 => 'q_47052) ` (xa::'q_47053 => bool) &
- (P::'q_47052 => bool) xb) =
-(EX xb::'q_47053. xb : xa & P (x xb))"
- by (import hollight EXISTS_IN_IMAGE)
-
-lemma FORALL_SUBSET_IMAGE: "(ALL xc<=(xa::'q_47140 => 'q_47156) ` (xb::'q_47140 => bool).
- (x::('q_47156 => bool) => bool) xc) =
-(ALL t<=xb. x (xa ` t))"
- by (import hollight FORALL_SUBSET_IMAGE)
-
-lemma EXISTS_SUBSET_IMAGE: "(EX xc<=(xa::'q_47183 => 'q_47199) ` (xb::'q_47183 => bool).
- (x::('q_47199 => bool) => bool) xc) =
-(EX t<=xb. x (xa ` t))"
- by (import hollight EXISTS_SUBSET_IMAGE)
-
-lemma SIMPLE_IMAGE: "{u::'q_47262.
- EX xb::'q_47258.
- xb : (xa::'q_47258 => bool) & u = (x::'q_47258 => 'q_47262) xb} =
-x ` xa"
- by (import hollight SIMPLE_IMAGE)
-
-lemma SIMPLE_IMAGE_GEN: "{u::'q_47292.
- EX xa::'q_47305.
- (P::'q_47305 => bool) xa & u = (x::'q_47305 => 'q_47292) xa} =
-x ` {u::'q_47305. EX x::'q_47305. P x & u = x}"
- by (import hollight SIMPLE_IMAGE_GEN)
-
-lemma IMAGE_UNIONS: "(x::'q_47323 => 'q_47332) ` Union (xa::('q_47323 => bool) => bool) =
-Union (op ` x ` xa)"
- by (import hollight IMAGE_UNIONS)
-
-lemma SURJECTIVE_IMAGE_EQ: "(ALL y::'q_47396.
- y : (xa::'q_47396 => bool) -->
- (EX x::'q_47400. (f::'q_47400 => 'q_47396) x = y)) &
-(ALL xb::'q_47400. (f xb : xa) = (xb : (x::'q_47400 => bool)))
-==> f ` x = xa"
- by (import hollight SURJECTIVE_IMAGE_EQ)
-
-lemma EMPTY_GSPEC: "{u::'q_47425. Ex (SETSPEC u False)} = {}"
- by (import hollight EMPTY_GSPEC)
-
-lemma SING_GSPEC: "(ALL x::'q_47454. {u::'q_47454. EX xa::'q_47454. xa = x & u = xa} = {x}) &
-(ALL x::'q_47480. {u::'q_47480. EX xa::'q_47480. x = xa & u = xa} = {x})"
- by (import hollight SING_GSPEC)
-
-lemma IN_ELIM_PAIR_THM: "((xa::'q_47526, xb::'q_47525)
- : {xa::'q_47526 * 'q_47525.
- EX (xb::'q_47526) y::'q_47525.
- (x::'q_47526 => 'q_47525 => bool) xb y & xa = (xb, y)}) =
-x xa xb"
- by (import hollight IN_ELIM_PAIR_THM)
-
-lemma SET_PAIR_THM: "{u::'q_47570 * 'q_47569.
- EX p::'q_47570 * 'q_47569. (x::'q_47570 * 'q_47569 => bool) p & u = p} =
-{u::'q_47570 * 'q_47569.
- EX (a::'q_47570) b::'q_47569. x (a, b) & u = (a, b)}"
- by (import hollight SET_PAIR_THM)
-
-lemma FORALL_IN_GSPEC: "(ALL (P::'q_47618 => bool) f::'q_47618 => 'q_47739.
- (ALL z::'q_47739.
- z : {u::'q_47739. EX x::'q_47618. P x & u = f x} -->
- (Q::'q_47739 => bool) z) =
- (ALL x::'q_47618. P x --> Q (f x))) &
-(ALL (P::'q_47675 => 'q_47674 => bool) f::'q_47675 => 'q_47674 => 'q_47739.
- (ALL z::'q_47739.
- z : {u::'q_47739.
- EX (x::'q_47675) y::'q_47674. P x y & u = f x y} -->
- Q z) =
- (ALL (x::'q_47675) y::'q_47674. P x y --> Q (f x y))) &
-(ALL (P::'q_47742 => 'q_47741 => 'q_47740 => bool)
- f::'q_47742 => 'q_47741 => 'q_47740 => 'q_47739.
- (ALL z::'q_47739.
- z : {u::'q_47739.
- EX (w::'q_47742) (x::'q_47741) y::'q_47740.
- P w x y & u = f w x y} -->
- Q z) =
- (ALL (w::'q_47742) (x::'q_47741) y::'q_47740. P w x y --> Q (f w x y)))"
- by (import hollight FORALL_IN_GSPEC)
-
-lemma EXISTS_IN_GSPEC: "(ALL (P::'q_47788 => bool) f::'q_47788 => 'q_47909.
- (EX z::'q_47909.
- z : {u::'q_47909. EX x::'q_47788. P x & u = f x} &
- (Q::'q_47909 => bool) z) =
- (EX x::'q_47788. P x & Q (f x))) &
-(ALL (P::'q_47845 => 'q_47844 => bool) f::'q_47845 => 'q_47844 => 'q_47909.
- (EX z::'q_47909.
- z : {u::'q_47909. EX (x::'q_47845) y::'q_47844. P x y & u = f x y} &
- Q z) =
- (EX (x::'q_47845) y::'q_47844. P x y & Q (f x y))) &
-(ALL (P::'q_47912 => 'q_47911 => 'q_47910 => bool)
- f::'q_47912 => 'q_47911 => 'q_47910 => 'q_47909.
- (EX z::'q_47909.
- z : {u::'q_47909.
- EX (w::'q_47912) (x::'q_47911) y::'q_47910.
- P w x y & u = f w x y} &
- Q z) =
- (EX (w::'q_47912) (x::'q_47911) y::'q_47910. P w x y & Q (f w x y)))"
- by (import hollight EXISTS_IN_GSPEC)
-
-lemma SET_PROVE_CASES: "(P::('A => bool) => bool) {} &
-(ALL (a::'A) s::'A => bool. a ~: s --> P (insert a s))
-==> P (x::'A => bool)"
- by (import hollight SET_PROVE_CASES)
-
-lemma UNIONS_IMAGE: "Union ((f::'q_47989 => 'q_47973 => bool) ` (s::'q_47989 => bool)) =
-{u::'q_47973. EX y::'q_47973. (EX x::'q_47989. x : s & y : f x) & u = y}"
- by (import hollight UNIONS_IMAGE)
-
-lemma INTERS_IMAGE: "Inter ((f::'q_48032 => 'q_48016 => bool) ` (s::'q_48032 => bool)) =
-{u::'q_48016. EX y::'q_48016. (ALL x::'q_48032. x : s --> y : f x) & u = y}"
- by (import hollight INTERS_IMAGE)
-
-lemma UNIONS_GSPEC: "(ALL (P::'q_48085 => bool) f::'q_48085 => 'q_48071 => bool.
- Union {u::'q_48071 => bool. EX x::'q_48085. P x & u = f x} =
- {u::'q_48071.
- EX a::'q_48071. (EX x::'q_48085. P x & a : f x) & u = a}) &
-(ALL (P::'q_48149 => 'q_48148 => bool)
- f::'q_48149 => 'q_48148 => 'q_48129 => bool.
- Union
- {u::'q_48129 => bool.
- EX (x::'q_48149) y::'q_48148. P x y & u = f x y} =
- {u::'q_48129.
- EX a::'q_48129.
- (EX (x::'q_48149) y::'q_48148. P x y & a : f x y) & u = a}) &
-(ALL (P::'q_48223 => 'q_48222 => 'q_48221 => bool)
- f::'q_48223 => 'q_48222 => 'q_48221 => 'q_48197 => bool.
- Union
- {u::'q_48197 => bool.
- EX (x::'q_48223) (y::'q_48222) z::'q_48221. P x y z & u = f x y z} =
- {u::'q_48197.
- EX a::'q_48197.
- (EX (x::'q_48223) (y::'q_48222) z::'q_48221.
- P x y z & a : f x y z) &
- u = a})"
- by (import hollight UNIONS_GSPEC)
-
-lemma INTERS_GSPEC: "(ALL (P::'q_48276 => bool) f::'q_48276 => 'q_48262 => bool.
- Inter {u::'q_48262 => bool. EX x::'q_48276. P x & u = f x} =
- {u::'q_48262.
- EX a::'q_48262. (ALL x::'q_48276. P x --> a : f x) & u = a}) &
-(ALL (P::'q_48340 => 'q_48339 => bool)
- f::'q_48340 => 'q_48339 => 'q_48320 => bool.
- Inter
- {u::'q_48320 => bool.
- EX (x::'q_48340) y::'q_48339. P x y & u = f x y} =
- {u::'q_48320.
- EX a::'q_48320.
- (ALL (x::'q_48340) y::'q_48339. P x y --> a : f x y) & u = a}) &
-(ALL (P::'q_48414 => 'q_48413 => 'q_48412 => bool)
- f::'q_48414 => 'q_48413 => 'q_48412 => 'q_48388 => bool.
- Inter
- {u::'q_48388 => bool.
- EX (x::'q_48414) (y::'q_48413) z::'q_48412. P x y z & u = f x y z} =
- {u::'q_48388.
- EX a::'q_48388.
- (ALL (x::'q_48414) (y::'q_48413) z::'q_48412.
- P x y z --> a : f x y z) &
- u = a})"
- by (import hollight INTERS_GSPEC)
-
-lemma DIFF_INTERS: "(x::'q_48451 => bool) - Inter (xa::('q_48451 => bool) => bool) =
-Union {u::'q_48451 => bool. EX xb::'q_48451 => bool. xb : xa & u = x - xb}"
- by (import hollight DIFF_INTERS)
-
-lemma INTERS_UNIONS: "Inter (x::('q_48486 => bool) => bool) =
-UNIV -
-Union {u::'q_48486 => bool. EX t::'q_48486 => bool. t : x & u = UNIV - t}"
- by (import hollight INTERS_UNIONS)
-
-lemma UNIONS_INTERS: "Union (s::('q_48521 => bool) => bool) =
-UNIV -
-Inter {u::'q_48521 => bool. EX t::'q_48521 => bool. t : s & u = UNIV - t}"
- by (import hollight UNIONS_INTERS)
-
-lemma FINITE_SING: "finite {x::'q_48799}"
- by (import hollight FINITE_SING)
-
-lemma FINITE_DELETE_IMP: "finite (s::'A => bool) ==> finite (s - {x::'A})"
- by (import hollight FINITE_DELETE_IMP)
-
-lemma FINITE_DELETE: "finite ((s::'A => bool) - {x::'A}) = finite s"
- by (import hollight FINITE_DELETE)
-
-lemma FINITE_FINITE_UNIONS: "finite (s::('q_48871 => bool) => bool)
-==> finite (Union s) = (ALL t::'q_48871 => bool. t : s --> finite t)"
- by (import hollight FINITE_FINITE_UNIONS)
-
-lemma FINITE_IMAGE_EXPAND: "finite (s::'A => bool)
-==> finite
- {u::'B. EX y::'B. (EX x::'A. x : s & y = (f::'A => 'B) x) & u = y}"
- by (import hollight FINITE_IMAGE_EXPAND)
-
-lemma FINITE_IMAGE_INJ_GENERAL: "(ALL (x::'A) y::'A.
- x : (s::'A => bool) & y : s & (f::'A => 'B) x = f y --> x = y) &
-finite (x::'B => bool)
-==> finite {u::'A. EX xa::'A. (xa : s & f xa : x) & u = xa}"
- by (import hollight FINITE_IMAGE_INJ_GENERAL)
-
-lemma FINITE_FINITE_PREIMAGE_GENERAL: "finite (t::'B => bool) &
-(ALL y::'B.
- y : t -->
- finite
- {u::'A. EX x::'A. (x : (s::'A => bool) & (f::'A => 'B) x = y) & u = x})
-==> finite {u::'A. EX x::'A. (x : s & f x : t) & u = x}"
- by (import hollight FINITE_FINITE_PREIMAGE_GENERAL)
-
-lemma FINITE_FINITE_PREIMAGE: "finite (t::'B => bool) &
-(ALL y::'B. y : t --> finite {u::'A. EX x::'A. (f::'A => 'B) x = y & u = x})
-==> finite {u::'A. EX x::'A. f x : t & u = x}"
- by (import hollight FINITE_FINITE_PREIMAGE)
-
-lemma FINITE_IMAGE_INJ_EQ: "(!!(x::'A) y::'A.
- x : (s::'A => bool) & y : s & (f::'A => 'B) x = f y ==> x = y)
-==> finite (f ` s) = finite s"
- by (import hollight FINITE_IMAGE_INJ_EQ)
-
-lemma FINITE_IMAGE_INJ: "(ALL (x::'A) y::'A. (f::'A => 'B) x = f y --> x = y) &
-finite (A::'B => bool)
-==> finite {u::'A. EX x::'A. f x : A & u = x}"
- by (import hollight FINITE_IMAGE_INJ)
-
-lemma INFINITE_IMAGE_INJ: "[| !!(x::'A) y::'A. (f::'A => 'B) x = f y ==> x = y;
- infinite (s::'A => bool) |]
-==> infinite (f ` s)"
- by (import hollight INFINITE_IMAGE_INJ)
-
-lemma FINITE_SUBSET_IMAGE: "(finite (t::'B => bool) & t <= (f::'A => 'B) ` (s::'A => bool)) =
-(EX x::'A => bool. finite x & x <= s & t = f ` x)"
- by (import hollight FINITE_SUBSET_IMAGE)
-
-lemma EXISTS_FINITE_SUBSET_IMAGE: "(EX xc::'q_49755 => bool.
- finite xc &
- xc <= (xa::'q_49735 => 'q_49755) ` (xb::'q_49735 => bool) &
- (x::('q_49755 => bool) => bool) xc) =
-(EX xc::'q_49735 => bool. finite xc & xc <= xb & x (xa ` xc))"
- by (import hollight EXISTS_FINITE_SUBSET_IMAGE)
-
-lemma FINITE_SUBSET_IMAGE_IMP: "finite (t::'B => bool) & t <= (f::'A => 'B) ` (s::'A => bool)
-==> EX s'::'A => bool. finite s' & s' <= s & t <= f ` s'"
- by (import hollight FINITE_SUBSET_IMAGE_IMP)
-
-definition
- FINREC :: "('A => 'B => 'B) => 'B => ('A => bool) => 'B => nat => bool" where
- "FINREC ==
-SOME FINREC::('A => 'B => 'B) => 'B => ('A => bool) => 'B => nat => bool.
- (ALL (f::'A => 'B => 'B) (s::'A => bool) (a::'B) b::'B.
- FINREC f b s a (0::nat) = (s = {} & a = b)) &
- (ALL (b::'B) (s::'A => bool) (n::nat) (a::'B) f::'A => 'B => 'B.
- FINREC f b s a (Suc n) =
- (EX (x::'A) c::'B. x : s & FINREC f b (s - {x}) c n & a = f x c))"
-
-lemma DEF_FINREC: "FINREC =
-(SOME FINREC::('A => 'B => 'B) => 'B => ('A => bool) => 'B => nat => bool.
- (ALL (f::'A => 'B => 'B) (s::'A => bool) (a::'B) b::'B.
- FINREC f b s a (0::nat) = (s = {} & a = b)) &
- (ALL (b::'B) (s::'A => bool) (n::nat) (a::'B) f::'A => 'B => 'B.
- FINREC f b s a (Suc n) =
- (EX (x::'A) c::'B. x : s & FINREC f b (s - {x}) c n & a = f x c)))"
- by (import hollight DEF_FINREC)
-
-lemma FINREC_1_LEMMA: "FINREC (x::'q_49919 => 'q_49918 => 'q_49918) (xa::'q_49918)
- (xb::'q_49919 => bool) (xc::'q_49918) (Suc (0::nat)) =
-(EX xd::'q_49919. xb = {xd} & xc = x xd xa)"
- by (import hollight FINREC_1_LEMMA)
-
-lemma FINREC_SUC_LEMMA: "[| !!(x::'A) (y::'A) s::'B.
- x ~= y ==> (f::'A => 'B => 'B) x (f y s) = f y (f x s);
- FINREC f (b::'B) (s::'A => bool) (z::'B) (Suc (n::nat)); (x::'A) : s |]
-==> EX w::'B. FINREC f b (s - {x}) w n & z = f x w"
- by (import hollight FINREC_SUC_LEMMA)
-
-lemma FINREC_UNIQUE_LEMMA: "[| !!(x::'A) (y::'A) s::'B.
- x ~= y ==> (f::'A => 'B => 'B) x (f y s) = f y (f x s);
- FINREC f (b::'B) (s::'A => bool) (a1::'B) (n1::nat) &
- FINREC f b s (a2::'B) (n2::nat) |]
-==> a1 = a2 & n1 = n2"
- by (import hollight FINREC_UNIQUE_LEMMA)
-
-lemma FINREC_EXISTS_LEMMA: "finite (s::'A => bool)
-==> EX a::'B. Ex (FINREC (f::'A => 'B => 'B) (b::'B) s a)"
- by (import hollight FINREC_EXISTS_LEMMA)
-
-lemma FINREC_FUN_LEMMA: "(ALL s::'A.
- (P::'A => bool) s -->
- (EX a::'B. Ex ((R::'A => 'B => 'C => bool) s a))) &
-(ALL (n1::'C) (n2::'C) (s::'A) (a1::'B) a2::'B.
- R s a1 n1 & R s a2 n2 --> a1 = a2 & n1 = n2)
-==> EX x::'A => 'B. ALL (s::'A) a::'B. P s --> Ex (R s a) = (x s = a)"
- by (import hollight FINREC_FUN_LEMMA)
-
-lemma FINREC_FUN: "(!!(x::'A) (y::'A) s::'B.
- x ~= y ==> (f::'A => 'B => 'B) x (f y s) = f y (f x s))
-==> EX g::('A => bool) => 'B.
- g {} = (b::'B) &
- (ALL (s::'A => bool) x::'A.
- finite s & x : s --> g s = f x (g (s - {x})))"
- by (import hollight FINREC_FUN)
-
-lemma SET_RECURSION_LEMMA: "(!!(x::'A) (y::'A) s::'B.
- x ~= y ==> (f::'A => 'B => 'B) x (f y s) = f y (f x s))
-==> EX g::('A => bool) => 'B.
- g {} = (b::'B) &
- (ALL (x::'A) s::'A => bool.
- finite s --> g (insert x s) = (if x : s then g s else f x (g s)))"
- by (import hollight SET_RECURSION_LEMMA)
-
-definition
- ITSET :: "('q_50575 => 'q_50574 => 'q_50574)
-=> ('q_50575 => bool) => 'q_50574 => 'q_50574" where
- "ITSET ==
-%(u::'q_50575 => 'q_50574 => 'q_50574) (ua::'q_50575 => bool) ub::'q_50574.
- (SOME g::('q_50575 => bool) => 'q_50574.
- g {} = ub &
- (ALL (x::'q_50575) s::'q_50575 => bool.
- finite s -->
- g (insert x s) = (if x : s then g s else u x (g s))))
- ua"
-
-lemma DEF_ITSET: "ITSET =
-(%(u::'q_50575 => 'q_50574 => 'q_50574) (ua::'q_50575 => bool) ub::'q_50574.
- (SOME g::('q_50575 => bool) => 'q_50574.
- g {} = ub &
- (ALL (x::'q_50575) s::'q_50575 => bool.
- finite s -->
- g (insert x s) = (if x : s then g s else u x (g s))))
- ua)"
- by (import hollight DEF_ITSET)
-
-lemma FINITE_RECURSION: "(!!(x::'A) (y::'A) s::'B.
- x ~= y ==> (f::'A => 'B => 'B) x (f y s) = f y (f x s))
-==> ITSET f {} (b::'B) = b &
- (ALL (x::'A) xa::'A => bool.
- finite xa -->
- ITSET f (insert x xa) b =
- (if x : xa then ITSET f xa b else f x (ITSET f xa b)))"
- by (import hollight FINITE_RECURSION)
-
-lemma FINITE_RECURSION_DELETE: "(!!(x::'A) (y::'A) s::'B.
- x ~= y ==> (f::'A => 'B => 'B) x (f y s) = f y (f x s))
-==> ITSET f {} (b::'B) = b &
- (ALL (x::'A) s::'A => bool.
- finite s -->
- ITSET f s b =
- (if x : s then f x (ITSET f (s - {x}) b) else ITSET f (s - {x}) b))"
- by (import hollight FINITE_RECURSION_DELETE)
-
-lemma ITSET_EQ: "finite (x::'q_50880 => bool) &
-(ALL xc::'q_50880.
- xc : x -->
- (xa::'q_50880 => 'q_50881 => 'q_50881) xc =
- (xb::'q_50880 => 'q_50881 => 'q_50881) xc) &
-(ALL (x::'q_50880) (y::'q_50880) s::'q_50881.
- x ~= y --> xa x (xa y s) = xa y (xa x s)) &
-(ALL (x::'q_50880) (y::'q_50880) s::'q_50881.
- x ~= y --> xb x (xb y s) = xb y (xb x s))
-==> ITSET xa x (xc::'q_50881) = ITSET xb x xc"
- by (import hollight ITSET_EQ)
-
-lemma SUBSET_RESTRICT: "{u::'q_50914.
- EX xb::'q_50914.
- (xb : (x::'q_50914 => bool) & (xa::'q_50914 => bool) xb) & u = xb}
-<= x"
- by (import hollight SUBSET_RESTRICT)
-
-lemma FINITE_RESTRICT: "finite (s::'A => bool)
-==> finite {u::'A. EX x::'A. (x : s & (P::'A => bool) x) & u = x}"
- by (import hollight FINITE_RESTRICT)
-
-definition
- CARD :: "('q_50968 => bool) => nat" where
- "CARD == %u::'q_50968 => bool. ITSET (%x::'q_50968. Suc) u (0::nat)"
-
-lemma DEF_CARD: "CARD = (%u::'q_50968 => bool. ITSET (%x::'q_50968. Suc) u (0::nat))"
- by (import hollight DEF_CARD)
-
-lemma CARD_CLAUSES: "CARD {} = (0::nat) &
-(ALL (x::'A::type) s::'A::type => bool.
- finite s -->
- CARD (insert x s) = (if x : s then CARD s else Suc (CARD s)))"
- by (import hollight CARD_CLAUSES)
-
-lemma CARD_UNION: "finite (x::'A => bool) & finite (xa::'A => bool) & x Int xa = {}
-==> CARD (x Un xa) = CARD x + CARD xa"
- by (import hollight CARD_UNION)
-
-lemma CARD_DELETE: "finite (s::'A => bool)
-==> CARD (s - {x::'A}) = (if x : s then CARD s - (1::nat) else CARD s)"
- by (import hollight CARD_DELETE)
-
-lemma CARD_UNION_EQ: "finite (u::'q_51213 => bool) &
-(s::'q_51213 => bool) Int (t::'q_51213 => bool) = {} & s Un t = u
-==> CARD s + CARD t = CARD u"
- by (import hollight CARD_UNION_EQ)
-
-lemma CARD_DIFF: "finite (s::'q_51270 => bool) & (t::'q_51270 => bool) <= s
-==> CARD (s - t) = CARD s - CARD t"
- by (import hollight CARD_DIFF)
-
-lemma CARD_EQ_0: "finite (s::'q_51308 => bool) ==> (CARD s = (0::nat)) = (s = {})"
- by (import hollight CARD_EQ_0)
-
-lemma FINITE_INDUCT_DELETE: "[| (P::('A => bool) => bool) {} &
- (ALL s::'A => bool.
- finite s & s ~= {} --> (EX x::'A. x : s & (P (s - {x}) --> P s)));
- finite (s::'A => bool) |]
-==> P s"
- by (import hollight FINITE_INDUCT_DELETE)
-
-definition
- HAS_SIZE :: "('q_51427 => bool) => nat => bool" where
- "HAS_SIZE == %(u::'q_51427 => bool) ua::nat. finite u & CARD u = ua"
-
-lemma DEF_HAS_SIZE: "HAS_SIZE = (%(u::'q_51427 => bool) ua::nat. finite u & CARD u = ua)"
- by (import hollight DEF_HAS_SIZE)
-
-lemma HAS_SIZE_CARD: "HAS_SIZE (x::'q_51446 => bool) (xa::nat) ==> CARD x = xa"
- by (import hollight HAS_SIZE_CARD)
-
-lemma HAS_SIZE_0: "HAS_SIZE (s::'A => bool) (0::nat) = (s = {})"
- by (import hollight HAS_SIZE_0)
-
-lemma HAS_SIZE_SUC: "HAS_SIZE (s::'A => bool) (Suc (n::nat)) =
-(s ~= {} & (ALL x::'A. x : s --> HAS_SIZE (s - {x}) n))"
- by (import hollight HAS_SIZE_SUC)
-
-lemma HAS_SIZE_UNION: "HAS_SIZE (x::'q_51584 => bool) (xb::nat) &
-HAS_SIZE (xa::'q_51584 => bool) (xc::nat) & x Int xa = {}
-==> HAS_SIZE (x Un xa) (xb + xc)"
- by (import hollight HAS_SIZE_UNION)
-
-lemma HAS_SIZE_DIFF: "HAS_SIZE (x::'q_51620 => bool) (xb::nat) &
-HAS_SIZE (xa::'q_51620 => bool) (xc::nat) & xa <= x
-==> HAS_SIZE (x - xa) (xb - xc)"
- by (import hollight HAS_SIZE_DIFF)
-
-lemma HAS_SIZE_UNIONS: "HAS_SIZE (x::'A => bool) (xb::nat) &
-(ALL xb::'A. xb : x --> HAS_SIZE ((xa::'A => 'B => bool) xb) (xc::nat)) &
-(ALL (xb::'A) y::'A. xb : x & y : x & xb ~= y --> xa xb Int xa y = {})
-==> HAS_SIZE (Union {u::'B => bool. EX xb::'A. xb : x & u = xa xb})
- (xb * xc)"
- by (import hollight HAS_SIZE_UNIONS)
-
-lemma FINITE_HAS_SIZE: "finite (x::'q_51824 => bool) = HAS_SIZE x (CARD x)"
- by (import hollight FINITE_HAS_SIZE)
-
-lemma HAS_SIZE_CLAUSES: "HAS_SIZE (s::'q_51886 => bool) (0::nat) = (s = {}) &
-HAS_SIZE s (Suc (n::nat)) =
-(EX (a::'q_51886) t::'q_51886 => bool.
- HAS_SIZE t n & a ~: t & s = insert a t)"
- by (import hollight HAS_SIZE_CLAUSES)
-
-lemma CARD_SUBSET_EQ: "finite (b::'A => bool) & (a::'A => bool) <= b & CARD a = CARD b ==> a = b"
- by (import hollight CARD_SUBSET_EQ)
-
-lemma CARD_SUBSET: "(a::'A => bool) <= (b::'A => bool) & finite b ==> CARD a <= CARD b"
- by (import hollight CARD_SUBSET)
-
-lemma CARD_SUBSET_LE: "finite (b::'A => bool) & (a::'A => bool) <= b & CARD b <= CARD a ==> a = b"
- by (import hollight CARD_SUBSET_LE)
-
-lemma SUBSET_CARD_EQ: "finite (t::'q_52197 => bool) & (s::'q_52197 => bool) <= t
-==> (CARD s = CARD t) = (s = t)"
- by (import hollight SUBSET_CARD_EQ)
-
-lemma CARD_PSUBSET: "(a::'A => bool) < (b::'A => bool) & finite b ==> CARD a < CARD b"
- by (import hollight CARD_PSUBSET)
-
-lemma CARD_UNION_LE: "finite (s::'A => bool) & finite (t::'A => bool)
-==> CARD (s Un t) <= CARD s + CARD t"
- by (import hollight CARD_UNION_LE)
-
-lemma CARD_UNIONS_LE: "HAS_SIZE (x::'A => bool) (xb::nat) &
-(ALL xb::'A.
- xb : x -->
- finite ((xa::'A => 'B => bool) xb) & CARD (xa xb) <= (xc::nat))
-==> CARD (Union {u::'B => bool. EX xb::'A. xb : x & u = xa xb}) <= xb * xc"
- by (import hollight CARD_UNIONS_LE)
-
-lemma CARD_UNION_GEN: "finite (s::'q_52620 => bool) & finite (t::'q_52620 => bool)
-==> CARD (s Un t) = CARD s + CARD t - CARD (s Int t)"
- by (import hollight CARD_UNION_GEN)
-
-lemma CARD_UNION_OVERLAP_EQ: "finite (s::'q_52701 => bool) & finite (t::'q_52701 => bool)
-==> (CARD (s Un t) = CARD s + CARD t) = (s Int t = {})"
- by (import hollight CARD_UNION_OVERLAP_EQ)
-
-lemma CARD_UNION_OVERLAP: "finite (x::'q_52743 => bool) &
-finite (xa::'q_52743 => bool) & CARD (x Un xa) < CARD x + CARD xa
-==> x Int xa ~= {}"
- by (import hollight CARD_UNION_OVERLAP)
-
-lemma CARD_IMAGE_INJ: "(ALL (xa::'A) y::'A.
- xa : (x::'A => bool) & y : x & (f::'A => 'B) xa = f y --> xa = y) &
-finite x
-==> CARD (f ` x) = CARD x"
- by (import hollight CARD_IMAGE_INJ)
-
-lemma HAS_SIZE_IMAGE_INJ: "(ALL (xb::'A) y::'A.
- xb : (xa::'A => bool) & y : xa & (x::'A => 'B) xb = x y --> xb = y) &
-HAS_SIZE xa (xb::nat)
-==> HAS_SIZE (x ` xa) xb"
- by (import hollight HAS_SIZE_IMAGE_INJ)
-
-lemma CARD_IMAGE_LE: "finite (s::'A => bool) ==> CARD ((f::'A => 'B) ` s) <= CARD s"
- by (import hollight CARD_IMAGE_LE)
-
-lemma CARD_IMAGE_INJ_EQ: "finite (s::'A => bool) &
-(ALL x::'A. x : s --> (f::'A => 'B) x : (t::'B => bool)) &
-(ALL y::'B. y : t --> (EX! x::'A. x : s & f x = y))
-==> CARD t = CARD s"
- by (import hollight CARD_IMAGE_INJ_EQ)
-
-lemma CARD_SUBSET_IMAGE: "finite (t::'q_52977 => bool) &
-(s::'q_52984 => bool) <= (f::'q_52977 => 'q_52984) ` t
-==> CARD s <= CARD t"
- by (import hollight CARD_SUBSET_IMAGE)
-
-lemma HAS_SIZE_IMAGE_INJ_EQ: "(!!(x::'q_53049) y::'q_53049.
- x : (s::'q_53049 => bool) & y : s & (f::'q_53049 => 'q_53044) x = f y
- ==> x = y)
-==> HAS_SIZE (f ` s) (n::nat) = HAS_SIZE s n"
- by (import hollight HAS_SIZE_IMAGE_INJ_EQ)
-
-lemma CHOOSE_SUBSET_STRONG: "(finite (s::'A => bool) ==> (n::nat) <= CARD s) ==> EX t<=s. HAS_SIZE t n"
- by (import hollight CHOOSE_SUBSET_STRONG)
-
-lemma CHOOSE_SUBSET: "[| finite (s::'A => bool); (n::nat) <= CARD s |] ==> EX t<=s. HAS_SIZE t n"
- by (import hollight CHOOSE_SUBSET)
-
-lemma HAS_SIZE_PRODUCT_DEPENDENT: "HAS_SIZE (x::'A => bool) (xa::nat) &
-(ALL xa::'A. xa : x --> HAS_SIZE ((xb::'A => 'B => bool) xa) (xc::nat))
-==> HAS_SIZE
- {u::'A * 'B. EX (xa::'A) y::'B. (xa : x & y : xb xa) & u = (xa, y)}
- (xa * xc)"
- by (import hollight HAS_SIZE_PRODUCT_DEPENDENT)
-
-lemma FINITE_PRODUCT_DEPENDENT: "finite (s::'A => bool) &
-(ALL x::'A. x : s --> finite ((t::'A => 'B => bool) x))
-==> finite
- {u::'C.
- EX (x::'A) y::'B. (x : s & y : t x) & u = (f::'A => 'B => 'C) x y}"
- by (import hollight FINITE_PRODUCT_DEPENDENT)
-
-lemma FINITE_PRODUCT: "finite (x::'A => bool) & finite (xa::'B => bool)
-==> finite {u::'A * 'B. EX (xb::'A) y::'B. (xb : x & y : xa) & u = (xb, y)}"
- by (import hollight FINITE_PRODUCT)
-
-lemma CARD_PRODUCT: "finite (s::'A => bool) & finite (t::'B => bool)
-==> CARD {u::'A * 'B. EX (x::'A) y::'B. (x : s & y : t) & u = (x, y)} =
- CARD s * CARD t"
- by (import hollight CARD_PRODUCT)
-
-lemma HAS_SIZE_PRODUCT: "HAS_SIZE (x::'A => bool) (xa::nat) & HAS_SIZE (xb::'B => bool) (xc::nat)
-==> HAS_SIZE
- {u::'A * 'B. EX (xa::'A) y::'B. (xa : x & y : xb) & u = (xa, y)}
- (xa * xc)"
- by (import hollight HAS_SIZE_PRODUCT)
-
-definition
- CROSS :: "('q_53759 => bool) => ('q_53758 => bool) => 'q_53759 * 'q_53758 => bool" where
- "CROSS ==
-%(u::'q_53759 => bool) ua::'q_53758 => bool.
- {ub::'q_53759 * 'q_53758.
- EX (x::'q_53759) y::'q_53758. (x : u & y : ua) & ub = (x, y)}"
-
-lemma DEF_CROSS: "CROSS =
-(%(u::'q_53759 => bool) ua::'q_53758 => bool.
- {ub::'q_53759 * 'q_53758.
- EX (x::'q_53759) y::'q_53758. (x : u & y : ua) & ub = (x, y)})"
- by (import hollight DEF_CROSS)
-
-lemma IN_CROSS: "((x::'q_53795, xa::'q_53798)
- : CROSS (xb::'q_53795 => bool) (xc::'q_53798 => bool)) =
-(x : xb & xa : xc)"
- by (import hollight IN_CROSS)
-
-lemma HAS_SIZE_CROSS: "HAS_SIZE (x::'q_53823 => bool) (xb::nat) &
-HAS_SIZE (xa::'q_53826 => bool) (xc::nat)
-==> HAS_SIZE (CROSS x xa) (xb * xc)"
- by (import hollight HAS_SIZE_CROSS)
-
-lemma FINITE_CROSS: "finite (x::'q_53851 => bool) & finite (xa::'q_53853 => bool)
-==> finite (CROSS x xa)"
- by (import hollight FINITE_CROSS)
-
-lemma CARD_CROSS: "finite (x::'q_53874 => bool) & finite (xa::'q_53876 => bool)
-==> CARD (CROSS x xa) = CARD x * CARD xa"
- by (import hollight CARD_CROSS)
-
-lemma CROSS_EQ_EMPTY: "(CROSS (x::'q_53917 => bool) (xa::'q_53921 => bool) = {}) =
-(x = {} | xa = {})"
- by (import hollight CROSS_EQ_EMPTY)
-
-lemma HAS_SIZE_FUNSPACE: "HAS_SIZE (s::'A => bool) (m::nat) & HAS_SIZE (t::'B => bool) (n::nat)
-==> HAS_SIZE
- {u::'A => 'B.
- EX f::'A => 'B.
- ((ALL x::'A. x : s --> f x : t) &
- (ALL x::'A. x ~: s --> f x = (d::'B))) &
- u = f}
- (n ^ m)"
- by (import hollight HAS_SIZE_FUNSPACE)
-
-lemma CARD_FUNSPACE: "finite (s::'q_54227 => bool) & finite (t::'q_54224 => bool)
-==> CARD
- {u::'q_54227 => 'q_54224.
- EX f::'q_54227 => 'q_54224.
- ((ALL x::'q_54227. x : s --> f x : t) &
- (ALL x::'q_54227. x ~: s --> f x = (d::'q_54224))) &
- u = f} =
- CARD t ^ CARD s"
- by (import hollight CARD_FUNSPACE)
-
-lemma FINITE_FUNSPACE: "finite (s::'q_54293 => bool) & finite (t::'q_54290 => bool)
-==> finite
- {u::'q_54293 => 'q_54290.
- EX f::'q_54293 => 'q_54290.
- ((ALL x::'q_54293. x : s --> f x : t) &
- (ALL x::'q_54293. x ~: s --> f x = (d::'q_54290))) &
- u = f}"
- by (import hollight FINITE_FUNSPACE)
-
-lemma HAS_SIZE_POWERSET: "HAS_SIZE (s::'A => bool) (n::nat)