# HG changeset patch # User Cezary Kaliszyk # Date 1310484573 -32400 # Node ID fea3ed6951e3d30aa1ce31598e8189a5250accf7 # Parent 2bd54d4b5f3dccbf7672d6e6286e4701b93c3faa Update files generated in HOL/Import/HOLLight diff -r 2bd54d4b5f3d -r fea3ed6951e3 src/HOL/Import/HOLLight/HOLLight.thy --- a/src/HOL/Import/HOLLight/HOLLight.thy Wed Jul 13 00:23:24 2011 +0900 +++ b/src/HOL/Import/HOLLight/HOLLight.thy Wed Jul 13 00:29:33 2011 +0900 @@ -1,6 +1,6 @@ (* AUTOMATICALLY GENERATED, DO NOT EDIT! *) -theory HOLLight = "../HOLLightCompat" + "../HOL4Syntax": +theory HOLLight imports "../HOLLightCompat" "../HOL4Syntax" begin ;setup_theory hollight @@ -13,235 +13,290 @@ lemma DEF__FALSITY_: "_FALSITY_ = False" by (import hollight DEF__FALSITY_) -lemma CONJ_ACI: "((p::bool) & (q::bool)) = (q & p) & -((p & q) & (r::bool)) = (p & q & r) & +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::bool) | (q::bool)) = (q | p) & -((p | q) | (r::bool)) = (p | q | r) & +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 EQ_CLAUSES: "ALL t::bool. - (True = t) = t & - (t = True) = t & (False = t) = (~ t) & (t = False) = (~ t)" +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: "ALL t::bool. - (True & t) = t & - (t & True) = t & (False & t) = False & (t & False) = False & (t & t) = t" +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: "ALL t::bool. - (True | t) = True & - (t | True) = True & (False | t) = t & (t | False) = t & (t | t) = t" +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: "ALL t::bool. - (True --> t) = t & - (t --> True) = True & - (False --> t) = True & (t --> t) = True & (t --> False) = (~ t)" +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_864::type) = x --> (p::bool)) = p" +lemma IMP_EQ_CLAUSE: "((x::'q_851) = x --> (p::bool)) = p" by (import hollight IMP_EQ_CLAUSE) -lemma SWAP_FORALL_THM: "ALL P::'A::type => 'B::type => bool. - (ALL x::'A::type. All (P x)) = (ALL (y::'B::type) x::'A::type. P x y)" - by (import hollight SWAP_FORALL_THM) - -lemma SWAP_EXISTS_THM: "ALL P::'A::type => 'B::type => bool. - (EX x::'A::type. Ex (P x)) = (EX (x::'B::type) xa::'A::type. P xa x)" - by (import hollight SWAP_EXISTS_THM) - -lemma TRIV_EXISTS_AND_THM: "ALL (P::bool) Q::bool. - (EX x::'A::type. P & Q) = ((EX x::'A::type. P) & (EX x::'A::type. Q))" +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: "ALL (P::bool) Q::bool. - ((EX x::'A::type. P) & (EX x::'A::type. Q)) = (EX x::'A::type. P & Q)" +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 (P::bool) Q::bool. - (ALL x::'A::type. P | Q) = ((ALL x::'A::type. P) | (ALL x::'A::type. Q))" +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 (P::bool) Q::bool. - ((ALL x::'A::type. P) | (ALL x::'A::type. Q)) = (ALL x::'A::type. P | Q)" +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 (P::bool) Q::bool. - (ALL x::'A::type. P --> Q) = - ((EX x::'A::type. P) --> (ALL x::'A::type. Q))" +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: "ALL (P::bool) Q::bool. - (EX x::'A::type. P --> Q) = - ((ALL x::'A::type. P) --> (EX x::'A::type. Q))" +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: "ALL P::'A::type => bool. - Ex1 P = (EX x::'A::type. ALL y::'A::type. P y = (x = y))" +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: "ALL (P::'A::type => bool) x::'A::type. - (ALL y::'A::type. P y = (y = x)) --> Eps P = x" +lemma SELECT_UNIQUE: "(!!y::'A. (P::'A => bool) y = (y = (x::'A))) ==> Eps P = x" by (import hollight SELECT_UNIQUE) -lemma EXCLUDED_MIDDLE: "ALL t::bool. t | ~ t" +lemma EXCLUDED_MIDDLE: "t | ~ t" by (import hollight EXCLUDED_MIDDLE) -definition COND :: "bool => 'A => 'A => 'A" where - "COND == -%(t::bool) (t1::'A::type) t2::'A::type. - SOME x::'A::type. (t = True --> x = t1) & (t = False --> x = t2)" - -lemma DEF_COND: "COND = -(%(t::bool) (t1::'A::type) t2::'A::type. - SOME x::'A::type. (t = True --> x = t1) & (t = False --> x = t2))" - by (import hollight DEF_COND) - -lemma COND_CLAUSES: "ALL (x::'A::type) xa::'A::type. COND True x xa = x & COND False x xa = xa" +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: "ALL (b::bool) (t1::bool) t2::bool. COND b t1 t2 = ((~ b | t1) & (b | t2))" +lemma COND_EXPAND: "(if b then t1 else t2) = ((~ b | t1) & (b | t2))" by (import hollight COND_EXPAND) -lemma COND_ID: "ALL (b::bool) t::'A::type. COND b t t = t" - by (import hollight COND_ID) - -lemma COND_RAND: "ALL (b::bool) (f::'A::type => 'B::type) (x::'A::type) y::'A::type. - f (COND b x y) = COND b (f x) (f y)" - by (import hollight COND_RAND) - -lemma COND_RATOR: "ALL (b::bool) (f::'A::type => 'B::type) (g::'A::type => 'B::type) - x::'A::type. COND b f g x = COND b (f x) (g x)" +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: "ALL (b::bool) (f::'A::type => 'B::type) g::'A::type => 'B::type. - (%x::'A::type. COND b (f x) (g x)) = COND b f g" +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::bool) --> (B::bool)) & ((C::bool) --> (D::bool)) --> -COND (b::bool) A C --> COND b B D" +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 COND_ELIM_THM: "(P::'A::type => bool) (COND (c::bool) (x::'A::type) (y::'A::type)) = -((c --> P x) & (~ c --> P y))" - by (import hollight COND_ELIM_THM) - -lemma SKOLEM_THM: "ALL P::'A::type => 'B::type => bool. - (ALL x::'A::type. Ex (P x)) = - (EX x::'A::type => 'B::type. ALL xa::'A::type. P xa (x xa))" +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 P::'A::type => 'B::type => bool. - (ALL x::'A::type. Ex1 (P x)) = - (EX f::'A::type => 'B::type. - ALL (x::'A::type) y::'B::type. P x y = (f x = y))" +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: "COND ((x::'q_3000::type) = x) (y::'q_2993::type) (z::'q_2993::type) = y" +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 o_ASSOC: "ALL (f::'C::type => 'D::type) (g::'B::type => 'C::type) - h::'A::type => 'B::type. f o (g o h) = f o g o h" +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: "ALL f::'A::type => 'B::type. id o f = f & f o id = f" +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::bool. x" +lemma EXISTS_ONE_REP: "EX x. x" by (import hollight EXISTS_ONE_REP) -lemma one_axiom: "ALL f::'A::type => unit. All (op = f)" +lemma one_axiom: "(f::'A => unit) = (x::'A => unit)" by (import hollight one_axiom) -lemma one_RECURSION: "ALL e::'A::type. EX x::unit => 'A::type. x () = e" +lemma one_RECURSION: "EX x::unit => 'A. x () = (e::'A)" by (import hollight one_RECURSION) -lemma one_Axiom: "ALL e::'A::type. EX! fn::unit => 'A::type. fn () = e" +lemma one_Axiom: "EX! fn::unit => 'A. fn () = (e::'A)" by (import hollight one_Axiom) -lemma th_cond: "(P::'A::type => bool => bool) (COND (b::bool) (x::'A::type) (y::'A::type)) - b = -(b & P x True | ~ b & P y False)" +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::type. t" - -lemma DEF_LET_END: "LET_END = (%t::'A::type. t)" +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) -definition GABS :: "('A => bool) => 'A" where - "(op ==::(('A::type => bool) => 'A::type) - => (('A::type => bool) => 'A::type) => prop) - (GABS::('A::type => bool) => 'A::type) - (Eps::('A::type => bool) => 'A::type)" - -lemma DEF_GABS: "(op =::(('A::type => bool) => 'A::type) - => (('A::type => bool) => 'A::type) => bool) - (GABS::('A::type => bool) => 'A::type) - (Eps::('A::type => bool) => 'A::type)" - by (import hollight DEF_GABS) - -definition GEQ :: "'A => 'A => bool" where - "(op ==::('A::type => 'A::type => bool) - => ('A::type => 'A::type => bool) => prop) - (GEQ::'A::type => 'A::type => bool) (op =::'A::type => 'A::type => bool)" - -lemma DEF_GEQ: "(op =::('A::type => 'A::type => bool) - => ('A::type => 'A::type => bool) => bool) - (GEQ::'A::type => 'A::type => bool) (op =::'A::type => 'A::type => bool)" - by (import hollight DEF_GEQ) - -lemma PAIR_EXISTS_THM: "EX (x::'A::type => 'B::type => bool) (a::'A::type) b::'B::type. - x = Pair_Rep a b" +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) -definition CURRY :: "('A * 'B => 'C) => 'A => 'B => 'C" where - "CURRY == -%(u::'A::type * 'B::type => 'C::type) (ua::'A::type) ub::'B::type. - u (ua, ub)" - -lemma DEF_CURRY: "CURRY = -(%(u::'A::type * 'B::type => 'C::type) (ua::'A::type) ub::'B::type. - u (ua, ub))" - by (import hollight DEF_CURRY) - -definition UNCURRY :: "('A => 'B => 'C) => 'A * 'B => 'C" where - "UNCURRY == -%(u::'A::type => 'B::type => 'C::type) ua::'A::type * 'B::type. - u (fst ua) (snd ua)" - -lemma DEF_UNCURRY: "UNCURRY = -(%(u::'A::type => 'B::type => 'C::type) ua::'A::type * 'B::type. - u (fst ua) (snd ua))" +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 +definition + PASSOC :: "(('A * 'B) * 'C => 'D) => 'A * 'B * 'C => 'D" where "PASSOC == -%(u::('A::type * 'B::type) * 'C::type => 'D::type) - ua::'A::type * 'B::type * 'C::type. +%(u::('A * 'B) * 'C => 'D) ua::'A * 'B * 'C. u ((fst ua, fst (snd ua)), snd (snd ua))" lemma DEF_PASSOC: "PASSOC = -(%(u::('A::type * 'B::type) * 'C::type => 'D::type) - ua::'A::type * 'B::type * 'C::type. +(%(u::('A * 'B) * 'C => 'D) ua::'A * 'B * 'C. u ((fst ua, fst (snd ua)), snd (snd ua)))" by (import hollight DEF_PASSOC) -lemma num_Axiom: "ALL (e::'A::type) f::'A::type => nat => 'A::type. - EX! fn::nat => 'A::type. fn 0 = e & (ALL n::nat. fn (Suc n) = f (fn n) n)" - by (import hollight num_Axiom) - -lemma ADD_CLAUSES: "(ALL x::nat. 0 + x = x) & -(ALL x::nat. x + 0 = x) & +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) @@ -250,25 +305,25 @@ m + n + (p::nat) = m + (n + p) & m + (n + p) = n + (m + p)" by (import hollight ADD_AC) -lemma EQ_ADD_LCANCEL_0: "ALL (m::nat) n::nat. (m + n = m) = (n = 0)" +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: "ALL (x::nat) xa::nat. (x + xa = xa) = (x = 0)" +lemma EQ_ADD_RCANCEL_0: "((x::nat) + (xa::nat) = xa) = (x = (0::nat))" by (import hollight EQ_ADD_RCANCEL_0) -lemma ONE: "NUMERAL_BIT1 0 = Suc 0" - by (import hollight ONE) - -lemma TWO: "NUMERAL_BIT0 (NUMERAL_BIT1 0) = Suc (NUMERAL_BIT1 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 ADD1: "ALL x::nat. Suc x = x + NUMERAL_BIT1 0" - by (import hollight ADD1) - -lemma MULT_CLAUSES: "(ALL x::nat. 0 * x = 0) & -(ALL x::nat. x * 0 = 0) & -(ALL x::nat. NUMERAL_BIT1 0 * x = x) & -(ALL x::nat. x * NUMERAL_BIT1 0 = x) & +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) @@ -277,1496 +332,831 @@ m * n * (p::nat) = m * (n * p) & m * (n * p) = n * (m * p)" by (import hollight MULT_AC) -lemma MULT_2: "ALL n::nat. NUMERAL_BIT0 (NUMERAL_BIT1 0) * n = n + n" - by (import hollight MULT_2) - -lemma MULT_EQ_1: "ALL (m::nat) n::nat. - (m * n = NUMERAL_BIT1 0) = (m = NUMERAL_BIT1 0 & n = NUMERAL_BIT1 0)" - by (import hollight MULT_EQ_1) - -definition EXP :: "nat => nat => nat" where - "EXP == -SOME EXP::nat => nat => nat. - (ALL m::nat. EXP m 0 = NUMERAL_BIT1 0) & - (ALL (m::nat) n::nat. EXP m (Suc n) = m * EXP m n)" - -lemma DEF_EXP: "EXP = -(SOME EXP::nat => nat => nat. - (ALL m::nat. EXP m 0 = NUMERAL_BIT1 0) & - (ALL (m::nat) n::nat. EXP m (Suc n) = m * EXP m n))" - by (import hollight DEF_EXP) - -lemma EXP_EQ_0: "ALL (m::nat) n::nat. (EXP m n = 0) = (m = 0 & n ~= 0)" - by (import hollight EXP_EQ_0) - -lemma EXP_ADD: "ALL (m::nat) (n::nat) p::nat. EXP m (n + p) = EXP m n * EXP m p" - by (import hollight EXP_ADD) - -lemma EXP_ONE: "ALL n::nat. EXP (NUMERAL_BIT1 0) n = NUMERAL_BIT1 0" - by (import hollight EXP_ONE) - -lemma EXP_1: "ALL x::nat. EXP x (NUMERAL_BIT1 0) = x" - by (import hollight EXP_1) - -lemma EXP_2: "ALL x::nat. EXP x (NUMERAL_BIT0 (NUMERAL_BIT1 0)) = x * x" - by (import hollight EXP_2) - -lemma MULT_EXP: "ALL (p::nat) (m::nat) n::nat. EXP (m * n) p = EXP m p * EXP n p" - by (import hollight MULT_EXP) - -lemma EXP_MULT: "ALL (m::nat) (n::nat) p::nat. EXP m (n * p) = EXP (EXP m n) p" - by (import hollight EXP_MULT) - -consts - "<=" :: "nat => nat => bool" ("<=") - -defs - "<=_def": "<= == -SOME u::nat => nat => bool. - (ALL m::nat. u m 0 = (m = 0)) & - (ALL (m::nat) n::nat. u m (Suc n) = (m = Suc n | u m n))" - -lemma DEF__lessthan__equal_: "<= = -(SOME u::nat => nat => bool. - (ALL m::nat. u m 0 = (m = 0)) & - (ALL (m::nat) n::nat. u m (Suc n) = (m = Suc n | u m n)))" - by (import hollight DEF__lessthan__equal_) - -consts - "<" :: "nat => nat => bool" ("<") - -defs - "<_def": "< == -SOME u::nat => nat => bool. - (ALL m::nat. u m 0 = False) & - (ALL (m::nat) n::nat. u m (Suc n) = (m = n | u m n))" - -lemma DEF__lessthan_: "< = -(SOME u::nat => nat => bool. - (ALL m::nat. u m 0 = False) & - (ALL (m::nat) n::nat. u m (Suc n) = (m = n | u m n)))" - by (import hollight DEF__lessthan_) - -consts - ">=" :: "nat => nat => bool" (">=") - -defs - ">=_def": ">= == %(u::nat) ua::nat. <= ua u" - -lemma DEF__greaterthan__equal_: ">= = (%(u::nat) ua::nat. <= ua u)" - by (import hollight DEF__greaterthan__equal_) - -consts - ">" :: "nat => nat => bool" (">") - -defs - ">_def": "> == %(u::nat) ua::nat. < ua u" - -lemma DEF__greaterthan_: "> = (%(u::nat) ua::nat. < ua u)" - by (import hollight DEF__greaterthan_) - -lemma LE_SUC_LT: "ALL (m::nat) n::nat. <= (Suc m) n = < m n" - by (import hollight LE_SUC_LT) - -lemma LT_SUC_LE: "ALL (m::nat) n::nat. < m (Suc n) = <= m n" - by (import hollight LT_SUC_LE) - -lemma LE_SUC: "ALL (x::nat) xa::nat. <= (Suc x) (Suc xa) = <= x xa" - by (import hollight LE_SUC) - -lemma LT_SUC: "ALL (x::nat) xa::nat. < (Suc x) (Suc xa) = < x xa" - by (import hollight LT_SUC) - -lemma LE_0: "All (<= 0)" - by (import hollight LE_0) - -lemma LT_0: "ALL x::nat. < 0 (Suc x)" - by (import hollight LT_0) - -lemma LE_REFL: "ALL n::nat. <= n n" - by (import hollight LE_REFL) - -lemma LT_REFL: "ALL n::nat. ~ < n n" - by (import hollight LT_REFL) - -lemma LE_ANTISYM: "ALL (m::nat) n::nat. (<= m n & <= n m) = (m = n)" - by (import hollight LE_ANTISYM) - -lemma LT_ANTISYM: "ALL (m::nat) n::nat. ~ (< m n & < n m)" +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: "ALL (m::nat) n::nat. ~ (<= m n & < n m)" +lemma LET_ANTISYM: "~ ((m::nat) <= (n::nat) & n < m)" by (import hollight LET_ANTISYM) -lemma LTE_ANTISYM: "ALL (x::nat) xa::nat. ~ (< x xa & <= xa x)" +lemma LTE_ANTISYM: "~ ((x::nat) < (xa::nat) & xa <= x)" by (import hollight LTE_ANTISYM) -lemma LE_TRANS: "ALL (m::nat) (n::nat) p::nat. <= m n & <= n p --> <= m p" - by (import hollight LE_TRANS) - -lemma LT_TRANS: "ALL (m::nat) (n::nat) p::nat. < m n & < n p --> < m p" - by (import hollight LT_TRANS) - -lemma LET_TRANS: "ALL (m::nat) (n::nat) p::nat. <= m n & < n p --> < m p" - by (import hollight LET_TRANS) - -lemma LTE_TRANS: "ALL (m::nat) (n::nat) p::nat. < m n & <= n p --> < m p" - by (import hollight LTE_TRANS) - -lemma LE_CASES: "ALL (m::nat) n::nat. <= m n | <= n m" - by (import hollight LE_CASES) - -lemma LT_CASES: "ALL (m::nat) n::nat. < m n | < n m | m = n" +lemma LT_CASES: "(m::nat) < (n::nat) | n < m | m = n" by (import hollight LT_CASES) -lemma LET_CASES: "ALL (m::nat) n::nat. <= m n | < n m" - by (import hollight LET_CASES) - -lemma LTE_CASES: "ALL (x::nat) xa::nat. < x xa | <= xa x" +lemma LTE_CASES: "(x::nat) < (xa::nat) | xa <= x" by (import hollight LTE_CASES) -lemma LT_NZ: "ALL n::nat. < 0 n = (n ~= 0)" - by (import hollight LT_NZ) - -lemma LE_LT: "ALL (m::nat) n::nat. <= m n = (< m n | m = n)" - by (import hollight LE_LT) - -lemma LT_LE: "ALL (x::nat) xa::nat. < x xa = (<= x xa & x ~= xa)" - by (import hollight LT_LE) - -lemma NOT_LE: "ALL (m::nat) n::nat. (~ <= m n) = < n m" - by (import hollight NOT_LE) - -lemma NOT_LT: "ALL (m::nat) n::nat. (~ < m n) = <= n m" - by (import hollight NOT_LT) - -lemma LT_IMP_LE: "ALL (x::nat) xa::nat. < x xa --> <= x xa" - by (import hollight LT_IMP_LE) - -lemma EQ_IMP_LE: "ALL (m::nat) n::nat. m = n --> <= m n" - by (import hollight EQ_IMP_LE) - -lemma LE_EXISTS: "ALL (m::nat) n::nat. <= m n = (EX d::nat. n = m + d)" - by (import hollight LE_EXISTS) - -lemma LT_EXISTS: "ALL (m::nat) n::nat. < m n = (EX d::nat. n = m + Suc d)" +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 LE_ADD: "ALL (m::nat) n::nat. <= m (m + n)" - by (import hollight LE_ADD) - -lemma LE_ADDR: "ALL (x::nat) xa::nat. <= xa (x + xa)" - by (import hollight LE_ADDR) - -lemma LT_ADD: "ALL (m::nat) n::nat. < m (m + n) = < 0 n" +lemma LT_ADD: "((m::nat) < m + (n::nat)) = ((0::nat) < n)" by (import hollight LT_ADD) -lemma LT_ADDR: "ALL (x::nat) xa::nat. < xa (x + xa) = < 0 x" +lemma LT_ADDR: "((xa::nat) < (x::nat) + xa) = ((0::nat) < x)" by (import hollight LT_ADDR) -lemma LE_ADD_LCANCEL: "ALL (x::nat) (xa::nat) xb::nat. <= (x + xa) (x + xb) = <= xa xb" - by (import hollight LE_ADD_LCANCEL) - -lemma LE_ADD_RCANCEL: "ALL (x::nat) (xa::nat) xb::nat. <= (x + xb) (xa + xb) = <= x xa" - by (import hollight LE_ADD_RCANCEL) - -lemma LT_ADD_LCANCEL: "ALL (x::nat) (xa::nat) xb::nat. < (x + xa) (x + xb) = < xa xb" - by (import hollight LT_ADD_LCANCEL) - -lemma LT_ADD_RCANCEL: "ALL (x::nat) (xa::nat) xb::nat. < (x + xb) (xa + xb) = < x xa" - by (import hollight LT_ADD_RCANCEL) - -lemma LE_ADD2: "ALL (m::nat) (n::nat) (p::nat) q::nat. - <= m p & <= n q --> <= (m + n) (p + q)" - by (import hollight LE_ADD2) - -lemma LET_ADD2: "ALL (m::nat) (n::nat) (p::nat) q::nat. <= m p & < n q --> < (m + n) (p + q)" - by (import hollight LET_ADD2) - -lemma LTE_ADD2: "ALL (x::nat) (xa::nat) (xb::nat) xc::nat. - < x xb & <= xa xc --> < (x + xa) (xb + xc)" - by (import hollight LTE_ADD2) - -lemma LT_ADD2: "ALL (m::nat) (n::nat) (p::nat) q::nat. < m p & < n q --> < (m + n) (p + q)" - by (import hollight LT_ADD2) - -lemma LT_MULT: "ALL (m::nat) n::nat. < 0 (m * n) = (< 0 m & < 0 n)" - by (import hollight LT_MULT) - -lemma LE_MULT2: "ALL (m::nat) (n::nat) (p::nat) q::nat. - <= m n & <= p q --> <= (m * p) (n * q)" - by (import hollight LE_MULT2) - -lemma LT_LMULT: "ALL (m::nat) (n::nat) p::nat. m ~= 0 & < n p --> < (m * n) (m * p)" +lemma LT_LMULT: "(m::nat) ~= (0::nat) & (n::nat) < (p::nat) ==> m * n < m * p" by (import hollight LT_LMULT) -lemma LE_MULT_LCANCEL: "ALL (m::nat) (n::nat) p::nat. <= (m * n) (m * p) = (m = 0 | <= n p)" +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: "ALL (x::nat) (xa::nat) xb::nat. <= (x * xb) (xa * xb) = (<= x xa | xb = 0)" +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: "ALL (m::nat) (n::nat) p::nat. < (m * n) (m * p) = (m ~= 0 & < n p)" +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: "ALL (x::nat) (xa::nat) xb::nat. < (x * xb) (xa * xb) = (< x xa & xb ~= 0)" +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: "ALL (m::nat) (n::nat) (p::nat) q::nat. < m n & < p q --> < (m * p) (n * q)" +lemma LT_MULT2: "(m::nat) < (n::nat) & (p::nat) < (q::nat) ==> m * p < n * q" by (import hollight LT_MULT2) -lemma LE_SQUARE_REFL: "ALL n::nat. <= n (n * n)" - by (import hollight LE_SQUARE_REFL) - 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) --> -(ALL m::nat. All (P 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) --> -(ALL m::nat. All (P m))" +(ALL (m::nat) n::nat. m < n --> P m n) +==> P (m::nat) (x::nat)" by (import hollight WLOG_LT) -lemma num_WF: "ALL P::nat => bool. - (ALL n::nat. (ALL m::nat. < m n --> P m) --> P n) --> All P" - by (import hollight num_WF) - -lemma num_WOP: "ALL P::nat => bool. Ex P = (EX n::nat. P n & (ALL m::nat. < m n --> ~ P m))" +lemma num_WOP: "Ex (P::nat => bool) = (EX n::nat. P n & (ALL m bool. - (Ex P & (EX M::nat. ALL x::nat. P x --> <= x M)) = - (EX m::nat. P m & (ALL x::nat. P x --> <= x m))" +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) -definition EVEN :: "nat => bool" where - "EVEN == -SOME EVEN::nat => bool. - EVEN 0 = True & (ALL n::nat. EVEN (Suc n) = (~ EVEN n))" - -lemma DEF_EVEN: "EVEN = -(SOME EVEN::nat => bool. - EVEN 0 = True & (ALL n::nat. EVEN (Suc n) = (~ EVEN n)))" - by (import hollight DEF_EVEN) - -definition ODD :: "nat => bool" where - "ODD == -SOME ODD::nat => bool. ODD 0 = False & (ALL n::nat. ODD (Suc n) = (~ ODD n))" - -lemma DEF_ODD: "ODD = -(SOME ODD::nat => bool. - ODD 0 = False & (ALL n::nat. ODD (Suc n) = (~ ODD n)))" - by (import hollight DEF_ODD) - -lemma NOT_EVEN: "ALL n::nat. (~ EVEN n) = ODD n" +lemma NOT_EVEN: "odd (n::nat) = odd n" by (import hollight NOT_EVEN) -lemma NOT_ODD: "ALL n::nat. (~ ODD n) = EVEN n" +lemma NOT_ODD: "(~ odd (n::nat)) = even n" by (import hollight NOT_ODD) -lemma EVEN_OR_ODD: "ALL n::nat. EVEN n | ODD n" +lemma EVEN_OR_ODD: "even (n::nat) | odd n" by (import hollight EVEN_OR_ODD) -lemma EVEN_AND_ODD: "ALL x::nat. ~ (EVEN x & ODD x)" +lemma EVEN_AND_ODD: "~ (even (x::nat) & odd x)" by (import hollight EVEN_AND_ODD) -lemma EVEN_ADD: "ALL (m::nat) n::nat. EVEN (m + n) = (EVEN m = EVEN n)" - by (import hollight EVEN_ADD) - -lemma EVEN_MULT: "ALL (m::nat) n::nat. EVEN (m * n) = (EVEN m | EVEN n)" - by (import hollight EVEN_MULT) - -lemma EVEN_EXP: "ALL (m::nat) n::nat. EVEN (EXP m n) = (EVEN m & n ~= 0)" +lemma EVEN_EXP: "even ((m::nat) ^ (n::nat)) = (even m & n ~= (0::nat))" by (import hollight EVEN_EXP) -lemma ODD_ADD: "ALL (m::nat) n::nat. ODD (m + n) = (ODD m ~= ODD n)" - by (import hollight ODD_ADD) - -lemma ODD_MULT: "ALL (m::nat) n::nat. ODD (m * n) = (ODD m & ODD n)" +lemma ODD_MULT: "odd ((m::nat) * (n::nat)) = (odd m & odd n)" by (import hollight ODD_MULT) -lemma ODD_EXP: "ALL (m::nat) n::nat. ODD (EXP m n) = (ODD m | n = 0)" +lemma ODD_EXP: "odd ((m::nat) ^ (n::nat)) = (odd m | n = (0::nat))" by (import hollight ODD_EXP) -lemma EVEN_DOUBLE: "ALL n::nat. EVEN (NUMERAL_BIT0 (NUMERAL_BIT1 0) * n)" +lemma EVEN_DOUBLE: "even ((2::nat) * (n::nat))" by (import hollight EVEN_DOUBLE) -lemma ODD_DOUBLE: "ALL x::nat. ODD (Suc (NUMERAL_BIT0 (NUMERAL_BIT1 0) * x))" +lemma ODD_DOUBLE: "odd (Suc (2 * x))" by (import hollight ODD_DOUBLE) -lemma EVEN_EXISTS_LEMMA: "ALL n::nat. - (EVEN n --> (EX m::nat. n = NUMERAL_BIT0 (NUMERAL_BIT1 0) * m)) & - (~ EVEN n --> (EX m::nat. n = Suc (NUMERAL_BIT0 (NUMERAL_BIT1 0) * m)))" +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_EXISTS: "ALL n::nat. EVEN n = (EX m::nat. n = NUMERAL_BIT0 (NUMERAL_BIT1 0) * m)" - by (import hollight EVEN_EXISTS) - -lemma ODD_EXISTS: "ALL n::nat. ODD n = (EX m::nat. n = Suc (NUMERAL_BIT0 (NUMERAL_BIT1 0) * m))" - by (import hollight ODD_EXISTS) - -lemma EVEN_ODD_DECOMPOSITION: "ALL n::nat. - (EX (k::nat) m::nat. - ODD m & n = EXP (NUMERAL_BIT0 (NUMERAL_BIT1 0)) k * m) = - (n ~= 0)" +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: "ALL x::nat. 0 - x = 0 & x - 0 = x" +lemma SUB_0: "(0::nat) - (x::nat) = (0::nat) & x - (0::nat) = x" by (import hollight SUB_0) -lemma SUB_PRESUC: "ALL (m::nat) n::nat. Pred (Suc m - n) = m - n" +lemma SUB_PRESUC: "Suc m - n - Suc 0 = m - n" by (import hollight SUB_PRESUC) -lemma SUB_EQ_0: "ALL (m::nat) n::nat. (m - n = 0) = <= m n" - by (import hollight SUB_EQ_0) - -lemma ADD_SUBR: "ALL (x::nat) xa::nat. xa - (x + xa) = 0" +lemma ADD_SUBR: "(xa::nat) - ((x::nat) + xa) = (0::nat)" by (import hollight ADD_SUBR) -lemma SUB_ADD: "ALL (x::nat) xa::nat. <= xa x --> x - xa + xa = x" - by (import hollight SUB_ADD) - -lemma SUC_SUB1: "ALL x::nat. Suc x - NUMERAL_BIT1 0 = x" - by (import hollight SUC_SUB1) - -definition FACT :: "nat => nat" where - "FACT == -SOME FACT::nat => nat. - FACT 0 = NUMERAL_BIT1 0 & (ALL n::nat. FACT (Suc n) = Suc n * FACT n)" - -lemma DEF_FACT: "FACT = -(SOME FACT::nat => nat. - FACT 0 = NUMERAL_BIT1 0 & (ALL n::nat. FACT (Suc n) = Suc n * FACT n))" - by (import hollight DEF_FACT) - -lemma FACT_LT: "ALL n::nat. < 0 (FACT n)" - by (import hollight FACT_LT) - -lemma FACT_LE: "ALL x::nat. <= (NUMERAL_BIT1 0) (FACT x)" - by (import hollight FACT_LE) - -lemma FACT_MONO: "ALL (m::nat) n::nat. <= m n --> <= (FACT m) (FACT n)" - by (import hollight FACT_MONO) - -lemma DIVMOD_EXIST: "ALL (m::nat) n::nat. n ~= 0 --> (EX (q::nat) r::nat. m = q * n + r & < r n)" +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: "ALL (m::nat) n::nat. - EX (x::nat) xa::nat. - COND (n = 0) (x = 0 & xa = 0) (m = x * n + xa & < xa n)" +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) -definition DIV :: "nat => nat => nat" where - "DIV == -SOME q::nat => nat => nat. - EX r::nat => nat => nat. - ALL (m::nat) n::nat. - COND (n = 0) (q m n = 0 & r m n = 0) - (m = q m n * n + r m n & < (r m n) n)" - -lemma DEF_DIV: "DIV = -(SOME q::nat => nat => nat. - EX r::nat => nat => nat. - ALL (m::nat) n::nat. - COND (n = 0) (q m n = 0 & r m n = 0) - (m = q m n * n + r m n & < (r m n) n))" - by (import hollight DEF_DIV) - -definition MOD :: "nat => nat => nat" where - "MOD == -SOME r::nat => nat => nat. - ALL (m::nat) n::nat. - COND (n = 0) (DIV m n = 0 & r m n = 0) - (m = DIV m n * n + r m n & < (r m n) n)" - -lemma DEF_MOD: "MOD = -(SOME r::nat => nat => nat. - ALL (m::nat) n::nat. - COND (n = 0) (DIV m n = 0 & r m n = 0) - (m = DIV m n * n + r m n & < (r m n) n))" - by (import hollight DEF_MOD) - -lemma DIVISION: "ALL (m::nat) n::nat. n ~= 0 --> m = DIV m n * n + MOD m n & < (MOD m n) n" +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: "ALL (m::nat) (n::nat) (q1::nat) (r1::nat) (q2::nat) r2::nat. - (m = q1 * n + r1 & < r1 n) & m = q2 * n + r2 & < r2 n --> - q1 = q2 & r1 = r2" +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: "ALL (m::nat) (n::nat) (q::nat) r::nat. - m = q * n + r & < r n --> DIV m n = q & MOD m n = r" +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: "ALL (m::nat) (n::nat) (q::nat) r::nat. m = q * n + r & < r n --> MOD m n = r" +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: "ALL (m::nat) (n::nat) (q::nat) r::nat. m = q * n + r & < r n --> DIV m n = q" +lemma DIV_UNIQ: "(m::nat) = (q::nat) * (n::nat) + (r::nat) & r < n ==> m div n = q" by (import hollight DIV_UNIQ) -lemma MOD_MULT: "ALL (x::nat) xa::nat. x ~= 0 --> MOD (x * xa) x = 0" - by (import hollight MOD_MULT) - -lemma DIV_MULT: "ALL (x::nat) xa::nat. x ~= 0 --> DIV (x * xa) x = xa" - by (import hollight DIV_MULT) - -lemma DIV_DIV: "ALL (m::nat) (n::nat) p::nat. n * p ~= 0 --> DIV (DIV m n) p = DIV m (n * p)" - by (import hollight DIV_DIV) - -lemma MOD_LT: "ALL (m::nat) n::nat. < m n --> MOD m n = m" - by (import hollight MOD_LT) - -lemma MOD_EQ: "ALL (m::nat) (n::nat) (p::nat) q::nat. m = n + q * p --> MOD m p = MOD n p" +lemma MOD_EQ: "(m::nat) = (n::nat) + (q::nat) * (p::nat) ==> m mod p = n mod p" by (import hollight MOD_EQ) -lemma DIV_MOD: "ALL (m::nat) (n::nat) p::nat. - n * p ~= 0 --> MOD (DIV m n) p = DIV (MOD m (n * p)) n" - by (import hollight DIV_MOD) - -lemma DIV_1: "ALL n::nat. DIV n (NUMERAL_BIT1 0) = n" - by (import hollight DIV_1) - -lemma EXP_LT_0: "ALL (x::nat) xa::nat. < 0 (EXP xa x) = (xa ~= 0 | x = 0)" - by (import hollight EXP_LT_0) - -lemma DIV_LE: "ALL (m::nat) n::nat. n ~= 0 --> <= (DIV m n) m" +lemma DIV_LE: "(n::nat) ~= (0::nat) ==> (m::nat) div n <= m" by (import hollight DIV_LE) -lemma DIV_MUL_LE: "ALL (m::nat) n::nat. <= (n * DIV m n) m" +lemma DIV_MUL_LE: "(n::nat) * ((m::nat) div n) <= m" by (import hollight DIV_MUL_LE) -lemma DIV_0: "ALL n::nat. n ~= 0 --> DIV 0 n = 0" - by (import hollight DIV_0) - -lemma MOD_0: "ALL n::nat. n ~= 0 --> MOD 0 n = 0" - by (import hollight MOD_0) - -lemma DIV_LT: "ALL (m::nat) n::nat. < m n --> DIV m n = 0" - by (import hollight DIV_LT) - -lemma MOD_MOD: "ALL (m::nat) (n::nat) p::nat. n * p ~= 0 --> MOD (MOD m (n * p)) n = MOD m n" +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: "ALL (m::nat) n::nat. n ~= 0 --> MOD (MOD m n) n = MOD m n" +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: "ALL (x::nat) (xa::nat) xb::nat. - x * xb ~= 0 --> DIV (x * xa) (x * xb) = DIV xa xb" +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: "ALL (x::nat) (xa::nat) xb::nat. - x * xb ~= 0 --> MOD (x * xa) (x * xb) = x * MOD xa xb" +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_1: "ALL n::nat. MOD n (NUMERAL_BIT1 0) = 0" - by (import hollight MOD_1) - -lemma MOD_EXISTS: "ALL (m::nat) n::nat. - (EX q::nat. m = n * q) = COND (n = 0) (m = 0) (MOD m n = 0)" +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 LT_EXP: "ALL (x::nat) (m::nat) n::nat. - < (EXP x m) (EXP x n) = - (<= (NUMERAL_BIT0 (NUMERAL_BIT1 0)) x & < m n | x = 0 & m ~= 0 & n = 0)" - by (import hollight LT_EXP) - -lemma LE_EXP: "ALL (x::nat) (m::nat) n::nat. - <= (EXP x m) (EXP x n) = - COND (x = 0) (m = 0 --> n = 0) (x = NUMERAL_BIT1 0 | <= m n)" - by (import hollight LE_EXP) - -lemma DIV_MONO: "ALL (m::nat) (n::nat) p::nat. p ~= 0 & <= m n --> <= (DIV m p) (DIV n p)" +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: "ALL (m::nat) (n::nat) p::nat. - p ~= 0 & <= (m + p) n --> < (DIV m p) (DIV n p)" +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 LE_LDIV: "ALL (a::nat) (b::nat) n::nat. a ~= 0 & <= b (a * n) --> <= (DIV b a) n" - by (import hollight LE_LDIV) - -lemma LE_RDIV_EQ: "ALL (a::nat) (b::nat) n::nat. a ~= 0 --> <= n (DIV b a) = <= (a * n) b" - by (import hollight LE_RDIV_EQ) - -lemma LE_LDIV_EQ: "ALL (a::nat) (b::nat) n::nat. - a ~= 0 --> <= (DIV b a) n = < b (a * (n + NUMERAL_BIT1 0))" - by (import hollight LE_LDIV_EQ) - -lemma DIV_EQ_0: "ALL (m::nat) n::nat. n ~= 0 --> (DIV m n = 0) = < m n" +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: "ALL (m::nat) n::nat. n ~= 0 --> (MOD m n = 0) = (EX q::nat. m = q * n)" +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: "ALL n::nat. EVEN n = (MOD n (NUMERAL_BIT0 (NUMERAL_BIT1 0)) = 0)" +lemma EVEN_MOD: "even (n::nat) = (n mod (2::nat) = (0::nat))" by (import hollight EVEN_MOD) -lemma ODD_MOD: "ALL n::nat. ODD n = (MOD n (NUMERAL_BIT0 (NUMERAL_BIT1 0)) = NUMERAL_BIT1 0)" +lemma ODD_MOD: "odd (n::nat) = (n mod (2::nat) = (1::nat))" by (import hollight ODD_MOD) -lemma MOD_MULT_RMOD: "ALL (m::nat) (n::nat) p::nat. n ~= 0 --> MOD (m * MOD p n) n = MOD (m * p) n" +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: "ALL (x::nat) (xa::nat) xb::nat. - xa ~= 0 --> MOD (MOD x xa * xb) xa = MOD (x * xb) xa" +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: "ALL (x::nat) (xa::nat) xb::nat. - xa ~= 0 --> MOD (MOD x xa * MOD xb xa) xa = MOD (x * xb) xa" +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: "ALL (m::nat) (n::nat) p::nat. - n ~= 0 --> MOD (EXP (MOD m n) p) n = MOD (EXP m p) n" +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_MULT_ADD: "ALL (m::nat) (n::nat) p::nat. MOD (m * n + p) n = MOD p n" - by (import hollight MOD_MULT_ADD) - -lemma MOD_ADD_MOD: "ALL (a::nat) (b::nat) n::nat. - n ~= 0 --> MOD (MOD a n + MOD b n) n = MOD (a + b) n" +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: "ALL (a::nat) (b::nat) n::nat. - n ~= 0 --> - (MOD (a + b) n = MOD a n + MOD b n) = (DIV (a + b) n = DIV a n + DIV b n)" +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 DIV_REFL: "ALL n::nat. n ~= 0 --> DIV n n = NUMERAL_BIT1 0" - by (import hollight DIV_REFL) - -lemma MOD_LE: "ALL (m::nat) n::nat. n ~= 0 --> <= (MOD m n) m" +lemma MOD_LE: "(n::nat) ~= (0::nat) ==> (m::nat) mod n <= m" by (import hollight MOD_LE) -lemma DIV_MONO2: "ALL (m::nat) (n::nat) p::nat. p ~= 0 & <= p m --> <= (DIV n m) (DIV n p)" +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: "ALL (a::nat) (b::nat) (c::nat) d::nat. - b ~= 0 & < (b * c) ((a + NUMERAL_BIT1 0) * d) --> <= (DIV c d) (DIV a b)" +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) + NUMERAL_BIT1 0) * (d::nat)) & -< (a * d) ((c + NUMERAL_BIT1 0) * b) --> -DIV a b = DIV c d" +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 x::nat. (b = a + x --> P 0) & (a = b + x --> P x))" +(ALL d::nat. a = b + d | a < b & d = (0::nat) --> P d)" by (import hollight SUB_ELIM_THM) -lemma PRE_ELIM_THM: "(P::nat => bool) (Pred (n::nat)) = -(ALL m::nat. (n = 0 --> P 0) & (n = Suc m --> P m))" - by (import hollight PRE_ELIM_THM) - -lemma DIVMOD_ELIM_THM: "(P::nat => nat => bool) (DIV (m::nat) (n::nat)) (MOD m n) = -(n = 0 & P 0 0 | - n ~= 0 & (ALL (q::nat) r::nat. m = q * n + r & < r n --> P q r))" +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 eqeq :: "'q_9910 => 'q_9909 => ('q_9910 => 'q_9909 => bool) => bool" where - "eqeq == -%(u::'q_9910::type) (ua::'q_9909::type) - ub::'q_9910::type => 'q_9909::type => bool. ub u ua" - -lemma DEF__equal__equal_: "eqeq = -(%(u::'q_9910::type) (ua::'q_9909::type) - ub::'q_9910::type => 'q_9909::type => bool. ub u ua)" - by (import hollight DEF__equal__equal_) - -definition mod_nat :: "nat => nat => nat => bool" where - "mod_nat == -%(u::nat) (ua::nat) ub::nat. EX (q1::nat) q2::nat. ua + u * q1 = ub + u * q2" - -lemma DEF_mod_nat: "mod_nat = -(%(u::nat) (ua::nat) ub::nat. - EX (q1::nat) q2::nat. ua + u * q1 = ub + u * q2)" - by (import hollight DEF_mod_nat) - -definition minimal :: "(nat => bool) => nat" where - "minimal == %u::nat => bool. SOME n::nat. u n & (ALL m::nat. < m n --> ~ u m)" - -lemma DEF_minimal: "minimal = -(%u::nat => bool. SOME n::nat. u n & (ALL m::nat. < m n --> ~ u m))" +definition + minimal :: "(nat => bool) => nat" where + "minimal == %u. SOME n. u n & (ALL m bool. - Ex P = (P (minimal P) & (ALL x::nat. < x (minimal P) --> ~ P x))" +lemma MINIMAL: "Ex P = (P (minimal P) & (ALL x 'A => bool) => bool" where - "WF == -%u::'A::type => 'A::type => bool. - ALL P::'A::type => bool. - Ex P --> (EX x::'A::type. P x & (ALL y::'A::type. u y x --> ~ P y))" - -lemma DEF_WF: "WF = -(%u::'A::type => 'A::type => bool. - ALL P::'A::type => bool. - Ex P --> (EX x::'A::type. P x & (ALL y::'A::type. u y x --> ~ P y)))" - by (import hollight DEF_WF) - -lemma WF_EQ: "WF (u_353::'A::type => 'A::type => bool) = -(ALL P::'A::type => bool. - Ex P = (EX x::'A::type. P x & (ALL y::'A::type. u_353 y x --> ~ P y)))" +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: "WF (u_353::'A::type => 'A::type => bool) = -(ALL P::'A::type => bool. - (ALL x::'A::type. (ALL y::'A::type. u_353 y x --> P y) --> P x) --> - All P)" +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: "WF (u_353::'A::type => 'A::type => bool) = -(~ (EX s::nat => 'A::type. ALL n::nat. u_353 (s (Suc n)) (s n)))" +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: "WF (u_353::'A::type => 'A::type => bool) --> -(ALL H::('A::type => 'B::type) => 'A::type => 'B::type. - (ALL (f::'A::type => 'B::type) (g::'A::type => 'B::type) x::'A::type. - (ALL z::'A::type. u_353 z x --> f z = g z) --> H f x = H g x) --> - (ALL (f::'A::type => 'B::type) g::'A::type => 'B::type. - (ALL x::'A::type. f x = H f x) & (ALL x::'A::type. g x = H g x) --> - f = g))" +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: "(ALL H::('A::type => bool) => 'A::type => bool. - (ALL (f::'A::type => bool) (g::'A::type => bool) x::'A::type. - (ALL z::'A::type. - (u_353::'A::type => 'A::type => bool) z x --> f z = g z) --> - H f x = H g x) --> - (ALL (f::'A::type => bool) g::'A::type => bool. - (ALL x::'A::type. f x = H f x) & (ALL x::'A::type. g x = H g x) --> - f = g)) --> -WF u_353" +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: "WF (u_353::'A::type => 'A::type => bool) --> -(ALL (H::('A::type => 'B::type) => 'A::type => 'B::type) - S::'A::type => 'B::type => bool. - (ALL (f::'A::type => 'B::type) (g::'A::type => 'B::type) x::'A::type. - (ALL z::'A::type. u_353 z x --> f z = g z & S z (f z)) --> - H f x = H g x & S x (H f x)) --> - (EX f::'A::type => 'B::type. ALL x::'A::type. f x = H f x))" +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: "WF (u_353::'A::type => 'A::type => bool) --> -(ALL H::('A::type => 'B::type) => 'A::type => 'B::type. - (ALL (f::'A::type => 'B::type) (g::'A::type => 'B::type) x::'A::type. - (ALL z::'A::type. u_353 z x --> f z = g z) --> H f x = H g x) --> - (EX f::'A::type => 'B::type. ALL x::'A::type. f x = H f x))" +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: "(ALL H::('A::type => nat) => 'A::type => nat. - (ALL (f::'A::type => nat) (g::'A::type => nat) x::'A::type. - (ALL z::'A::type. - (u_353::'A::type => 'A::type => bool) z x --> f z = g z) --> - H f x = H g x) --> - (EX f::'A::type => nat. ALL x::'A::type. f x = H f x)) --> -WF u_353" +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: "WF (u_353::'A::type => 'A::type => bool) --> -(ALL H::('A::type => 'B::type) => 'A::type => 'B::type. - (ALL (f::'A::type => 'B::type) (g::'A::type => 'B::type) x::'A::type. - (ALL z::'A::type. u_353 z x --> f z = g z) --> H f x = H g x) --> - (EX! f::'A::type => 'B::type. ALL x::'A::type. f x = H f x))" +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::type) y::'A::type. - (u_353::'A::type => 'A::type => bool) x y --> - (u_472::'A::type => 'A::type => bool) x y) & -WF u_472 --> -WF u_353" +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: "ALL m::'A::type => 'B::type. - WF (u_353::'B::type => 'B::type => bool) --> - WF (%(x::'A::type) x'::'A::type. u_353 (m x) (m x'))" +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: "ALL (R::'A::type => 'A::type => bool) - S::'A::type => 'B::type => 'B::type => bool. - WF R & (ALL x::'A::type. WF (S x)) --> - WF (GABS - (%f::'A::type * 'B::type => 'A::type * 'B::type => bool. - ALL (r1::'A::type) s1::'B::type. - GEQ (f (r1, s1)) - (GABS - (%f::'A::type * 'B::type => bool. - ALL (r2::'A::type) s2::'B::type. - GEQ (f (r2, s2)) - (R r1 r2 | r1 = r2 & S r1 s1 s2)))))" +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: "ALL (x::'A::type => 'A::type => bool) xa::'B::type => 'B::type => bool. - WF x & WF xa --> - WF (GABS - (%f::'A::type * 'B::type => 'A::type * 'B::type => bool. - ALL (r1::'A::type) s1::'B::type. - GEQ (f (r1, s1)) - (GABS - (%f::'A::type * 'B::type => bool. - ALL (r2::'A::type) s2::'B::type. - GEQ (f (r2, s2)) (x r1 r2 | r1 = r2 & xa s1 s2)))))" +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: "WF (u_353::'A::type => 'A::type => bool) & -WF (u_472::'B::type => 'B::type => bool) --> -WF (GABS - (%f::'A::type * 'B::type => 'A::type * 'B::type => bool. - ALL (x1::'A::type) y1::'B::type. - GEQ (f (x1, y1)) - (GABS - (%f::'A::type * 'B::type => bool. - ALL (x2::'A::type) y2::'B::type. - GEQ (f (x2, y2)) (u_353 x1 x2 & u_472 y1 y2)))))" +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: "WF <" +lemma WF_num: "(wfP::(nat => nat => bool) => bool) (op <::nat => nat => bool)" by (import hollight WF_num) -lemma WF_REC_num: "ALL H::(nat => 'A::type) => nat => 'A::type. - (ALL (f::nat => 'A::type) (g::nat => 'A::type) x::nat. - (ALL z::nat. < z x --> f z = g z) --> H f x = H g x) --> - (EX f::nat => 'A::type. ALL x::nat. f x = H f x)" +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) -consts - measure :: "('q_11107 => nat) => 'q_11107 => 'q_11107 => bool" - -defs - measure_def: "hollight.measure == -%(u::'q_11107::type => nat) (x::'q_11107::type) y::'q_11107::type. - < (u x) (u y)" - -lemma DEF_measure: "hollight.measure = -(%(u::'q_11107::type => nat) (x::'q_11107::type) y::'q_11107::type. - < (u x) (u y))" - by (import hollight DEF_measure) - -lemma WF_MEASURE: "ALL m::'A::type => nat. WF (hollight.measure m)" +lemma WF_MEASURE: "wfP (%(a::'A) b::'A. measure (m::'A => nat) (a, b))" by (import hollight WF_MEASURE) -lemma MEASURE_LE: "(ALL x::'q_11137::type. - hollight.measure (m::'q_11137::type => nat) x (a::'q_11137::type) --> - hollight.measure m x (b::'q_11137::type)) = -<= (m a) (m b)" +lemma MEASURE_LE: "(ALL x::'q_12099. + measure (m::'q_12099 => nat) (x, a::'q_12099) --> + measure m (x, b::'q_12099)) = +(m a <= m b)" by (import hollight MEASURE_LE) -lemma WF_REFL: "ALL x::'A::type. WF (u_353::'A::type => 'A::type => bool) --> ~ u_353 x x" +lemma WF_REFL: "wfP (u_556::'A => 'A => bool) ==> ~ u_556 (x::'A) x" by (import hollight WF_REFL) -lemma WF_FALSE: "WF (%(x::'A::type) y::'A::type. False)" - by (import hollight WF_FALSE) - -lemma WF_REC_TAIL: "ALL (P::'A::type => bool) (g::'A::type => 'A::type) h::'A::type => 'B::type. - EX f::'A::type => 'B::type. - ALL x::'A::type. f x = COND (P x) (f (g x)) (h x)" +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: "ALL (P::('A::type => 'B::type) => 'A::type => bool) - (G::('A::type => 'B::type) => 'A::type => 'A::type) - H::('A::type => 'B::type) => 'A::type => 'B::type. - WF (u_353::'A::type => 'A::type => bool) & - (ALL (f::'A::type => 'B::type) (g::'A::type => 'B::type) x::'A::type. - (ALL z::'A::type. u_353 z x --> f z = g z) --> - P f x = P g x & G f x = G g x & H f x = H g x) & - (ALL (f::'A::type => 'B::type) (g::'A::type => 'B::type) x::'A::type. - (ALL z::'A::type. u_353 z x --> f z = g z) --> H f x = H g x) & - (ALL (f::'A::type => 'B::type) (x::'A::type) y::'A::type. - P f x & u_353 y (G f x) --> u_353 y x) --> - (EX f::'A::type => 'B::type. - ALL x::'A::type. f x = COND (P f x) (f (G f x)) (H f x))" +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: "(op &::bool => bool => bool) ((op =::nat => nat => bool) (0::nat) (0::nat)) - ((op =::nat => nat => bool) ((NUMERAL_BIT0::nat => nat) (0::nat)) (0::nat))" +lemma ARITH_ZERO: "(0::nat) = (0::nat) & (0::nat) = (0::nat)" by (import hollight ARITH_ZERO) -lemma ARITH_SUC: "(ALL x::nat. Suc x = Suc x) & -Suc 0 = NUMERAL_BIT1 0 & -(ALL x::nat. Suc (NUMERAL_BIT0 x) = NUMERAL_BIT1 x) & -(ALL x::nat. Suc (NUMERAL_BIT1 x) = NUMERAL_BIT0 (Suc x))" +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::nat. Pred x = Pred x) & -Pred 0 = 0 & -(ALL x::nat. - Pred (NUMERAL_BIT0 x) = COND (x = 0) 0 (NUMERAL_BIT1 (Pred x))) & -(ALL x::nat. Pred (NUMERAL_BIT1 x) = NUMERAL_BIT0 x)" +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: "(op &::bool => bool => bool) - ((All::(nat => bool) => bool) - (%x::nat. - (All::(nat => bool) => bool) - (%xa::nat. - (op =::nat => nat => bool) ((plus::nat => nat => nat) x xa) - ((plus::nat => nat => nat) x xa)))) - ((op &::bool => bool => bool) - ((op =::nat => nat => bool) ((plus::nat => nat => nat) (0::nat) (0::nat)) - (0::nat)) - ((op &::bool => bool => bool) - ((All::(nat => bool) => bool) - (%x::nat. - (op =::nat => nat => bool) - ((plus::nat => nat => nat) (0::nat) - ((NUMERAL_BIT0::nat => nat) x)) - ((NUMERAL_BIT0::nat => nat) x))) - ((op &::bool => bool => bool) - ((All::(nat => bool) => bool) - (%x::nat. - (op =::nat => nat => bool) - ((plus::nat => nat => nat) (0::nat) - ((NUMERAL_BIT1::nat => nat) x)) - ((NUMERAL_BIT1::nat => nat) x))) - ((op &::bool => bool => bool) - ((All::(nat => bool) => bool) - (%x::nat. - (op =::nat => nat => bool) - ((plus::nat => nat => nat) ((NUMERAL_BIT0::nat => nat) x) - (0::nat)) - ((NUMERAL_BIT0::nat => nat) x))) - ((op &::bool => bool => bool) - ((All::(nat => bool) => bool) - (%x::nat. - (op =::nat => nat => bool) - ((plus::nat => nat => nat) ((NUMERAL_BIT1::nat => nat) x) - (0::nat)) - ((NUMERAL_BIT1::nat => nat) x))) - ((op &::bool => bool => bool) - ((All::(nat => bool) => bool) - (%x::nat. - (All::(nat => bool) => bool) - (%xa::nat. - (op =::nat => nat => bool) - ((plus::nat => nat => nat) - ((NUMERAL_BIT0::nat => nat) x) - ((NUMERAL_BIT0::nat => nat) xa)) - ((NUMERAL_BIT0::nat => nat) - ((plus::nat => nat => nat) x xa))))) - ((op &::bool => bool => bool) - ((All::(nat => bool) => bool) - (%x::nat. - (All::(nat => bool) => bool) - (%xa::nat. - (op =::nat => nat => bool) - ((plus::nat => nat => nat) - ((NUMERAL_BIT0::nat => nat) x) - ((NUMERAL_BIT1::nat => nat) xa)) - ((NUMERAL_BIT1::nat => nat) - ((plus::nat => nat => nat) x xa))))) - ((op &::bool => bool => bool) - ((All::(nat => bool) => bool) - (%x::nat. - (All::(nat => bool) => bool) - (%xa::nat. - (op =::nat => nat => bool) - ((plus::nat => nat => nat) - ((NUMERAL_BIT1::nat => nat) x) - ((NUMERAL_BIT0::nat => nat) xa)) - ((NUMERAL_BIT1::nat => nat) - ((plus::nat => nat => nat) x xa))))) - ((All::(nat => bool) => bool) - (%x::nat. - (All::(nat => bool) => bool) - (%xa::nat. - (op =::nat => nat => bool) - ((plus::nat => nat => nat) - ((NUMERAL_BIT1::nat => nat) x) - ((NUMERAL_BIT1::nat => nat) xa)) - ((NUMERAL_BIT0::nat => nat) - ((Suc::nat => nat) - ((plus::nat => nat => nat) x - xa))))))))))))))" +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: "(op &::bool => bool => bool) - ((All::(nat => bool) => bool) - (%x::nat. - (All::(nat => bool) => bool) - (%xa::nat. - (op =::nat => nat => bool) ((op *::nat => nat => nat) x xa) - ((op *::nat => nat => nat) x xa)))) - ((op &::bool => bool => bool) - ((op =::nat => nat => bool) ((op *::nat => nat => nat) (0::nat) (0::nat)) - (0::nat)) - ((op &::bool => bool => bool) - ((All::(nat => bool) => bool) - (%x::nat. - (op =::nat => nat => bool) - ((op *::nat => nat => nat) (0::nat) - ((NUMERAL_BIT0::nat => nat) x)) - (0::nat))) - ((op &::bool => bool => bool) - ((All::(nat => bool) => bool) - (%x::nat. - (op =::nat => nat => bool) - ((op *::nat => nat => nat) (0::nat) - ((NUMERAL_BIT1::nat => nat) x)) - (0::nat))) - ((op &::bool => bool => bool) - ((All::(nat => bool) => bool) - (%x::nat. - (op =::nat => nat => bool) - ((op *::nat => nat => nat) ((NUMERAL_BIT0::nat => nat) x) - (0::nat)) - (0::nat))) - ((op &::bool => bool => bool) - ((All::(nat => bool) => bool) - (%x::nat. - (op =::nat => nat => bool) - ((op *::nat => nat => nat) ((NUMERAL_BIT1::nat => nat) x) - (0::nat)) - (0::nat))) - ((op &::bool => bool => bool) - ((All::(nat => bool) => bool) - (%x::nat. - (All::(nat => bool) => bool) - (%xa::nat. - (op =::nat => nat => bool) - ((op *::nat => nat => nat) - ((NUMERAL_BIT0::nat => nat) x) - ((NUMERAL_BIT0::nat => nat) xa)) - ((NUMERAL_BIT0::nat => nat) - ((NUMERAL_BIT0::nat => nat) - ((op *::nat => nat => nat) x xa)))))) - ((op &::bool => bool => bool) - ((All::(nat => bool) => bool) - (%x::nat. - (All::(nat => bool) => bool) - (%xa::nat. - (op =::nat => nat => bool) - ((op *::nat => nat => nat) - ((NUMERAL_BIT0::nat => nat) x) - ((NUMERAL_BIT1::nat => nat) xa)) - ((plus::nat => nat => nat) - ((NUMERAL_BIT0::nat => nat) x) - ((NUMERAL_BIT0::nat => nat) - ((NUMERAL_BIT0::nat => nat) - ((op *::nat => nat => nat) x xa))))))) - ((op &::bool => bool => bool) - ((All::(nat => bool) => bool) - (%x::nat. - (All::(nat => bool) => bool) - (%xa::nat. - (op =::nat => nat => bool) - ((op *::nat => nat => nat) - ((NUMERAL_BIT1::nat => nat) x) - ((NUMERAL_BIT0::nat => nat) xa)) - ((plus::nat => nat => nat) - ((NUMERAL_BIT0::nat => nat) xa) - ((NUMERAL_BIT0::nat => nat) - ((NUMERAL_BIT0::nat => nat) - ((op *::nat => nat => nat) x xa))))))) - ((All::(nat => bool) => bool) - (%x::nat. - (All::(nat => bool) => bool) - (%xa::nat. - (op =::nat => nat => bool) - ((op *::nat => nat => nat) - ((NUMERAL_BIT1::nat => nat) x) - ((NUMERAL_BIT1::nat => nat) xa)) - ((plus::nat => nat => nat) - ((NUMERAL_BIT1::nat => nat) x) - ((plus::nat => nat => nat) - ((NUMERAL_BIT0::nat => nat) xa) - ((NUMERAL_BIT0::nat => nat) - ((NUMERAL_BIT0::nat => nat) - ((op *::nat => nat => nat) x - xa))))))))))))))))" +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. EXP x xa = EXP x xa) & -EXP 0 0 = NUMERAL_BIT1 0 & -(ALL m::nat. EXP (NUMERAL_BIT0 m) 0 = NUMERAL_BIT1 0) & -(ALL m::nat. EXP (NUMERAL_BIT1 m) 0 = NUMERAL_BIT1 0) & -(ALL n::nat. EXP 0 (NUMERAL_BIT0 n) = EXP 0 n * EXP 0 n) & +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. - EXP (NUMERAL_BIT0 m) (NUMERAL_BIT0 n) = - EXP (NUMERAL_BIT0 m) n * EXP (NUMERAL_BIT0 m) n) & + ((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. - EXP (NUMERAL_BIT1 m) (NUMERAL_BIT0 n) = - EXP (NUMERAL_BIT1 m) n * EXP (NUMERAL_BIT1 m) n) & -(ALL n::nat. EXP 0 (NUMERAL_BIT1 n) = 0) & -(ALL (m::nat) n::nat. - EXP (NUMERAL_BIT0 m) (NUMERAL_BIT1 n) = - NUMERAL_BIT0 m * (EXP (NUMERAL_BIT0 m) n * EXP (NUMERAL_BIT0 m) n)) & + ((2::nat) * m) ^ ((2::nat) * n + (1::nat)) = + (2::nat) * m * (((2::nat) * m) ^ n * ((2::nat) * m) ^ n)) & (ALL (m::nat) n::nat. - EXP (NUMERAL_BIT1 m) (NUMERAL_BIT1 n) = - NUMERAL_BIT1 m * (EXP (NUMERAL_BIT1 m) n * EXP (NUMERAL_BIT1 m) n))" + ((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 = True & -(ALL x::nat. EVEN (NUMERAL_BIT0 x) = True) & -(ALL x::nat. EVEN (NUMERAL_BIT1 x) = False)" +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 = False & -(ALL x::nat. ODD (NUMERAL_BIT0 x) = False) & -(ALL x::nat. ODD (NUMERAL_BIT1 x) = True)" +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 0 = True & -(ALL x::nat. <= (NUMERAL_BIT0 x) 0 = (x = 0)) & -(ALL x::nat. <= (NUMERAL_BIT1 x) 0 = False) & -(ALL x::nat. <= 0 (NUMERAL_BIT0 x) = True) & -(ALL x::nat. <= 0 (NUMERAL_BIT1 x) = True) & -(ALL (x::nat) xa::nat. <= (NUMERAL_BIT0 x) (NUMERAL_BIT0 xa) = <= x xa) & -(ALL (x::nat) xa::nat. <= (NUMERAL_BIT0 x) (NUMERAL_BIT1 xa) = <= x xa) & -(ALL (x::nat) xa::nat. <= (NUMERAL_BIT1 x) (NUMERAL_BIT0 xa) = < x xa) & -(ALL (x::nat) xa::nat. <= (NUMERAL_BIT1 x) (NUMERAL_BIT1 xa) = <= x xa)" +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 0 = False & -(ALL x::nat. < (NUMERAL_BIT0 x) 0 = False) & -(ALL x::nat. < (NUMERAL_BIT1 x) 0 = False) & -(ALL x::nat. < 0 (NUMERAL_BIT0 x) = < 0 x) & -(ALL x::nat. < 0 (NUMERAL_BIT1 x) = True) & -(ALL (x::nat) xa::nat. < (NUMERAL_BIT0 x) (NUMERAL_BIT0 xa) = < x xa) & -(ALL (x::nat) xa::nat. < (NUMERAL_BIT0 x) (NUMERAL_BIT1 xa) = <= x xa) & -(ALL (x::nat) xa::nat. < (NUMERAL_BIT1 x) (NUMERAL_BIT0 xa) = < x xa) & -(ALL (x::nat) xa::nat. < (NUMERAL_BIT1 x) (NUMERAL_BIT1 xa) = < x xa)" +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: "(op &::bool => bool => bool) - ((All::(nat => bool) => bool) - (%x::nat. - (All::(nat => bool) => bool) - (%xa::nat. - (op =::bool => bool => bool) ((op =::nat => nat => bool) x xa) - ((op =::nat => nat => bool) x xa)))) - ((op &::bool => bool => bool) - ((op =::bool => bool => bool) - ((op =::nat => nat => bool) (0::nat) (0::nat)) (True::bool)) - ((op &::bool => bool => bool) - ((All::(nat => bool) => bool) - (%x::nat. - (op =::bool => bool => bool) - ((op =::nat => nat => bool) ((NUMERAL_BIT0::nat => nat) x) - (0::nat)) - ((op =::nat => nat => bool) x (0::nat)))) - ((op &::bool => bool => bool) - ((All::(nat => bool) => bool) - (%x::nat. - (op =::bool => bool => bool) - ((op =::nat => nat => bool) ((NUMERAL_BIT1::nat => nat) x) - (0::nat)) - (False::bool))) - ((op &::bool => bool => bool) - ((All::(nat => bool) => bool) - (%x::nat. - (op =::bool => bool => bool) - ((op =::nat => nat => bool) (0::nat) - ((NUMERAL_BIT0::nat => nat) x)) - ((op =::nat => nat => bool) (0::nat) x))) - ((op &::bool => bool => bool) - ((All::(nat => bool) => bool) - (%x::nat. - (op =::bool => bool => bool) - ((op =::nat => nat => bool) (0::nat) - ((NUMERAL_BIT1::nat => nat) x)) - (False::bool))) - ((op &::bool => bool => bool) - ((All::(nat => bool) => bool) - (%x::nat. - (All::(nat => bool) => bool) - (%xa::nat. - (op =::bool => bool => bool) - ((op =::nat => nat => bool) - ((NUMERAL_BIT0::nat => nat) x) - ((NUMERAL_BIT0::nat => nat) xa)) - ((op =::nat => nat => bool) x xa)))) - ((op &::bool => bool => bool) - ((All::(nat => bool) => bool) - (%x::nat. - (All::(nat => bool) => bool) - (%xa::nat. - (op =::bool => bool => bool) - ((op =::nat => nat => bool) - ((NUMERAL_BIT0::nat => nat) x) - ((NUMERAL_BIT1::nat => nat) xa)) - (False::bool)))) - ((op &::bool => bool => bool) - ((All::(nat => bool) => bool) - (%x::nat. - (All::(nat => bool) => bool) - (%xa::nat. - (op =::bool => bool => bool) - ((op =::nat => nat => bool) - ((NUMERAL_BIT1::nat => nat) x) - ((NUMERAL_BIT0::nat => nat) xa)) - (False::bool)))) - ((All::(nat => bool) => bool) - (%x::nat. - (All::(nat => bool) => bool) - (%xa::nat. - (op =::bool => bool => bool) - ((op =::nat => nat => bool) - ((NUMERAL_BIT1::nat => nat) x) - ((NUMERAL_BIT1::nat => nat) xa)) - ((op =::nat => nat => bool) x xa))))))))))))" +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: "(op &::bool => bool => bool) - ((All::(nat => bool) => bool) - (%x::nat. - (All::(nat => bool) => bool) - (%xa::nat. - (op =::nat => nat => bool) ((op -::nat => nat => nat) x xa) - ((op -::nat => nat => nat) x xa)))) - ((op &::bool => bool => bool) - ((op =::nat => nat => bool) ((op -::nat => nat => nat) (0::nat) (0::nat)) - (0::nat)) - ((op &::bool => bool => bool) - ((All::(nat => bool) => bool) - (%x::nat. - (op =::nat => nat => bool) - ((op -::nat => nat => nat) (0::nat) - ((NUMERAL_BIT0::nat => nat) x)) - (0::nat))) - ((op &::bool => bool => bool) - ((All::(nat => bool) => bool) - (%x::nat. - (op =::nat => nat => bool) - ((op -::nat => nat => nat) (0::nat) - ((NUMERAL_BIT1::nat => nat) x)) - (0::nat))) - ((op &::bool => bool => bool) - ((All::(nat => bool) => bool) - (%x::nat. - (op =::nat => nat => bool) - ((op -::nat => nat => nat) ((NUMERAL_BIT0::nat => nat) x) - (0::nat)) - ((NUMERAL_BIT0::nat => nat) x))) - ((op &::bool => bool => bool) - ((All::(nat => bool) => bool) - (%x::nat. - (op =::nat => nat => bool) - ((op -::nat => nat => nat) ((NUMERAL_BIT1::nat => nat) x) - (0::nat)) - ((NUMERAL_BIT1::nat => nat) x))) - ((op &::bool => bool => bool) - ((All::(nat => bool) => bool) - (%m::nat. - (All::(nat => bool) => bool) - (%n::nat. - (op =::nat => nat => bool) - ((op -::nat => nat => nat) - ((NUMERAL_BIT0::nat => nat) m) - ((NUMERAL_BIT0::nat => nat) n)) - ((NUMERAL_BIT0::nat => nat) - ((op -::nat => nat => nat) m n))))) - ((op &::bool => bool => bool) - ((All::(nat => bool) => bool) - (%m::nat. - (All::(nat => bool) => bool) - (%n::nat. - (op =::nat => nat => bool) - ((op -::nat => nat => nat) - ((NUMERAL_BIT0::nat => nat) m) - ((NUMERAL_BIT1::nat => nat) n)) - ((Pred::nat => nat) - ((NUMERAL_BIT0::nat => nat) - ((op -::nat => nat => nat) m n)))))) - ((op &::bool => bool => bool) - ((All::(nat => bool) => bool) - (%m::nat. - (All::(nat => bool) => bool) - (%n::nat. - (op =::nat => nat => bool) - ((op -::nat => nat => nat) - ((NUMERAL_BIT1::nat => nat) m) - ((NUMERAL_BIT0::nat => nat) n)) - ((COND::bool => nat => nat => nat) - ((<=::nat => nat => bool) n m) - ((NUMERAL_BIT1::nat => nat) - ((op -::nat => nat => nat) m n)) - (0::nat))))) - ((All::(nat => bool) => bool) - (%m::nat. - (All::(nat => bool) => bool) - (%n::nat. - (op =::nat => nat => bool) - ((op -::nat => nat => nat) - ((NUMERAL_BIT1::nat => nat) m) - ((NUMERAL_BIT1::nat => nat) n)) - ((NUMERAL_BIT0::nat => nat) - ((op -::nat => nat => nat) m n)))))))))))))" +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) * NUMERAL_BIT1 (x::nat) = s + NUMERAL_BIT0 (s * x)" +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::type) (y::'A::type) z::'A::type. - (add::'A::type => 'A::type => 'A::type) x (add y z) = add (add x y) z) & -(ALL (x::'A::type) y::'A::type. add x y = add y x) & -(ALL x::'A::type. add (r0::'A::type) x = x) & -(ALL (x::'A::type) (y::'A::type) z::'A::type. - (mul::'A::type => 'A::type => 'A::type) x (mul y z) = mul (mul x y) z) & -(ALL (x::'A::type) y::'A::type. mul x y = mul y x) & -(ALL x::'A::type. mul (r1::'A::type) x = x) & -(ALL x::'A::type. mul r0 x = r0) & -(ALL (x::'A::type) (y::'A::type) z::'A::type. - mul x (add y z) = add (mul x y) (mul x z)) & -(ALL x::'A::type. (pwr::'A::type => nat => 'A::type) x 0 = r1) & -(ALL (x::'A::type) n::nat. pwr x (Suc n) = mul x (pwr x n)) --> -mul r1 (x::'A::type) = x & -add (mul (a::'A::type) (m::'A::type)) (mul (b::'A::type) 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::type) = 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::type) (ly::'A::type)) - (mul (rx::'A::type) (ry::'A::type)) = -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::type)) = 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 (NUMERAL_BIT0 (NUMERAL_BIT1 0)) & -pwr (mul x (y::'A::type)) q = mul (pwr x q) (pwr y q) & -pwr (pwr x p) q = pwr x (p * q) & -pwr x 0 = r1 & -pwr x (NUMERAL_BIT1 0) = x & -mul x (add y (z::'A::type)) = add (mul x y) (mul x z) & -pwr x (Suc q) = mul x (pwr x q)" +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 sth: "(ALL (x::nat) (y::nat) z::nat. x + (y + z) = x + y + z) & -(ALL (x::nat) y::nat. x + y = y + x) & -(ALL x::nat. 0 + x = x) & -(ALL (x::nat) (y::nat) z::nat. x * (y * z) = x * y * z) & -(ALL (x::nat) y::nat. x * y = y * x) & -(ALL x::nat. NUMERAL_BIT1 0 * x = x) & -(ALL x::nat. 0 * x = 0) & -(ALL (x::nat) (xa::nat) xb::nat. x * (xa + xb) = x * xa + x * xb) & -(ALL x::nat. EXP x 0 = NUMERAL_BIT1 0) & -(ALL (x::nat) xa::nat. EXP x (Suc xa) = x * EXP x xa)" - by (import hollight sth) - -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)" +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 * x = 0) & +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: "ALL P::'A::type => 'B::type => 'C::type. - (ALL (x1::'A::type) (y1::'B::type) (x2::'A::type) y2::'B::type. - (P x1 y1 = P x2 y2) = (x1 = x2 & y1 = y2)) --> - (EX (x::'C::type => 'A::type) Y::'C::type => 'B::type. - ALL (xa::'A::type) y::'B::type. x (P xa y) = xa & Y (P xa y) = y)" +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::nat) ua::nat. - EXP (NUMERAL_BIT0 (NUMERAL_BIT1 0)) u * - (NUMERAL_BIT0 (NUMERAL_BIT1 0) * ua + NUMERAL_BIT1 0)" - -lemma DEF_NUMPAIR: "NUMPAIR = -(%(u::nat) ua::nat. - EXP (NUMERAL_BIT0 (NUMERAL_BIT1 0)) u * - (NUMERAL_BIT0 (NUMERAL_BIT1 0) * ua + NUMERAL_BIT1 0))" +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: "ALL (x::nat) (xa::nat) (xb::nat) xc::nat. - NUMPAIR x xa = NUMPAIR xb xc --> x = xb" +lemma NUMPAIR_INJ_LEMMA: "NUMPAIR x xa = NUMPAIR xb xc ==> x = xb" by (import hollight NUMPAIR_INJ_LEMMA) -lemma NUMPAIR_INJ: "ALL (x1::nat) (y1::nat) (x2::nat) y2::nat. - (NUMPAIR x1 y1 = NUMPAIR x2 y2) = (x1 = x2 & y1 = y2)" +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::nat => nat. - EX Y::nat => nat. - ALL (x::nat) y::nat. X (NUMPAIR x y) = x & Y (NUMPAIR x y) = y" - -lemma DEF_NUMFST: "NUMFST = -(SOME X::nat => nat. - EX Y::nat => nat. - ALL (x::nat) y::nat. X (NUMPAIR x y) = x & Y (NUMPAIR x y) = y)" +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::nat => nat. - ALL (x::nat) y::nat. NUMFST (NUMPAIR x y) = x & Y (NUMPAIR x y) = y" - -lemma DEF_NUMSND: "NUMSND = -(SOME Y::nat => nat. - ALL (x::nat) y::nat. NUMFST (NUMPAIR x y) = x & Y (NUMPAIR x y) = y)" +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::bool) ua::nat. - COND u (Suc (NUMERAL_BIT0 (NUMERAL_BIT1 0) * ua)) - (NUMERAL_BIT0 (NUMERAL_BIT1 0) * ua)" - -lemma DEF_NUMSUM: "NUMSUM = -(%(u::bool) ua::nat. - COND u (Suc (NUMERAL_BIT0 (NUMERAL_BIT1 0) * ua)) - (NUMERAL_BIT0 (NUMERAL_BIT1 0) * ua))" +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: "ALL (b1::bool) (x1::nat) (b2::bool) x2::nat. - (NUMSUM b1 x1 = NUMSUM b2 x2) = (b1 = b2 & x1 = x2)" +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::nat => bool. - EX Y::nat => nat. - ALL (x::bool) y::nat. X (NUMSUM x y) = x & Y (NUMSUM x y) = y" - -lemma DEF_NUMLEFT: "NUMLEFT = -(SOME X::nat => bool. - EX Y::nat => nat. - ALL (x::bool) y::nat. X (NUMSUM x y) = x & Y (NUMSUM x y) = y)" +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::nat => nat. - ALL (x::bool) y::nat. NUMLEFT (NUMSUM x y) = x & Y (NUMSUM x y) = y" - -lemma DEF_NUMRIGHT: "NUMRIGHT = -(SOME Y::nat => nat. - ALL (x::bool) y::nat. NUMLEFT (NUMSUM x y) = x & Y (NUMSUM x y) = y)" +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::type. n = u" - -lemma DEF_INJN: "INJN = (%(u::nat) (n::nat) a::'A::type. n = u)" +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: "(All::(nat => bool) => bool) - (%n1::nat. - (All::(nat => bool) => bool) - (%n2::nat. - (op =::bool => bool => bool) - ((op =::(nat => 'A::type => bool) - => (nat => 'A::type => bool) => bool) - ((INJN::nat => nat => 'A::type => bool) n1) - ((INJN::nat => nat => 'A::type => bool) n2)) - ((op =::nat => nat => bool) n1 n2)))" +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::type) (n::nat) b::'A::type. b = u" - -lemma DEF_INJA: "INJA = (%(u::'A::type) (n::nat) b::'A::type. b = u)" +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: "ALL (a1::'A::type) a2::'A::type. (INJA a1 = INJA a2) = (a1 = a2)" +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::type => bool) n::nat. u (NUMFST n) (NUMSND n)" - -lemma DEF_INJF: "INJF = -(%(u::nat => nat => 'A::type => bool) n::nat. u (NUMFST n) (NUMSND n))" +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: "ALL (f1::nat => nat => 'A::type => bool) f2::nat => nat => 'A::type => bool. - (INJF f1 = INJF f2) = (f1 = f2)" +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 +definition + INJP :: "(nat => 'A => bool) => (nat => 'A => bool) => nat => 'A => bool" where "INJP == -%(u::nat => 'A::type => bool) (ua::nat => 'A::type => bool) (n::nat) - a::'A::type. COND (NUMLEFT n) (u (NUMRIGHT n) a) (ua (NUMRIGHT n) a)" +%(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::type => bool) (ua::nat => 'A::type => bool) (n::nat) - a::'A::type. COND (NUMLEFT n) (u (NUMRIGHT n) a) (ua (NUMRIGHT n) a))" +(%(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: "ALL (f1::nat => 'A::type => bool) (f1'::nat => 'A::type => bool) - (f2::nat => 'A::type => bool) f2'::nat => 'A::type => bool. - (INJP f1 f2 = INJP f1' f2') = (f1 = f1' & f2 = f2')" +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 +definition + ZCONSTR :: "nat => 'A => (nat => nat => 'A => bool) => nat => 'A => bool" where "ZCONSTR == -%(u::nat) (ua::'A::type) ub::nat => nat => 'A::type => bool. +%(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::type) ub::nat => nat => 'A::type => bool. +(%(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) (SOME z::nat => 'A::type => bool. True)" - -lemma DEF_ZBOT: "ZBOT = INJP (INJN 0) (SOME z::nat => 'A::type => bool. True)" +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: "ALL (x::nat) (xa::'A::type) xb::nat => nat => 'A::type => bool. - ZCONSTR x xa xb ~= 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 +definition + ZRECSPACE :: "(nat => 'A => bool) => bool" where "ZRECSPACE == -%a::nat => 'A::type => bool. - ALL ZRECSPACE'::(nat => 'A::type => bool) => bool. - (ALL a::nat => 'A::type => bool. +%a::nat => 'A => bool. + ALL ZRECSPACE'::(nat => 'A => bool) => bool. + (ALL a::nat => 'A => bool. a = ZBOT | - (EX (c::nat) (i::'A::type) r::nat => nat => 'A::type => bool. + (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::type => bool. - ALL ZRECSPACE'::(nat => 'A::type => bool) => bool. - (ALL a::nat => 'A::type => bool. +(%a::nat => 'A => bool. + ALL ZRECSPACE'::(nat => 'A => bool) => bool. + (ALL a::nat => 'A => bool. a = ZBOT | - (EX (c::nat) (i::'A::type) r::nat => nat => 'A::type => bool. + (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::((nat => 'A::type => bool) => bool) - => (nat => 'A::type => bool) set) - (ZRECSPACE::(nat => 'A::type => bool) => bool)" morphisms "_dest_rec" "_mk_rec" - apply (rule light_ex_imp_nonempty[where t="ZBOT::nat => 'A::type => bool"]) +typedef (open) ('A) recspace = "Collect ZRECSPACE" morphisms "_dest_rec" "_mk_rec" + apply (rule light_ex_imp_nonempty[where t="ZBOT"]) by (import hollight TYDEF_recspace) syntax @@ -1779,1315 +1169,236 @@ [where a="a :: 'A recspace" and r=r , OF type_definition_recspace] -definition BOTTOM :: "'A recspace" where - "(op ==::'A::type recspace => 'A::type recspace => prop) - (BOTTOM::'A::type recspace) - ((_mk_rec::(nat => 'A::type => bool) => 'A::type recspace) - (ZBOT::nat => 'A::type => bool))" - -lemma DEF_BOTTOM: "(op =::'A::type recspace => 'A::type recspace => bool) - (BOTTOM::'A::type recspace) - ((_mk_rec::(nat => 'A::type => bool) => 'A::type recspace) - (ZBOT::nat => 'A::type => bool))" +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 - "(op ==::(nat => 'A::type => (nat => 'A::type recspace) => 'A::type recspace) - => (nat - => 'A::type => (nat => 'A::type recspace) => 'A::type recspace) - => prop) - (CONSTR::nat - => 'A::type => (nat => 'A::type recspace) => 'A::type recspace) - (%(u::nat) (ua::'A::type) ub::nat => 'A::type recspace. - (_mk_rec::(nat => 'A::type => bool) => 'A::type recspace) - ((ZCONSTR::nat - => 'A::type - => (nat => nat => 'A::type => bool) - => nat => 'A::type => bool) - u ua - (%n::nat. - (_dest_rec::'A::type recspace => nat => 'A::type => bool) - (ub n))))" - -lemma DEF_CONSTR: "(op =::(nat => 'A::type => (nat => 'A::type recspace) => 'A::type recspace) - => (nat - => 'A::type => (nat => 'A::type recspace) => 'A::type recspace) - => bool) - (CONSTR::nat - => 'A::type => (nat => 'A::type recspace) => 'A::type recspace) - (%(u::nat) (ua::'A::type) ub::nat => 'A::type recspace. - (_mk_rec::(nat => 'A::type => bool) => 'A::type recspace) - ((ZCONSTR::nat - => 'A::type - => (nat => nat => 'A::type => bool) - => nat => 'A::type => bool) - u ua - (%n::nat. - (_dest_rec::'A::type recspace => nat => 'A::type => bool) - (ub n))))" +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: "(All::((nat => 'A::type => bool) => bool) => bool) - (%x::nat => 'A::type => bool. - (All::((nat => 'A::type => bool) => bool) => bool) - (%y::nat => 'A::type => bool. - (op -->::bool => bool => bool) - ((op =::'A::type recspace => 'A::type recspace => bool) - ((_mk_rec::(nat => 'A::type => bool) => 'A::type recspace) x) - ((_mk_rec::(nat => 'A::type => bool) => 'A::type recspace) y)) - ((op -->::bool => bool => bool) - ((op &::bool => bool => bool) - ((ZRECSPACE::(nat => 'A::type => bool) => bool) x) - ((ZRECSPACE::(nat => 'A::type => bool) => bool) y)) - ((op =::(nat => 'A::type => bool) - => (nat => 'A::type => bool) => bool) - x y))))" +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: "ALL (c::nat) (i::'A::type) r::nat => 'A::type recspace. - CONSTR c i r ~= BOTTOM" +lemma CONSTR_BOT: "CONSTR (c::nat) (i::'A) (r::nat => 'A recspace) ~= BOTTOM" by (import hollight CONSTR_BOT) -lemma CONSTR_INJ: "ALL (c1::nat) (i1::'A::type) (r1::nat => 'A::type recspace) (c2::nat) - (i2::'A::type) r2::nat => 'A::type recspace. - (CONSTR c1 i1 r1 = CONSTR c2 i2 r2) = (c1 = c2 & i1 = i2 & r1 = r2)" +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: "ALL P::'A::type recspace => bool. - P BOTTOM & - (ALL (c::nat) (i::'A::type) r::nat => 'A::type recspace. - (ALL n::nat. P (r n)) --> P (CONSTR c i r)) --> - All P" +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: "ALL Fn::nat - => 'A::type - => (nat => 'A::type recspace) => (nat => 'B::type) => 'B::type. - EX f::'A::type recspace => 'B::type. - ALL (c::nat) (i::'A::type) r::nat => 'A::type recspace. - f (CONSTR c i r) = Fn c i r (%n::nat. f (r n))" +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 +definition + FCONS :: "'A => (nat => 'A) => nat => 'A" where "FCONS == -SOME FCONS::'A::type => (nat => 'A::type) => nat => 'A::type. - (ALL (a::'A::type) f::nat => 'A::type. FCONS a f 0 = a) & - (ALL (a::'A::type) (f::nat => 'A::type) n::nat. FCONS a f (Suc n) = f n)" +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::type => (nat => 'A::type) => nat => 'A::type. - (ALL (a::'A::type) f::nat => 'A::type. FCONS a f 0 = a) & - (ALL (a::'A::type) (f::nat => 'A::type) n::nat. - FCONS a f (Suc n) = f n))" +(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: "ALL f::nat => 'A::type. f = FCONS (f 0) (f o Suc)" +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::type. True" - -lemma DEF_FNIL: "FNIL = (%u::nat. SOME x::'A::type. True)" +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) -consts - OUTL :: "'A + 'B => 'A" - -defs - OUTL_def: "hollight.OUTL == -SOME OUTL::'A::type + 'B::type => 'A::type. - ALL x::'A::type. OUTL (Inl x) = x" - -lemma DEF_OUTL: "hollight.OUTL = -(SOME OUTL::'A::type + 'B::type => 'A::type. - ALL x::'A::type. OUTL (Inl x) = x)" - by (import hollight DEF_OUTL) - -consts - OUTR :: "'A + 'B => 'B" - -defs - OUTR_def: "hollight.OUTR == -SOME OUTR::'A::type + 'B::type => 'B::type. - ALL y::'B::type. OUTR (Inr y) = y" - -lemma DEF_OUTR: "hollight.OUTR = -(SOME OUTR::'A::type + 'B::type => 'B::type. - ALL y::'B::type. OUTR (Inr y) = y)" - by (import hollight DEF_OUTR) - -typedef (open) ('A) option = "(Collect::('A::type recspace => bool) => 'A::type recspace set) - (%a::'A::type recspace. - (All::(('A::type recspace => bool) => bool) => bool) - (%option'::'A::type recspace => bool. - (op -->::bool => bool => bool) - ((All::('A::type recspace => bool) => bool) - (%a::'A::type recspace. - (op -->::bool => bool => bool) - ((op |::bool => bool => bool) - ((op =::'A::type recspace => 'A::type recspace => bool) - a ((CONSTR::nat - => 'A::type - => (nat => 'A::type recspace) - => 'A::type recspace) - ((NUMERAL::nat => nat) (0::nat)) - ((Eps::('A::type => bool) => 'A::type) - (%v::'A::type. True::bool)) - (%n::nat. BOTTOM::'A::type recspace))) - ((Ex::('A::type => bool) => bool) - (%aa::'A::type. - (op =::'A::type recspace - => 'A::type recspace => bool) - a ((CONSTR::nat - => 'A::type => (nat => 'A::type recspace) => 'A::type recspace) - ((Suc::nat => nat) - ((NUMERAL::nat => nat) (0::nat))) - aa (%n::nat. BOTTOM::'A::type recspace))))) - (option' a))) - (option' a)))" morphisms "_dest_option" "_mk_option" - apply (rule light_ex_imp_nonempty[where t="(CONSTR::nat => 'A::type => (nat => 'A::type recspace) => 'A::type recspace) - ((NUMERAL::nat => nat) (0::nat)) - ((Eps::('A::type => bool) => 'A::type) (%v::'A::type. True::bool)) - (%n::nat. BOTTOM::'A::type recspace)"]) - by (import hollight TYDEF_option) +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, a1, a2, a3, a4, a5, a6, a7) + (%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_option" :: _ ("'_dest'_option") + "_dest_char" :: _ ("'_dest'_char") syntax - "_mk_option" :: _ ("'_mk'_option") - -lemmas "TYDEF_option_@intern" = typedef_hol2hollight - [where a="a :: 'A hollight.option" and r=r , - OF type_definition_option] - -definition NONE :: "'A hollight.option" where - "(op ==::'A::type hollight.option => 'A::type hollight.option => prop) - (NONE::'A::type hollight.option) - ((_mk_option::'A::type recspace => 'A::type hollight.option) - ((CONSTR::nat - => 'A::type => (nat => 'A::type recspace) => 'A::type recspace) - (0::nat) - ((Eps::('A::type => bool) => 'A::type) (%v::'A::type. True::bool)) - (%n::nat. BOTTOM::'A::type recspace)))" - -lemma DEF_NONE: "(op =::'A::type hollight.option => 'A::type hollight.option => bool) - (NONE::'A::type hollight.option) - ((_mk_option::'A::type recspace => 'A::type hollight.option) - ((CONSTR::nat - => 'A::type => (nat => 'A::type recspace) => 'A::type recspace) - (0::nat) - ((Eps::('A::type => bool) => 'A::type) (%v::'A::type. True::bool)) - (%n::nat. BOTTOM::'A::type recspace)))" - by (import hollight DEF_NONE) + "_mk_char" :: _ ("'_mk'_char") + +lemmas "TYDEF_char_@intern" = typedef_hol2hollight + [where a="a :: hollight.char" and r=r , + OF type_definition_char] consts - SOME :: "'A => 'A hollight.option" ("SOME") - -defs - SOME_def: "(op ==::('A::type => 'A::type hollight.option) - => ('A::type => 'A::type hollight.option) => prop) - (SOME::'A::type => 'A::type hollight.option) - (%a::'A::type. - (_mk_option::'A::type recspace => 'A::type hollight.option) - ((CONSTR::nat - => 'A::type - => (nat => 'A::type recspace) => 'A::type recspace) - ((Suc::nat => nat) (0::nat)) a - (%n::nat. BOTTOM::'A::type recspace)))" - -lemma DEF_SOME: "(op =::('A::type => 'A::type hollight.option) - => ('A::type => 'A::type hollight.option) => bool) - (SOME::'A::type => 'A::type hollight.option) - (%a::'A::type. - (_mk_option::'A::type recspace => 'A::type hollight.option) - ((CONSTR::nat - => 'A::type - => (nat => 'A::type recspace) => 'A::type recspace) - ((Suc::nat => nat) (0::nat)) a - (%n::nat. BOTTOM::'A::type recspace)))" - by (import hollight DEF_SOME) - -typedef (open) ('A) list = "(Collect::('A::type recspace => bool) => 'A::type recspace set) - (%a::'A::type recspace. - (All::(('A::type recspace => bool) => bool) => bool) - (%list'::'A::type recspace => bool. - (op -->::bool => bool => bool) - ((All::('A::type recspace => bool) => bool) - (%a::'A::type recspace. - (op -->::bool => bool => bool) - ((op |::bool => bool => bool) - ((op =::'A::type recspace => 'A::type recspace => bool) - a ((CONSTR::nat - => 'A::type - => (nat => 'A::type recspace) - => 'A::type recspace) - ((NUMERAL::nat => nat) (0::nat)) - ((Eps::('A::type => bool) => 'A::type) - (%v::'A::type. True::bool)) - (%n::nat. BOTTOM::'A::type recspace))) - ((Ex::('A::type => bool) => bool) - (%a0::'A::type. - (Ex::('A::type recspace => bool) => bool) - (%a1::'A::type recspace. - (op &::bool => bool => bool) - ((op =::'A::type recspace - => 'A::type recspace => bool) - a ((CONSTR::nat - => 'A::type => (nat => 'A::type recspace) => 'A::type recspace) -((Suc::nat => nat) ((NUMERAL::nat => nat) (0::nat))) a0 -((FCONS::'A::type recspace - => (nat => 'A::type recspace) => nat => 'A::type recspace) - a1 (%n::nat. BOTTOM::'A::type recspace)))) - (list' a1))))) - (list' a))) - (list' a)))" morphisms "_dest_list" "_mk_list" - apply (rule light_ex_imp_nonempty[where t="(CONSTR::nat => 'A::type => (nat => 'A::type recspace) => 'A::type recspace) - ((NUMERAL::nat => nat) (0::nat)) - ((Eps::('A::type => bool) => 'A::type) (%v::'A::type. True::bool)) - (%n::nat. BOTTOM::'A::type recspace)"]) - by (import hollight TYDEF_list) - -syntax - "_dest_list" :: _ ("'_dest'_list") - -syntax - "_mk_list" :: _ ("'_mk'_list") - -lemmas "TYDEF_list_@intern" = typedef_hol2hollight - [where a="a :: 'A hollight.list" and r=r , - OF type_definition_list] - -definition NIL :: "'A hollight.list" where - "(op ==::'A::type hollight.list => 'A::type hollight.list => prop) - (NIL::'A::type hollight.list) - ((_mk_list::'A::type recspace => 'A::type hollight.list) - ((CONSTR::nat - => 'A::type => (nat => 'A::type recspace) => 'A::type recspace) - (0::nat) - ((Eps::('A::type => bool) => 'A::type) (%v::'A::type. True::bool)) - (%n::nat. BOTTOM::'A::type recspace)))" - -lemma DEF_NIL: "(op =::'A::type hollight.list => 'A::type hollight.list => bool) - (NIL::'A::type hollight.list) - ((_mk_list::'A::type recspace => 'A::type hollight.list) - ((CONSTR::nat - => 'A::type => (nat => 'A::type recspace) => 'A::type recspace) - (0::nat) - ((Eps::('A::type => bool) => 'A::type) (%v::'A::type. True::bool)) - (%n::nat. BOTTOM::'A::type recspace)))" - by (import hollight DEF_NIL) - -definition CONS :: "'A => 'A hollight.list => 'A hollight.list" where - "(op ==::('A::type => 'A::type hollight.list => 'A::type hollight.list) - => ('A::type => 'A::type hollight.list => 'A::type hollight.list) - => prop) - (CONS::'A::type => 'A::type hollight.list => 'A::type hollight.list) - (%(a0::'A::type) a1::'A::type hollight.list. - (_mk_list::'A::type recspace => 'A::type hollight.list) - ((CONSTR::nat - => 'A::type - => (nat => 'A::type recspace) => 'A::type recspace) - ((Suc::nat => nat) (0::nat)) a0 - ((FCONS::'A::type recspace - => (nat => 'A::type recspace) => nat => 'A::type recspace) - ((_dest_list::'A::type hollight.list => 'A::type recspace) a1) - (%n::nat. BOTTOM::'A::type recspace))))" - -lemma DEF_CONS: "(op =::('A::type => 'A::type hollight.list => 'A::type hollight.list) - => ('A::type => 'A::type hollight.list => 'A::type hollight.list) - => bool) - (CONS::'A::type => 'A::type hollight.list => 'A::type hollight.list) - (%(a0::'A::type) a1::'A::type hollight.list. - (_mk_list::'A::type recspace => 'A::type hollight.list) - ((CONSTR::nat - => 'A::type - => (nat => 'A::type recspace) => 'A::type recspace) - ((Suc::nat => nat) (0::nat)) a0 - ((FCONS::'A::type recspace - => (nat => 'A::type recspace) => nat => 'A::type recspace) - ((_dest_list::'A::type hollight.list => 'A::type recspace) a1) - (%n::nat. BOTTOM::'A::type recspace))))" - by (import hollight DEF_CONS) - -lemma pair_RECURSION: "ALL PAIR'::'A::type => 'B::type => 'C::type. - EX x::'A::type * 'B::type => 'C::type. - ALL (a0::'A::type) a1::'B::type. x (a0, a1) = PAIR' a0 a1" - by (import hollight pair_RECURSION) - -lemma num_RECURSION_STD: "ALL (e::'Z::type) f::nat => 'Z::type => 'Z::type. - EX fn::nat => 'Z::type. fn 0 = e & (ALL n::nat. fn (Suc n) = f n (fn n))" - by (import hollight num_RECURSION_STD) - -lemma bool_RECURSION: "ALL (a::'A::type) b::'A::type. - EX x::bool => 'A::type. x False = a & x True = b" - by (import hollight bool_RECURSION) - -definition ISO :: "('A => 'B) => ('B => 'A) => bool" where - "ISO == -%(u::'A::type => 'B::type) ua::'B::type => 'A::type. - (ALL x::'B::type. u (ua x) = x) & (ALL y::'A::type. ua (u y) = y)" - -lemma DEF_ISO: "ISO = -(%(u::'A::type => 'B::type) ua::'B::type => 'A::type. - (ALL x::'B::type. u (ua x) = x) & (ALL y::'A::type. ua (u y) = y))" - by (import hollight DEF_ISO) - -lemma ISO_REFL: "ISO (%x::'A::type. x) (%x::'A::type. x)" - by (import hollight ISO_REFL) - -lemma ISO_FUN: "ISO (f::'A::type => 'A'::type) (f'::'A'::type => 'A::type) & -ISO (g::'B::type => 'B'::type) (g'::'B'::type => 'B::type) --> -ISO (%(h::'A::type => 'B::type) a'::'A'::type. g (h (f' a'))) - (%(h::'A'::type => 'B'::type) a::'A::type. g' (h (f a)))" - by (import hollight ISO_FUN) - -lemma ISO_USAGE: "ISO (f::'q_16636::type => 'q_16633::type) - (g::'q_16633::type => 'q_16636::type) --> -(ALL P::'q_16636::type => bool. All P = (ALL x::'q_16633::type. P (g x))) & -(ALL P::'q_16636::type => bool. Ex P = (EX x::'q_16633::type. P (g x))) & -(ALL (a::'q_16636::type) b::'q_16633::type. (a = g b) = (f a = b))" - by (import hollight ISO_USAGE) - -typedef (open) N_2 = "{a::bool recspace. - ALL u::bool recspace => bool. - (ALL a::bool recspace. - a = CONSTR (NUMERAL 0) (SOME x::bool. True) (%n::nat. BOTTOM) | - a = - CONSTR (Suc (NUMERAL 0)) (SOME x::bool. True) (%n::nat. BOTTOM) --> - u a) --> - u a}" morphisms "_dest_2" "_mk_2" - apply (rule light_ex_imp_nonempty[where t="CONSTR (NUMERAL 0) (SOME x::bool. True) (%n::nat. BOTTOM)"]) - by (import hollight TYDEF_2) - -syntax - "_dest_2" :: _ ("'_dest'_2") - -syntax - "_mk_2" :: _ ("'_mk'_2") - -lemmas "TYDEF_2_@intern" = typedef_hol2hollight - [where a="a :: N_2" and r=r , - OF type_definition_N_2] - -consts - "_10302" :: "N_2" ("'_10302") - -defs - "_10302_def": "(op ==::N_2 => N_2 => prop) (_10302::N_2) - ((_mk_2::bool recspace => N_2) - ((CONSTR::nat => bool => (nat => bool recspace) => bool recspace) - (0::nat) ((Eps::(bool => bool) => bool) (%x::bool. True::bool)) - (%n::nat. BOTTOM::bool recspace)))" - -lemma DEF__10302: "(op =::N_2 => N_2 => bool) (_10302::N_2) - ((_mk_2::bool recspace => N_2) - ((CONSTR::nat => bool => (nat => bool recspace) => bool recspace) - (0::nat) ((Eps::(bool => bool) => bool) (%x::bool. True::bool)) - (%n::nat. BOTTOM::bool recspace)))" - by (import hollight DEF__10302) - -consts - "_10303" :: "N_2" ("'_10303") - -defs - "_10303_def": "(op ==::N_2 => N_2 => prop) (_10303::N_2) - ((_mk_2::bool recspace => N_2) - ((CONSTR::nat => bool => (nat => bool recspace) => bool recspace) - ((Suc::nat => nat) (0::nat)) - ((Eps::(bool => bool) => bool) (%x::bool. True::bool)) - (%n::nat. BOTTOM::bool recspace)))" - -lemma DEF__10303: "(op =::N_2 => N_2 => bool) (_10303::N_2) - ((_mk_2::bool recspace => N_2) - ((CONSTR::nat => bool => (nat => bool recspace) => bool recspace) - ((Suc::nat => nat) (0::nat)) - ((Eps::(bool => bool) => bool) (%x::bool. True::bool)) - (%n::nat. BOTTOM::bool recspace)))" - by (import hollight DEF__10303) - -definition two_1 :: "N_2" where - "two_1 == _10302" - -lemma DEF_two_1: "two_1 = _10302" - by (import hollight DEF_two_1) - -definition two_2 :: "N_2" where - "two_2 == _10303" - -lemma DEF_two_2: "two_2 = _10303" - by (import hollight DEF_two_2) - -typedef (open) N_3 = "{a::bool recspace. - ALL u::bool recspace => bool. - (ALL a::bool recspace. - a = CONSTR (NUMERAL 0) (SOME x::bool. True) (%n::nat. BOTTOM) | - a = - CONSTR (Suc (NUMERAL 0)) (SOME x::bool. True) (%n::nat. BOTTOM) | - a = - CONSTR (Suc (Suc (NUMERAL 0))) (SOME x::bool. True) - (%n::nat. BOTTOM) --> - u a) --> - u a}" morphisms "_dest_3" "_mk_3" - apply (rule light_ex_imp_nonempty[where t="CONSTR (NUMERAL 0) (SOME x::bool. True) (%n::nat. BOTTOM)"]) - by (import hollight TYDEF_3) - -syntax - "_dest_3" :: _ ("'_dest'_3") - -syntax - "_mk_3" :: _ ("'_mk'_3") - -lemmas "TYDEF_3_@intern" = typedef_hol2hollight - [where a="a :: N_3" and r=r , - OF type_definition_N_3] - -consts - "_10326" :: "N_3" ("'_10326") - -defs - "_10326_def": "(op ==::N_3 => N_3 => prop) (_10326::N_3) - ((_mk_3::bool recspace => N_3) - ((CONSTR::nat => bool => (nat => bool recspace) => bool recspace) - (0::nat) ((Eps::(bool => bool) => bool) (%x::bool. True::bool)) - (%n::nat. BOTTOM::bool recspace)))" - -lemma DEF__10326: "(op =::N_3 => N_3 => bool) (_10326::N_3) - ((_mk_3::bool recspace => N_3) - ((CONSTR::nat => bool => (nat => bool recspace) => bool recspace) - (0::nat) ((Eps::(bool => bool) => bool) (%x::bool. True::bool)) - (%n::nat. BOTTOM::bool recspace)))" - by (import hollight DEF__10326) - -consts - "_10327" :: "N_3" ("'_10327") - -defs - "_10327_def": "(op ==::N_3 => N_3 => prop) (_10327::N_3) - ((_mk_3::bool recspace => N_3) - ((CONSTR::nat => bool => (nat => bool recspace) => bool recspace) - ((Suc::nat => nat) (0::nat)) - ((Eps::(bool => bool) => bool) (%x::bool. True::bool)) - (%n::nat. BOTTOM::bool recspace)))" - -lemma DEF__10327: "(op =::N_3 => N_3 => bool) (_10327::N_3) - ((_mk_3::bool recspace => N_3) - ((CONSTR::nat => bool => (nat => bool recspace) => bool recspace) - ((Suc::nat => nat) (0::nat)) - ((Eps::(bool => bool) => bool) (%x::bool. True::bool)) - (%n::nat. BOTTOM::bool recspace)))" - by (import hollight DEF__10327) - -consts - "_10328" :: "N_3" ("'_10328") + "_11937" :: "bool +=> bool => bool => bool => bool => bool => bool => bool => hollight.char" ("'_11937") defs - "_10328_def": "(op ==::N_3 => N_3 => prop) (_10328::N_3) - ((_mk_3::bool recspace => N_3) - ((CONSTR::nat => bool => (nat => bool recspace) => bool recspace) - ((Suc::nat => nat) ((Suc::nat => nat) (0::nat))) - ((Eps::(bool => bool) => bool) (%x::bool. True::bool)) - (%n::nat. BOTTOM::bool recspace)))" - -lemma DEF__10328: "(op =::N_3 => N_3 => bool) (_10328::N_3) - ((_mk_3::bool recspace => N_3) - ((CONSTR::nat => bool => (nat => bool recspace) => bool recspace) - ((Suc::nat => nat) ((Suc::nat => nat) (0::nat))) - ((Eps::(bool => bool) => bool) (%x::bool. True::bool)) - (%n::nat. BOTTOM::bool recspace)))" - by (import hollight DEF__10328) - -definition three_1 :: "N_3" where - "three_1 == _10326" - -lemma DEF_three_1: "three_1 = _10326" - by (import hollight DEF_three_1) - -definition three_2 :: "N_3" where - "three_2 == _10327" - -lemma DEF_three_2: "three_2 = _10327" - by (import hollight DEF_three_2) - -definition three_3 :: "N_3" where - "three_3 == _10328" - -lemma DEF_three_3: "three_3 = _10328" - by (import hollight DEF_three_3) - -lemma list_INDUCT: "ALL P::'A::type hollight.list => bool. - P NIL & - (ALL (a0::'A::type) a1::'A::type hollight.list. - P a1 --> P (CONS a0 a1)) --> - All P" - by (import hollight list_INDUCT) - -definition HD :: "'A hollight.list => 'A" where - "HD == -SOME HD::'A::type hollight.list => 'A::type. - ALL (t::'A::type hollight.list) h::'A::type. HD (CONS h t) = h" - -lemma DEF_HD: "HD = -(SOME HD::'A::type hollight.list => 'A::type. - ALL (t::'A::type hollight.list) h::'A::type. HD (CONS h t) = h)" - by (import hollight DEF_HD) - -definition TL :: "'A hollight.list => 'A hollight.list" where - "TL == -SOME TL::'A::type hollight.list => 'A::type hollight.list. - ALL (h::'A::type) t::'A::type hollight.list. TL (CONS h t) = t" - -lemma DEF_TL: "TL = -(SOME TL::'A::type hollight.list => 'A::type hollight.list. - ALL (h::'A::type) t::'A::type hollight.list. TL (CONS h t) = t)" - by (import hollight DEF_TL) - -definition APPEND :: "'A hollight.list => 'A hollight.list => 'A hollight.list" where - "APPEND == -SOME APPEND::'A::type hollight.list - => 'A::type hollight.list => 'A::type hollight.list. - (ALL l::'A::type hollight.list. APPEND NIL l = l) & - (ALL (h::'A::type) (t::'A::type hollight.list) l::'A::type hollight.list. - APPEND (CONS h t) l = CONS h (APPEND t l))" - -lemma DEF_APPEND: "APPEND = -(SOME APPEND::'A::type hollight.list - => 'A::type hollight.list => 'A::type hollight.list. - (ALL l::'A::type hollight.list. APPEND NIL l = l) & - (ALL (h::'A::type) (t::'A::type hollight.list) - l::'A::type hollight.list. - APPEND (CONS h t) l = CONS h (APPEND t l)))" - by (import hollight DEF_APPEND) - -definition REVERSE :: "'A hollight.list => 'A hollight.list" where - "REVERSE == -SOME REVERSE::'A::type hollight.list => 'A::type hollight.list. - REVERSE NIL = NIL & - (ALL (l::'A::type hollight.list) x::'A::type. - REVERSE (CONS x l) = APPEND (REVERSE l) (CONS x NIL))" - -lemma DEF_REVERSE: "REVERSE = -(SOME REVERSE::'A::type hollight.list => 'A::type hollight.list. - REVERSE NIL = NIL & - (ALL (l::'A::type hollight.list) x::'A::type. - REVERSE (CONS x l) = APPEND (REVERSE l) (CONS x NIL)))" - by (import hollight DEF_REVERSE) - -definition LENGTH :: "'A hollight.list => nat" where - "LENGTH == -SOME LENGTH::'A::type hollight.list => nat. - LENGTH NIL = 0 & - (ALL (h::'A::type) t::'A::type hollight.list. - LENGTH (CONS h t) = Suc (LENGTH t))" - -lemma DEF_LENGTH: "LENGTH = -(SOME LENGTH::'A::type hollight.list => nat. - LENGTH NIL = 0 & - (ALL (h::'A::type) t::'A::type hollight.list. - LENGTH (CONS h t) = Suc (LENGTH t)))" - by (import hollight DEF_LENGTH) - -definition MAP :: "('A => 'B) => 'A hollight.list => 'B hollight.list" where - "MAP == -SOME MAP::('A::type => 'B::type) - => 'A::type hollight.list => 'B::type hollight.list. - (ALL f::'A::type => 'B::type. MAP f NIL = NIL) & - (ALL (f::'A::type => 'B::type) (h::'A::type) t::'A::type hollight.list. - MAP f (CONS h t) = CONS (f h) (MAP f t))" - -lemma DEF_MAP: "MAP = -(SOME MAP::('A::type => 'B::type) - => 'A::type hollight.list => 'B::type hollight.list. - (ALL f::'A::type => 'B::type. MAP f NIL = NIL) & - (ALL (f::'A::type => 'B::type) (h::'A::type) t::'A::type hollight.list. - MAP f (CONS h t) = CONS (f h) (MAP f t)))" - by (import hollight DEF_MAP) - -definition LAST :: "'A hollight.list => 'A" where - "LAST == -SOME LAST::'A::type hollight.list => 'A::type. - ALL (h::'A::type) t::'A::type hollight.list. - LAST (CONS h t) = COND (t = NIL) h (LAST t)" - -lemma DEF_LAST: "LAST = -(SOME LAST::'A::type hollight.list => 'A::type. - ALL (h::'A::type) t::'A::type hollight.list. - LAST (CONS h t) = COND (t = NIL) h (LAST t))" - by (import hollight DEF_LAST) - -definition REPLICATE :: "nat => 'q_16860 => 'q_16860 hollight.list" where - "REPLICATE == -SOME REPLICATE::nat => 'q_16860::type => 'q_16860::type hollight.list. - (ALL x::'q_16860::type. REPLICATE 0 x = NIL) & - (ALL (n::nat) x::'q_16860::type. - REPLICATE (Suc n) x = CONS x (REPLICATE n x))" - -lemma DEF_REPLICATE: "REPLICATE = -(SOME REPLICATE::nat => 'q_16860::type => 'q_16860::type hollight.list. - (ALL x::'q_16860::type. REPLICATE 0 x = NIL) & - (ALL (n::nat) x::'q_16860::type. - REPLICATE (Suc n) x = CONS x (REPLICATE n x)))" - by (import hollight DEF_REPLICATE) - -definition NULL :: "'q_16875 hollight.list => bool" where - "NULL == -SOME NULL::'q_16875::type hollight.list => bool. - NULL NIL = True & - (ALL (h::'q_16875::type) t::'q_16875::type hollight.list. - NULL (CONS h t) = False)" - -lemma DEF_NULL: "NULL = -(SOME NULL::'q_16875::type hollight.list => bool. - NULL NIL = True & - (ALL (h::'q_16875::type) t::'q_16875::type hollight.list. - NULL (CONS h t) = False))" - by (import hollight DEF_NULL) - -definition ALL_list :: "('q_16895 => bool) => 'q_16895 hollight.list => bool" where - "ALL_list == -SOME u::('q_16895::type => bool) => 'q_16895::type hollight.list => bool. - (ALL P::'q_16895::type => bool. u P NIL = True) & - (ALL (h::'q_16895::type) (P::'q_16895::type => bool) - t::'q_16895::type hollight.list. u P (CONS h t) = (P h & u P t))" - -lemma DEF_ALL: "ALL_list = -(SOME u::('q_16895::type => bool) => 'q_16895::type hollight.list => bool. - (ALL P::'q_16895::type => bool. u P NIL = True) & - (ALL (h::'q_16895::type) (P::'q_16895::type => bool) - t::'q_16895::type hollight.list. u P (CONS h t) = (P h & u P t)))" - by (import hollight DEF_ALL) - -consts - EX :: "('q_16916 => bool) => 'q_16916 hollight.list => bool" ("EX") - -defs - EX_def: "EX == -SOME u::('q_16916::type => bool) => 'q_16916::type hollight.list => bool. - (ALL P::'q_16916::type => bool. u P NIL = False) & - (ALL (h::'q_16916::type) (P::'q_16916::type => bool) - t::'q_16916::type hollight.list. u P (CONS h t) = (P h | u P t))" - -lemma DEF_EX: "EX = -(SOME u::('q_16916::type => bool) => 'q_16916::type hollight.list => bool. - (ALL P::'q_16916::type => bool. u P NIL = False) & - (ALL (h::'q_16916::type) (P::'q_16916::type => bool) - t::'q_16916::type hollight.list. u P (CONS h t) = (P h | u P t)))" - by (import hollight DEF_EX) - -definition ITLIST :: "('q_16939 => 'q_16938 => 'q_16938) -=> 'q_16939 hollight.list => 'q_16938 => 'q_16938" where - "ITLIST == -SOME ITLIST::('q_16939::type => 'q_16938::type => 'q_16938::type) - => 'q_16939::type hollight.list - => 'q_16938::type => 'q_16938::type. - (ALL (f::'q_16939::type => 'q_16938::type => 'q_16938::type) - b::'q_16938::type. ITLIST f NIL b = b) & - (ALL (h::'q_16939::type) - (f::'q_16939::type => 'q_16938::type => 'q_16938::type) - (t::'q_16939::type hollight.list) b::'q_16938::type. - ITLIST f (CONS h t) b = f h (ITLIST f t b))" - -lemma DEF_ITLIST: "ITLIST = -(SOME ITLIST::('q_16939::type => 'q_16938::type => 'q_16938::type) - => 'q_16939::type hollight.list - => 'q_16938::type => 'q_16938::type. - (ALL (f::'q_16939::type => 'q_16938::type => 'q_16938::type) - b::'q_16938::type. ITLIST f NIL b = b) & - (ALL (h::'q_16939::type) - (f::'q_16939::type => 'q_16938::type => 'q_16938::type) - (t::'q_16939::type hollight.list) b::'q_16938::type. - ITLIST f (CONS h t) b = f h (ITLIST f t b)))" - by (import hollight DEF_ITLIST) - -definition MEM :: "'q_16964 => 'q_16964 hollight.list => bool" where - "MEM == -SOME MEM::'q_16964::type => 'q_16964::type hollight.list => bool. - (ALL x::'q_16964::type. MEM x NIL = False) & - (ALL (h::'q_16964::type) (x::'q_16964::type) - t::'q_16964::type hollight.list. - MEM x (CONS h t) = (x = h | MEM x t))" - -lemma DEF_MEM: "MEM = -(SOME MEM::'q_16964::type => 'q_16964::type hollight.list => bool. - (ALL x::'q_16964::type. MEM x NIL = False) & - (ALL (h::'q_16964::type) (x::'q_16964::type) - t::'q_16964::type hollight.list. - MEM x (CONS h t) = (x = h | MEM x t)))" - by (import hollight DEF_MEM) - -definition ALL2 :: "('q_16997 => 'q_17004 => bool) -=> 'q_16997 hollight.list => 'q_17004 hollight.list => bool" where - "ALL2 == -SOME ALL2::('q_16997::type => 'q_17004::type => bool) - => 'q_16997::type hollight.list - => 'q_17004::type hollight.list => bool. - (ALL (P::'q_16997::type => 'q_17004::type => bool) - l2::'q_17004::type hollight.list. ALL2 P NIL l2 = (l2 = NIL)) & - (ALL (h1::'q_16997::type) (P::'q_16997::type => 'q_17004::type => bool) - (t1::'q_16997::type hollight.list) l2::'q_17004::type hollight.list. - ALL2 P (CONS h1 t1) l2 = - COND (l2 = NIL) False (P h1 (HD l2) & ALL2 P t1 (TL l2)))" - -lemma DEF_ALL2: "ALL2 = -(SOME ALL2::('q_16997::type => 'q_17004::type => bool) - => 'q_16997::type hollight.list - => 'q_17004::type hollight.list => bool. - (ALL (P::'q_16997::type => 'q_17004::type => bool) - l2::'q_17004::type hollight.list. ALL2 P NIL l2 = (l2 = NIL)) & - (ALL (h1::'q_16997::type) (P::'q_16997::type => 'q_17004::type => bool) - (t1::'q_16997::type hollight.list) l2::'q_17004::type hollight.list. - ALL2 P (CONS h1 t1) l2 = - COND (l2 = NIL) False (P h1 (HD l2) & ALL2 P t1 (TL l2))))" - by (import hollight DEF_ALL2) - -lemma ALL2: "ALL2 (P::'q_17059::type => 'q_17058::type => bool) NIL NIL = True & -ALL2 P (CONS (h1::'q_17059::type) (t1::'q_17059::type hollight.list)) NIL = -False & -ALL2 P NIL (CONS (h2::'q_17058::type) (t2::'q_17058::type hollight.list)) = -False & -ALL2 P (CONS h1 t1) (CONS h2 t2) = (P h1 h2 & ALL2 P t1 t2)" - by (import hollight ALL2) - -definition MAP2 :: "('q_17089 => 'q_17096 => 'q_17086) -=> 'q_17089 hollight.list - => 'q_17096 hollight.list => 'q_17086 hollight.list" where - "MAP2 == -SOME MAP2::('q_17089::type => 'q_17096::type => 'q_17086::type) - => 'q_17089::type hollight.list - => 'q_17096::type hollight.list - => 'q_17086::type hollight.list. - (ALL (f::'q_17089::type => 'q_17096::type => 'q_17086::type) - l::'q_17096::type hollight.list. MAP2 f NIL l = NIL) & - (ALL (h1::'q_17089::type) - (f::'q_17089::type => 'q_17096::type => 'q_17086::type) - (t1::'q_17089::type hollight.list) l::'q_17096::type hollight.list. - MAP2 f (CONS h1 t1) l = CONS (f h1 (HD l)) (MAP2 f t1 (TL l)))" - -lemma DEF_MAP2: "MAP2 = -(SOME MAP2::('q_17089::type => 'q_17096::type => 'q_17086::type) - => 'q_17089::type hollight.list - => 'q_17096::type hollight.list - => 'q_17086::type hollight.list. - (ALL (f::'q_17089::type => 'q_17096::type => 'q_17086::type) - l::'q_17096::type hollight.list. MAP2 f NIL l = NIL) & - (ALL (h1::'q_17089::type) - (f::'q_17089::type => 'q_17096::type => 'q_17086::type) - (t1::'q_17089::type hollight.list) l::'q_17096::type hollight.list. - MAP2 f (CONS h1 t1) l = CONS (f h1 (HD l)) (MAP2 f t1 (TL l))))" - by (import hollight DEF_MAP2) - -lemma MAP2: "MAP2 (f::'q_17131::type => 'q_17130::type => 'q_17137::type) NIL NIL = NIL & -MAP2 f (CONS (h1::'q_17131::type) (t1::'q_17131::type hollight.list)) - (CONS (h2::'q_17130::type) (t2::'q_17130::type hollight.list)) = -CONS (f h1 h2) (MAP2 f t1 t2)" - by (import hollight MAP2) - -definition EL :: "nat => 'q_17157 hollight.list => 'q_17157" where - "EL == -SOME EL::nat => 'q_17157::type hollight.list => 'q_17157::type. - (ALL l::'q_17157::type hollight.list. EL 0 l = HD l) & - (ALL (n::nat) l::'q_17157::type hollight.list. - EL (Suc n) l = EL n (TL l))" - -lemma DEF_EL: "EL = -(SOME EL::nat => 'q_17157::type hollight.list => 'q_17157::type. - (ALL l::'q_17157::type hollight.list. EL 0 l = HD l) & - (ALL (n::nat) l::'q_17157::type hollight.list. - EL (Suc n) l = EL n (TL l)))" - by (import hollight DEF_EL) - -definition FILTER :: "('q_17182 => bool) => 'q_17182 hollight.list => 'q_17182 hollight.list" where - "FILTER == -SOME FILTER::('q_17182::type => bool) - => 'q_17182::type hollight.list - => 'q_17182::type hollight.list. - (ALL P::'q_17182::type => bool. FILTER P NIL = NIL) & - (ALL (h::'q_17182::type) (P::'q_17182::type => bool) - t::'q_17182::type hollight.list. - FILTER P (CONS h t) = COND (P h) (CONS h (FILTER P t)) (FILTER P t))" - -lemma DEF_FILTER: "FILTER = -(SOME FILTER::('q_17182::type => bool) - => 'q_17182::type hollight.list - => 'q_17182::type hollight.list. - (ALL P::'q_17182::type => bool. FILTER P NIL = NIL) & - (ALL (h::'q_17182::type) (P::'q_17182::type => bool) - t::'q_17182::type hollight.list. - FILTER P (CONS h t) = - COND (P h) (CONS h (FILTER P t)) (FILTER P t)))" - by (import hollight DEF_FILTER) - -definition ASSOC :: "'q_17211 => ('q_17211 * 'q_17205) hollight.list => 'q_17205" where - "ASSOC == -SOME ASSOC::'q_17211::type - => ('q_17211::type * 'q_17205::type) hollight.list - => 'q_17205::type. - ALL (h::'q_17211::type * 'q_17205::type) (a::'q_17211::type) - t::('q_17211::type * 'q_17205::type) hollight.list. - ASSOC a (CONS h t) = COND (fst h = a) (snd h) (ASSOC a t)" - -lemma DEF_ASSOC: "ASSOC = -(SOME ASSOC::'q_17211::type - => ('q_17211::type * 'q_17205::type) hollight.list - => 'q_17205::type. - ALL (h::'q_17211::type * 'q_17205::type) (a::'q_17211::type) - t::('q_17211::type * 'q_17205::type) hollight.list. - ASSOC a (CONS h t) = COND (fst h = a) (snd h) (ASSOC a t))" - by (import hollight DEF_ASSOC) - -definition ITLIST2 :: "('q_17235 => 'q_17243 => 'q_17233 => 'q_17233) -=> 'q_17235 hollight.list => 'q_17243 hollight.list => 'q_17233 => 'q_17233" where - "ITLIST2 == -SOME ITLIST2::('q_17235::type - => 'q_17243::type => 'q_17233::type => 'q_17233::type) - => 'q_17235::type hollight.list - => 'q_17243::type hollight.list - => 'q_17233::type => 'q_17233::type. - (ALL (f::'q_17235::type - => 'q_17243::type => 'q_17233::type => 'q_17233::type) - (l2::'q_17243::type hollight.list) b::'q_17233::type. - ITLIST2 f NIL l2 b = b) & - (ALL (h1::'q_17235::type) - (f::'q_17235::type - => 'q_17243::type => 'q_17233::type => 'q_17233::type) - (t1::'q_17235::type hollight.list) (l2::'q_17243::type hollight.list) - b::'q_17233::type. - ITLIST2 f (CONS h1 t1) l2 b = f h1 (HD l2) (ITLIST2 f t1 (TL l2) b))" - -lemma DEF_ITLIST2: "ITLIST2 = -(SOME ITLIST2::('q_17235::type - => 'q_17243::type => 'q_17233::type => 'q_17233::type) - => 'q_17235::type hollight.list - => 'q_17243::type hollight.list - => 'q_17233::type => 'q_17233::type. - (ALL (f::'q_17235::type - => 'q_17243::type => 'q_17233::type => 'q_17233::type) - (l2::'q_17243::type hollight.list) b::'q_17233::type. - ITLIST2 f NIL l2 b = b) & - (ALL (h1::'q_17235::type) - (f::'q_17235::type - => 'q_17243::type => 'q_17233::type => 'q_17233::type) - (t1::'q_17235::type hollight.list) - (l2::'q_17243::type hollight.list) b::'q_17233::type. - ITLIST2 f (CONS h1 t1) l2 b = - f h1 (HD l2) (ITLIST2 f t1 (TL l2) b)))" - by (import hollight DEF_ITLIST2) - -lemma ITLIST2: "ITLIST2 - (f::'q_17277::type => 'q_17276::type => 'q_17275::type => 'q_17275::type) - NIL NIL (b::'q_17275::type) = -b & -ITLIST2 f (CONS (h1::'q_17277::type) (t1::'q_17277::type hollight.list)) - (CONS (h2::'q_17276::type) (t2::'q_17276::type hollight.list)) b = -f h1 h2 (ITLIST2 f t1 t2 b)" - by (import hollight ITLIST2) + "_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 - ZIP :: "'q_17307 hollight.list -=> 'q_17315 hollight.list => ('q_17307 * 'q_17315) hollight.list" + dist :: "nat * nat => nat" defs - ZIP_def: "hollight.ZIP == -SOME ZIP::'q_17307::type hollight.list - => 'q_17315::type hollight.list - => ('q_17307::type * 'q_17315::type) hollight.list. - (ALL l2::'q_17315::type hollight.list. ZIP NIL l2 = NIL) & - (ALL (h1::'q_17307::type) (t1::'q_17307::type hollight.list) - l2::'q_17315::type hollight.list. - ZIP (CONS h1 t1) l2 = CONS (h1, HD l2) (ZIP t1 (TL l2)))" - -lemma DEF_ZIP: "hollight.ZIP = -(SOME ZIP::'q_17307::type hollight.list - => 'q_17315::type hollight.list - => ('q_17307::type * 'q_17315::type) hollight.list. - (ALL l2::'q_17315::type hollight.list. ZIP NIL l2 = NIL) & - (ALL (h1::'q_17307::type) (t1::'q_17307::type hollight.list) - l2::'q_17315::type hollight.list. - ZIP (CONS h1 t1) l2 = CONS (h1, HD l2) (ZIP t1 (TL l2))))" - by (import hollight DEF_ZIP) - -lemma ZIP: "(op &::bool => bool => bool) - ((op =::('q_17326::type * 'q_17327::type) hollight.list - => ('q_17326::type * 'q_17327::type) hollight.list => bool) - ((hollight.ZIP::'q_17326::type hollight.list - => 'q_17327::type hollight.list - => ('q_17326::type * 'q_17327::type) hollight.list) - (NIL::'q_17326::type hollight.list) - (NIL::'q_17327::type hollight.list)) - (NIL::('q_17326::type * 'q_17327::type) hollight.list)) - ((op =::('q_17351::type * 'q_17352::type) hollight.list - => ('q_17351::type * 'q_17352::type) hollight.list => bool) - ((hollight.ZIP::'q_17351::type hollight.list - => 'q_17352::type hollight.list - => ('q_17351::type * 'q_17352::type) hollight.list) - ((CONS::'q_17351::type - => 'q_17351::type hollight.list - => 'q_17351::type hollight.list) - (h1::'q_17351::type) (t1::'q_17351::type hollight.list)) - ((CONS::'q_17352::type - => 'q_17352::type hollight.list - => 'q_17352::type hollight.list) - (h2::'q_17352::type) (t2::'q_17352::type hollight.list))) - ((CONS::'q_17351::type * 'q_17352::type - => ('q_17351::type * 'q_17352::type) hollight.list - => ('q_17351::type * 'q_17352::type) hollight.list) - ((Pair::'q_17351::type - => 'q_17352::type => 'q_17351::type * 'q_17352::type) - h1 h2) - ((hollight.ZIP::'q_17351::type hollight.list - => 'q_17352::type hollight.list - => ('q_17351::type * 'q_17352::type) hollight.list) - t1 t2)))" - by (import hollight ZIP) - -lemma NOT_CONS_NIL: "ALL (x::'A::type) xa::'A::type hollight.list. CONS x xa ~= NIL" - by (import hollight NOT_CONS_NIL) - -lemma LAST_CLAUSES: "LAST (CONS (h::'A::type) NIL) = h & -LAST (CONS h (CONS (k::'A::type) (t::'A::type hollight.list))) = -LAST (CONS k t)" - by (import hollight LAST_CLAUSES) - -lemma APPEND_NIL: "ALL l::'A::type hollight.list. APPEND l NIL = l" - by (import hollight APPEND_NIL) - -lemma APPEND_ASSOC: "ALL (l::'A::type hollight.list) (m::'A::type hollight.list) - n::'A::type hollight.list. APPEND l (APPEND m n) = APPEND (APPEND l m) n" - by (import hollight APPEND_ASSOC) - -lemma REVERSE_APPEND: "ALL (l::'A::type hollight.list) m::'A::type hollight.list. - REVERSE (APPEND l m) = APPEND (REVERSE m) (REVERSE l)" - by (import hollight REVERSE_APPEND) - -lemma REVERSE_REVERSE: "ALL l::'A::type hollight.list. REVERSE (REVERSE l) = l" - by (import hollight REVERSE_REVERSE) - -lemma CONS_11: "ALL (x::'A::type) (xa::'A::type) (xb::'A::type hollight.list) - xc::'A::type hollight.list. (CONS x xb = CONS xa xc) = (x = xa & xb = xc)" - by (import hollight CONS_11) - -lemma list_CASES: "ALL l::'A::type hollight.list. - l = NIL | (EX (h::'A::type) t::'A::type hollight.list. l = CONS h t)" - by (import hollight list_CASES) - -lemma LENGTH_APPEND: "ALL (l::'A::type hollight.list) m::'A::type hollight.list. - LENGTH (APPEND l m) = LENGTH l + LENGTH m" - by (import hollight LENGTH_APPEND) - -lemma MAP_APPEND: "ALL (f::'A::type => 'B::type) (l1::'A::type hollight.list) - l2::'A::type hollight.list. - MAP f (APPEND l1 l2) = APPEND (MAP f l1) (MAP f l2)" - by (import hollight MAP_APPEND) - -lemma LENGTH_MAP: "ALL (l::'A::type hollight.list) f::'A::type => 'B::type. - LENGTH (MAP f l) = LENGTH l" - by (import hollight LENGTH_MAP) - -lemma LENGTH_EQ_NIL: "ALL l::'A::type hollight.list. (LENGTH l = 0) = (l = NIL)" - by (import hollight LENGTH_EQ_NIL) - -lemma LENGTH_EQ_CONS: "ALL (l::'q_17659::type hollight.list) n::nat. - (LENGTH l = Suc n) = - (EX (h::'q_17659::type) t::'q_17659::type hollight.list. - l = CONS h t & LENGTH t = n)" - by (import hollight LENGTH_EQ_CONS) - -lemma MAP_o: "ALL (f::'A::type => 'B::type) (g::'B::type => 'C::type) - l::'A::type hollight.list. MAP (g o f) l = MAP g (MAP f l)" - by (import hollight MAP_o) - -lemma MAP_EQ: "ALL (f::'q_17723::type => 'q_17734::type) - (g::'q_17723::type => 'q_17734::type) l::'q_17723::type hollight.list. - ALL_list (%x::'q_17723::type. f x = g x) l --> MAP f l = MAP g l" - by (import hollight MAP_EQ) - -lemma ALL_IMP: "ALL (P::'q_17764::type => bool) (Q::'q_17764::type => bool) - l::'q_17764::type hollight.list. - (ALL x::'q_17764::type. MEM x l & P x --> Q x) & ALL_list P l --> - ALL_list Q l" - by (import hollight ALL_IMP) - -lemma NOT_EX: "ALL (P::'q_17792::type => bool) l::'q_17792::type hollight.list. - (~ EX P l) = ALL_list (%x::'q_17792::type. ~ P x) l" - by (import hollight NOT_EX) - -lemma NOT_ALL: "ALL (P::'q_17814::type => bool) l::'q_17814::type hollight.list. - (~ ALL_list P l) = EX (%x::'q_17814::type. ~ P x) l" - by (import hollight NOT_ALL) - -lemma ALL_MAP: "ALL (P::'q_17836::type => bool) (f::'q_17835::type => 'q_17836::type) - l::'q_17835::type hollight.list. - ALL_list P (MAP f l) = ALL_list (P o f) l" - by (import hollight ALL_MAP) - -lemma ALL_T: "All (ALL_list (%x::'q_17854::type. True))" - by (import hollight ALL_T) - -lemma MAP_EQ_ALL2: "ALL (l::'q_17879::type hollight.list) m::'q_17879::type hollight.list. - ALL2 - (%(x::'q_17879::type) y::'q_17879::type. - (f::'q_17879::type => 'q_17890::type) x = f y) - l m --> - MAP f l = MAP f m" - by (import hollight MAP_EQ_ALL2) - -lemma ALL2_MAP: "ALL (P::'q_17921::type => 'q_17922::type => bool) - (f::'q_17922::type => 'q_17921::type) l::'q_17922::type hollight.list. - ALL2 P (MAP f l) l = ALL_list (%a::'q_17922::type. P (f a) a) l" - by (import hollight ALL2_MAP) - -lemma MAP_EQ_DEGEN: "ALL (l::'q_17939::type hollight.list) f::'q_17939::type => 'q_17939::type. - ALL_list (%x::'q_17939::type. f x = x) l --> MAP f l = l" - by (import hollight MAP_EQ_DEGEN) - -lemma ALL2_AND_RIGHT: "ALL (l::'q_17982::type hollight.list) (m::'q_17981::type hollight.list) - (P::'q_17982::type => bool) Q::'q_17982::type => 'q_17981::type => bool. - ALL2 (%(x::'q_17982::type) y::'q_17981::type. P x & Q x y) l m = - (ALL_list P l & ALL2 Q l m)" - by (import hollight ALL2_AND_RIGHT) - -lemma ITLIST_EXTRA: "ALL l::'q_18019::type hollight.list. - ITLIST (f::'q_18019::type => 'q_18018::type => 'q_18018::type) - (APPEND l (CONS (a::'q_18019::type) NIL)) (b::'q_18018::type) = - ITLIST f l (f a b)" - by (import hollight ITLIST_EXTRA) - -lemma ALL_MP: "ALL (P::'q_18045::type => bool) (Q::'q_18045::type => bool) - l::'q_18045::type hollight.list. - ALL_list (%x::'q_18045::type. P x --> Q x) l & ALL_list P l --> - ALL_list Q l" - by (import hollight ALL_MP) - -lemma AND_ALL: "ALL x::'q_18075::type hollight.list. - (ALL_list (P::'q_18075::type => bool) x & - ALL_list (Q::'q_18075::type => bool) x) = - ALL_list (%x::'q_18075::type. P x & Q x) x" - by (import hollight AND_ALL) - -lemma EX_IMP: "ALL (P::'q_18105::type => bool) (Q::'q_18105::type => bool) - l::'q_18105::type hollight.list. - (ALL x::'q_18105::type. MEM x l & P x --> Q x) & EX P l --> EX Q l" - by (import hollight EX_IMP) - -lemma ALL_MEM: "ALL (P::'q_18132::type => bool) l::'q_18132::type hollight.list. - (ALL x::'q_18132::type. MEM x l --> P x) = ALL_list P l" - by (import hollight ALL_MEM) - -lemma LENGTH_REPLICATE: "ALL (n::nat) x::'q_18150::type. LENGTH (REPLICATE n x) = n" - by (import hollight LENGTH_REPLICATE) - -lemma EX_MAP: "ALL (P::'q_18174::type => bool) (f::'q_18173::type => 'q_18174::type) - l::'q_18173::type hollight.list. EX P (MAP f l) = EX (P o f) l" - by (import hollight EX_MAP) - -lemma EXISTS_EX: "ALL (P::'q_18212::type => 'q_18211::type => bool) - l::'q_18211::type hollight.list. - (EX x::'q_18212::type. EX (P x) l) = - EX (%s::'q_18211::type. EX x::'q_18212::type. P x s) l" - by (import hollight EXISTS_EX) - -lemma FORALL_ALL: "ALL (P::'q_18242::type => 'q_18241::type => bool) - l::'q_18241::type hollight.list. - (ALL x::'q_18242::type. ALL_list (P x) l) = - ALL_list (%s::'q_18241::type. ALL x::'q_18242::type. P x s) l" - by (import hollight FORALL_ALL) - -lemma MEM_APPEND: "ALL (x::'q_18270::type) (l1::'q_18270::type hollight.list) - l2::'q_18270::type hollight.list. - MEM x (APPEND l1 l2) = (MEM x l1 | MEM x l2)" - by (import hollight MEM_APPEND) - -lemma MEM_MAP: "ALL (f::'q_18306::type => 'q_18303::type) (y::'q_18303::type) - l::'q_18306::type hollight.list. - MEM y (MAP f l) = (EX x::'q_18306::type. MEM x l & y = f x)" - by (import hollight MEM_MAP) - -lemma FILTER_APPEND: "ALL (P::'q_18337::type => bool) (l1::'q_18337::type hollight.list) - l2::'q_18337::type hollight.list. - FILTER P (APPEND l1 l2) = APPEND (FILTER P l1) (FILTER P l2)" - by (import hollight FILTER_APPEND) - -lemma FILTER_MAP: "ALL (P::'q_18364::type => bool) (f::'q_18371::type => 'q_18364::type) - l::'q_18371::type hollight.list. - FILTER P (MAP f l) = MAP f (FILTER (P o f) l)" - by (import hollight FILTER_MAP) - -lemma MEM_FILTER: "ALL (P::'q_18398::type => bool) (l::'q_18398::type hollight.list) - x::'q_18398::type. MEM x (FILTER P l) = (P x & MEM x l)" - by (import hollight MEM_FILTER) - -lemma EX_MEM: "ALL (P::'q_18419::type => bool) l::'q_18419::type hollight.list. - (EX x::'q_18419::type. P x & MEM x l) = EX P l" - by (import hollight EX_MEM) - -lemma MAP_FST_ZIP: "ALL (l1::'q_18439::type hollight.list) l2::'q_18441::type hollight.list. - LENGTH l1 = LENGTH l2 --> MAP fst (hollight.ZIP l1 l2) = l1" - by (import hollight MAP_FST_ZIP) - -lemma MAP_SND_ZIP: "ALL (l1::'q_18470::type hollight.list) l2::'q_18472::type hollight.list. - LENGTH l1 = LENGTH l2 --> MAP snd (hollight.ZIP l1 l2) = l2" - by (import hollight MAP_SND_ZIP) - -lemma MEM_ASSOC: "ALL (l::('q_18516::type * 'q_18500::type) hollight.list) x::'q_18516::type. - MEM (x, ASSOC x l) l = MEM x (MAP fst l)" - by (import hollight MEM_ASSOC) - -lemma ALL_APPEND: "ALL (P::'q_18537::type => bool) (l1::'q_18537::type hollight.list) - l2::'q_18537::type hollight.list. - ALL_list P (APPEND l1 l2) = (ALL_list P l1 & ALL_list P l2)" - by (import hollight ALL_APPEND) - -lemma MEM_EL: "ALL (l::'q_18560::type hollight.list) n::nat. - < n (LENGTH l) --> MEM (EL n l) l" - by (import hollight MEM_EL) - -lemma ALL2_MAP2: "ALL (l::'q_18603::type hollight.list) m::'q_18604::type hollight.list. - ALL2 (P::'q_18602::type => 'q_18601::type => bool) - (MAP (f::'q_18603::type => 'q_18602::type) l) - (MAP (g::'q_18604::type => 'q_18601::type) m) = - ALL2 (%(x::'q_18603::type) y::'q_18604::type. P (f x) (g y)) l m" - by (import hollight ALL2_MAP2) - -lemma AND_ALL2: "ALL (P::'q_18650::type => 'q_18649::type => bool) - (Q::'q_18650::type => 'q_18649::type => bool) - (x::'q_18650::type hollight.list) xa::'q_18649::type hollight.list. - (ALL2 P x xa & ALL2 Q x xa) = - ALL2 (%(x::'q_18650::type) y::'q_18649::type. P x y & Q x y) x xa" - by (import hollight AND_ALL2) - -lemma ALL2_ALL: "ALL (P::'q_18672::type => 'q_18672::type => bool) - l::'q_18672::type hollight.list. - ALL2 P l l = ALL_list (%x::'q_18672::type. P x x) l" - by (import hollight ALL2_ALL) - -lemma APPEND_EQ_NIL: "ALL (x::'q_18701::type hollight.list) xa::'q_18701::type hollight.list. - (APPEND x xa = NIL) = (x = NIL & xa = NIL)" - by (import hollight APPEND_EQ_NIL) - -lemma LENGTH_MAP2: "ALL (f::'q_18721::type => 'q_18723::type => 'q_18734::type) - (l::'q_18721::type hollight.list) m::'q_18723::type hollight.list. - LENGTH l = LENGTH m --> LENGTH (MAP2 f l m) = LENGTH m" - by (import hollight LENGTH_MAP2) - -lemma MONO_ALL: "(ALL x::'A::type. (P::'A::type => bool) x --> (Q::'A::type => bool) x) --> -ALL_list P (l::'A::type hollight.list) --> ALL_list Q l" - by (import hollight MONO_ALL) - -lemma MONO_ALL2: "(ALL (x::'A::type) y::'B::type. - (P::'A::type => 'B::type => bool) x y --> - (Q::'A::type => 'B::type => bool) x y) --> -ALL2 P (l::'A::type hollight.list) (l'::'B::type hollight.list) --> -ALL2 Q l l'" - by (import hollight MONO_ALL2) - -definition dist :: "nat * nat => nat" where - "dist == %u::nat * nat. fst u - snd u + (snd u - fst u)" - -lemma DEF_dist: "dist = (%u::nat * nat. fst u - snd u + (snd u - fst u))" + 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: "ALL x::nat. dist (x, x) = 0" +lemma DIST_REFL: "hollight.dist (x, x) = 0" by (import hollight DIST_REFL) -lemma DIST_LZERO: "ALL x::nat. dist (0, x) = x" +lemma DIST_LZERO: "hollight.dist (0, x) = x" by (import hollight DIST_LZERO) -lemma DIST_RZERO: "ALL x::nat. dist (x, 0) = x" +lemma DIST_RZERO: "hollight.dist (x, 0) = x" by (import hollight DIST_RZERO) -lemma DIST_SYM: "ALL (x::nat) xa::nat. dist (x, xa) = dist (xa, x)" +lemma DIST_SYM: "hollight.dist (x, xa) = hollight.dist (xa, x)" by (import hollight DIST_SYM) -lemma DIST_LADD: "ALL (x::nat) (xa::nat) xb::nat. dist (x + xb, x + xa) = dist (xb, xa)" +lemma DIST_LADD: "hollight.dist (x + xb, x + xa) = hollight.dist (xb, xa)" by (import hollight DIST_LADD) -lemma DIST_RADD: "ALL (x::nat) (xa::nat) xb::nat. dist (x + xa, xb + xa) = dist (x, xb)" +lemma DIST_RADD: "hollight.dist (x + xa, xb + xa) = hollight.dist (x, xb)" by (import hollight DIST_RADD) -lemma DIST_LADD_0: "ALL (x::nat) xa::nat. dist (x + xa, x) = xa" +lemma DIST_LADD_0: "hollight.dist (x + xa, x) = xa" by (import hollight DIST_LADD_0) -lemma DIST_RADD_0: "ALL (x::nat) xa::nat. dist (x, x + xa) = xa" +lemma DIST_RADD_0: "hollight.dist (x, x + xa) = xa" by (import hollight DIST_RADD_0) -lemma DIST_LMUL: "ALL (x::nat) (xa::nat) xb::nat. x * dist (xa, xb) = dist (x * xa, x * xb)" +lemma DIST_LMUL: "x * hollight.dist (xa, xb) = hollight.dist (x * xa, x * xb)" by (import hollight DIST_LMUL) -lemma DIST_RMUL: "ALL (x::nat) (xa::nat) xb::nat. dist (x, xa) * xb = dist (x * xb, xa * xb)" +lemma DIST_RMUL: "hollight.dist (x, xa) * xb = hollight.dist (x * xb, xa * xb)" by (import hollight DIST_RMUL) -lemma DIST_EQ_0: "ALL (x::nat) xa::nat. (dist (x, xa) = 0) = (x = xa)" +lemma DIST_EQ_0: "(hollight.dist (x, xa) = 0) = (x = xa)" by (import hollight DIST_EQ_0) -lemma DIST_ELIM_THM: "(P::nat => bool) (dist (x::nat, y::nat)) = -(ALL d::nat. (x = y + d --> P d) & (y = x + d --> P d))" +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: "ALL (m::nat) (n::nat) p::nat. - <= (dist (m, n)) p = (<= m (n + p) & <= n (m + p))" +lemma DIST_LE_CASES: "(hollight.dist (m, n) <= p) = (m <= n + p & n <= m + p)" by (import hollight DIST_LE_CASES) -lemma DIST_ADDBOUND: "ALL (m::nat) n::nat. <= (dist (m, n)) (m + n)" - by (import hollight DIST_ADDBOUND) - -lemma DIST_TRIANGLE: "ALL (m::nat) (n::nat) p::nat. <= (dist (m, p)) (dist (m, n) + dist (n, p))" - by (import hollight DIST_TRIANGLE) - -lemma DIST_ADD2: "ALL (m::nat) (n::nat) (p::nat) q::nat. - <= (dist (m + n, p + q)) (dist (m, p) + dist (n, q))" - by (import hollight DIST_ADD2) - -lemma DIST_ADD2_REV: "ALL (m::nat) (n::nat) (p::nat) q::nat. - <= (dist (m, p)) (dist (m + n, p + q) + dist (n, q))" - by (import hollight DIST_ADD2_REV) - -lemma DIST_TRIANGLE_LE: "ALL (m::nat) (n::nat) (p::nat) q::nat. - <= (dist (m, n) + dist (n, p)) q --> <= (dist (m, p)) q" +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: "ALL (m::nat) (n::nat) (p::nat) (q::nat) (r::nat) s::nat. - <= (dist (m, n)) r & <= (dist (p, q)) s --> - <= (dist (m, p)) (dist (n, q) + (r + s))" +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 (A::nat) (B::nat) C::nat. (ALL n::nat. <= (A * n) (B * n + C)) = <= A B" +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 (A::nat) B::nat. (ALL n::nat. <= (A * n) B) = (A = 0)" +lemma BOUNDS_LINEAR_0: "(ALL n::nat. (A::nat) * n <= (B::nat)) = (A = (0::nat))" by (import hollight BOUNDS_LINEAR_0) -lemma BOUNDS_DIVIDED: "ALL P::nat => nat. - (EX B::nat. ALL n::nat. <= (P n) B) = - (EX (x::nat) B::nat. ALL n::nat. <= (n * P n) (x * n + B))" +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: "ALL (P::nat => nat => nat) (A::nat) B::nat. - P 0 0 = 0 & (ALL (m::nat) n::nat. <= (P m n) (A * (m + n) + B)) --> - (EX x::nat. ALL (m::nat) n::nat. <= (P m n) (x * (m + n)))" +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: "ALL (P::nat => nat) Q::nat => nat. - (EX B::nat. ALL i::nat. <= (P i) (Q i + B)) = - (EX (x::nat) N::nat. ALL i::nat. <= N i --> <= (P i) (Q i + x))" +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 +definition + is_nadd :: "(nat => nat) => bool" where "is_nadd == -%u::nat => nat. - EX B::nat. - ALL (m::nat) n::nat. <= (dist (m * u n, n * u m)) (B * (m + n))" +%u. EX B. ALL m n. hollight.dist (m * u n, n * u m) <= B * (m + n)" lemma DEF_is_nadd: "is_nadd = -(%u::nat => nat. - EX B::nat. - ALL (m::nat) n::nat. <= (dist (m * u n, n * u m)) (B * (m + n)))" +(%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::nat. 0)" +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::nat. NUMERAL 0"]) + apply (rule light_ex_imp_nonempty[where t="%n. NUMERAL 0"]) by (import hollight TYDEF_nadd) syntax @@ -3100,394 +1411,329 @@ [where a="a :: nadd" and r=r , OF type_definition_nadd] -lemma NADD_CAUCHY: "ALL x::nadd. - EX xa::nat. - ALL (xb::nat) xc::nat. - <= (dist (xb * dest_nadd x xc, xc * dest_nadd x xb)) - (xa * (xb + xc))" +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: "ALL x::nadd. - EX (xa::nat) B::nat. ALL n::nat. <= (dest_nadd x n) (xa * n + B)" +lemma NADD_BOUND: "EX xa B. ALL n. dest_nadd x n <= xa * n + B" by (import hollight NADD_BOUND) -lemma NADD_MULTIPLICATIVE: "ALL x::nadd. - EX xa::nat. - ALL (m::nat) n::nat. - <= (dist (dest_nadd x (m * n), m * dest_nadd x n)) (xa * m + xa)" +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: "ALL x::nadd. - EX xa::nat. - ALL (m::nat) n::nat. - <= (dist (dest_nadd x (m + n), dest_nadd x m + dest_nadd x n)) xa" +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: "ALL x::nadd. - EX xa::nat. ALL n::nat. <= (dist (dest_nadd x (Suc n), dest_nadd x n)) xa" +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: "ALL x::nadd. - EX xa::nat. - ALL (m::nat) n::nat. - <= (dist (dest_nadd x (m + n), dest_nadd x m)) (xa * n)" +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: "ALL x::nadd. - EX xa::nat. - ALL (m::nat) n::nat. - <= (dist (dest_nadd x m, dest_nadd x n)) (xa * dist (m, n))" +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: "ALL (x::nadd) y::nadd. - EX (A::nat) B::nat. - ALL n::nat. - <= (dist - (n * dest_nadd x (dest_nadd y n), - dest_nadd x n * dest_nadd y n)) - (A * n + B)" +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 +definition + nadd_eq :: "nadd => nadd => bool" where "nadd_eq == -%(u::nadd) ua::nadd. - EX B::nat. ALL n::nat. <= (dist (dest_nadd u n, dest_nadd ua n)) B" +%u ua. EX B. ALL n. hollight.dist (dest_nadd u n, dest_nadd ua n) <= B" lemma DEF_nadd_eq: "nadd_eq = -(%(u::nadd) ua::nadd. - EX B::nat. ALL n::nat. <= (dist (dest_nadd u n, dest_nadd ua n)) B)" +(%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: "ALL x::nadd. nadd_eq x x" +lemma NADD_EQ_REFL: "nadd_eq x x" by (import hollight NADD_EQ_REFL) -lemma NADD_EQ_SYM: "ALL (x::nadd) y::nadd. nadd_eq x y = nadd_eq y x" +lemma NADD_EQ_SYM: "nadd_eq x y = nadd_eq y x" by (import hollight NADD_EQ_SYM) -lemma NADD_EQ_TRANS: "ALL (x::nadd) (y::nadd) z::nadd. nadd_eq x y & nadd_eq y z --> nadd_eq x z" +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::nat. mk_nadd (op * u)" - -lemma DEF_nadd_of_num: "nadd_of_num = (%u::nat. mk_nadd (op * u))" +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: "ALL x::nat. dest_nadd (nadd_of_num x) = op * x" +lemma NADD_OF_NUM: "dest_nadd (nadd_of_num x) = op * x" by (import hollight NADD_OF_NUM) -lemma NADD_OF_NUM_WELLDEF: "ALL (m::nat) n::nat. m = n --> nadd_eq (nadd_of_num m) (nadd_of_num n)" +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: "ALL (m::nat) n::nat. nadd_eq (nadd_of_num m) (nadd_of_num n) = (m = n)" +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::nadd) ua::nadd. - EX B::nat. ALL n::nat. <= (dest_nadd u n) (dest_nadd ua n + B)" - -lemma DEF_nadd_le: "nadd_le = -(%(u::nadd) ua::nadd. - EX B::nat. ALL n::nat. <= (dest_nadd u n) (dest_nadd ua n + B))" +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: "ALL (x::nadd) (x'::nadd) (y::nadd) y'::nadd. - nadd_eq x x' & nadd_eq y y' & nadd_le x y --> nadd_le x' y'" +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: "ALL (x::nadd) (x'::nadd) (y::nadd) y'::nadd. - nadd_eq x x' & nadd_eq y y' --> nadd_le x y = nadd_le x' y'" +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: "ALL x::nadd. nadd_le x x" +lemma NADD_LE_REFL: "nadd_le x x" by (import hollight NADD_LE_REFL) -lemma NADD_LE_TRANS: "ALL (x::nadd) (y::nadd) z::nadd. nadd_le x y & nadd_le y z --> nadd_le x z" +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: "ALL (x::nadd) y::nadd. (nadd_le x y & nadd_le y x) = nadd_eq x y" +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: "ALL (x::nadd) y::nadd. - ~ nadd_le x y --> - (ALL B::nat. EX n::nat. n ~= 0 & < (dest_nadd y n + B) (dest_nadd x n))" +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: "ALL (x::nadd) y::nadd. nadd_le x y | nadd_le y x" +lemma NADD_LE_TOTAL: "nadd_le x y | nadd_le y x" by (import hollight NADD_LE_TOTAL) -lemma NADD_ARCH: "ALL x::nadd. EX xa::nat. nadd_le x (nadd_of_num xa)" +lemma NADD_ARCH: "EX xa. nadd_le x (nadd_of_num xa)" by (import hollight NADD_ARCH) -lemma NADD_OF_NUM_LE: "ALL (m::nat) n::nat. nadd_le (nadd_of_num m) (nadd_of_num n) = <= m n" +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::nadd) ua::nadd. mk_nadd (%n::nat. dest_nadd u n + dest_nadd ua n)" - -lemma DEF_nadd_add: "nadd_add = -(%(u::nadd) ua::nadd. mk_nadd (%n::nat. dest_nadd u n + dest_nadd ua n))" +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: "ALL (x::nadd) y::nadd. - dest_nadd (nadd_add x y) = (%n::nat. dest_nadd x n + dest_nadd y n)" +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: "ALL (x::nadd) (x'::nadd) (y::nadd) y'::nadd. - nadd_eq x x' & nadd_eq y y' --> nadd_eq (nadd_add x y) (nadd_add x' y')" +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: "ALL (x::nadd) y::nadd. nadd_eq (nadd_add x y) (nadd_add y x)" +lemma NADD_ADD_SYM: "nadd_eq (nadd_add x y) (nadd_add y x)" by (import hollight NADD_ADD_SYM) -lemma NADD_ADD_ASSOC: "ALL (x::nadd) (y::nadd) z::nadd. - nadd_eq (nadd_add x (nadd_add y z)) (nadd_add (nadd_add x y) z)" +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: "ALL x::nadd. nadd_eq (nadd_add (nadd_of_num 0) x) x" +lemma NADD_ADD_LID: "nadd_eq (nadd_add (nadd_of_num 0) x) x" by (import hollight NADD_ADD_LID) -lemma NADD_ADD_LCANCEL: "ALL (x::nadd) (y::nadd) z::nadd. - nadd_eq (nadd_add x y) (nadd_add x z) --> nadd_eq y z" +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: "ALL (x::nadd) y::nadd. nadd_le x (nadd_add x y)" +lemma NADD_LE_ADD: "nadd_le x (nadd_add x y)" by (import hollight NADD_LE_ADD) -lemma NADD_LE_EXISTS: "ALL (x::nadd) y::nadd. - nadd_le x y --> (EX d::nadd. nadd_eq y (nadd_add x d))" +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: "ALL (x::nat) xa::nat. - nadd_eq (nadd_add (nadd_of_num x) (nadd_of_num xa)) - (nadd_of_num (x + xa))" +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::nadd) ua::nadd. mk_nadd (%n::nat. dest_nadd u (dest_nadd ua n))" - -lemma DEF_nadd_mul: "nadd_mul = -(%(u::nadd) ua::nadd. mk_nadd (%n::nat. dest_nadd u (dest_nadd ua n)))" +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: "ALL (x::nadd) y::nadd. - dest_nadd (nadd_mul x y) = (%n::nat. dest_nadd x (dest_nadd y n))" +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: "ALL (x::nadd) y::nadd. nadd_eq (nadd_mul x y) (nadd_mul y x)" +lemma NADD_MUL_SYM: "nadd_eq (nadd_mul x y) (nadd_mul y x)" by (import hollight NADD_MUL_SYM) -lemma NADD_MUL_ASSOC: "ALL (x::nadd) (y::nadd) z::nadd. - nadd_eq (nadd_mul x (nadd_mul y z)) (nadd_mul (nadd_mul x y) z)" +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: "ALL x::nadd. nadd_eq (nadd_mul (nadd_of_num (NUMERAL_BIT1 0)) x) x" +lemma NADD_MUL_LID: "nadd_eq (nadd_mul (nadd_of_num 1) x) x" by (import hollight NADD_MUL_LID) -lemma NADD_LDISTRIB: "ALL (x::nadd) (y::nadd) z::nadd. - nadd_eq (nadd_mul x (nadd_add y z)) - (nadd_add (nadd_mul x y) (nadd_mul x z))" +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: "ALL (x::nadd) (y::nadd) y'::nadd. - nadd_eq y y' --> nadd_eq (nadd_mul x y) (nadd_mul x y')" +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: "ALL (x::nadd) (x'::nadd) (y::nadd) y'::nadd. - nadd_eq x x' & nadd_eq y y' --> nadd_eq (nadd_mul x y) (nadd_mul x' y')" +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: "ALL (x::nat) xa::nat. - nadd_eq (nadd_mul (nadd_of_num x) (nadd_of_num xa)) - (nadd_of_num (x * xa))" +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: "All (nadd_le (nadd_of_num 0))" +lemma NADD_LE_0: "nadd_le (nadd_of_num 0) x" by (import hollight NADD_LE_0) -lemma NADD_EQ_IMP_LE: "ALL (x::nadd) y::nadd. nadd_eq x y --> nadd_le x y" +lemma NADD_EQ_IMP_LE: "nadd_eq x y ==> nadd_le x y" by (import hollight NADD_EQ_IMP_LE) -lemma NADD_LE_LMUL: "ALL (x::nadd) (y::nadd) z::nadd. - nadd_le y z --> nadd_le (nadd_mul x y) (nadd_mul x z)" +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: "ALL (x::nadd) (y::nadd) z::nadd. - nadd_le x y --> nadd_le (nadd_mul x z) (nadd_mul y z)" +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: "ALL (x::nadd) (y::nadd) z::nadd. - nadd_le (nadd_add x z) (nadd_add y z) = nadd_le x y" +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: "ALL (x::nadd) (y::nadd) z::nadd. - nadd_le (nadd_add x y) (nadd_add x z) = nadd_le y z" +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: "ALL (x::nadd) (y::nadd) z::nadd. - nadd_eq (nadd_mul (nadd_add x y) z) - (nadd_add (nadd_mul x z) (nadd_mul y z))" +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: "ALL (x::nadd) k::nat. - ~ nadd_eq x (nadd_of_num 0) --> - (EX xa::nat. nadd_le (nadd_of_num k) (nadd_mul (nadd_of_num xa) x))" +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: "ALL (x::nadd) k::nadd. - (ALL n::nat. nadd_le (nadd_mul (nadd_of_num n) x) k) --> - nadd_eq x (nadd_of_num 0)" +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: "ALL (x::nadd) (y::nadd) z::nadd. - (ALL n::nat. - nadd_le (nadd_mul (nadd_of_num n) x) - (nadd_add (nadd_mul (nadd_of_num n) y) z)) --> - nadd_le x y" +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: "ALL P::nadd => bool. - Ex P & (EX M::nadd. ALL x::nadd. P x --> nadd_le x M) --> - (EX M::nadd. - (ALL x::nadd. P x --> nadd_le x M) & - (ALL M'::nadd. (ALL x::nadd. P x --> nadd_le x M') --> nadd_le M M'))" +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: "ALL x::nadd. - EX (xa::nat) N::nat. ALL n::nat. <= N n --> <= (dest_nadd x n) (xa * n)" +lemma NADD_UBOUND: "EX xa N. ALL n>=N. dest_nadd x n <= xa * n" by (import hollight NADD_UBOUND) -lemma NADD_NONZERO: "ALL x::nadd. - ~ nadd_eq x (nadd_of_num 0) --> - (EX N::nat. ALL n::nat. <= N n --> dest_nadd x n ~= 0)" +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: "ALL x::nadd. - ~ nadd_eq x (nadd_of_num 0) --> - (EX (A::nat) N::nat. ALL n::nat. <= N n --> <= n (A * dest_nadd x n))" +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::nadd) n::nat. DIV (n * n) (dest_nadd u n)" - -lemma DEF_nadd_rinv: "nadd_rinv = (%(u::nadd) n::nat. DIV (n * n) (dest_nadd u n))" +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: "ALL x::nadd. - ~ nadd_eq x (nadd_of_num 0) --> - (EX (xa::nat) B::nat. ALL i::nat. <= (nadd_rinv x i) (xa * i + B))" +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: "ALL (x::nadd) n::nat. - dest_nadd x n ~= 0 --> - <= (dist (dest_nadd x n * nadd_rinv x n, n * n)) (dest_nadd x n)" +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: "ALL x::nadd. - ~ nadd_eq x (nadd_of_num 0) --> - (EX N::nat. - ALL n::nat. - <= N n --> - <= (dist (dest_nadd x n * nadd_rinv x n, n * n)) (dest_nadd x n))" +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: "ALL x::nadd. - ~ nadd_eq x (nadd_of_num 0) --> - (EX N::nat. - ALL (m::nat) n::nat. - <= N n --> - <= (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)))" +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: "ALL x::nadd. - ~ nadd_eq x (nadd_of_num 0) --> - (EX N::nat. - ALL (m::nat) n::nat. - <= N m & <= N n --> - <= (dest_nadd x m * dest_nadd x n * - dist (m * nadd_rinv x n, n * nadd_rinv x m)) - (m * n * dist (m * dest_nadd x n, n * dest_nadd x m) + - dest_nadd x m * dest_nadd x n * (m + n)))" +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: "ALL x::nadd. - ~ nadd_eq x (nadd_of_num 0) --> - (EX (B::nat) N::nat. - ALL (m::nat) n::nat. - <= N m & <= N n --> - <= (dest_nadd x m * dest_nadd x n * - dist (m * nadd_rinv x n, n * nadd_rinv x m)) - (B * (m * n * (m + n))))" +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: "ALL x::nadd. - ~ nadd_eq x (nadd_of_num 0) --> - (EX (B::nat) N::nat. - ALL (m::nat) n::nat. - <= N m & <= N n --> - <= (m * n * dist (m * nadd_rinv x n, n * nadd_rinv x m)) - (B * (m * n * (m + n))))" +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: "ALL x::nadd. - ~ nadd_eq x (nadd_of_num 0) --> - (EX (B::nat) N::nat. - ALL (m::nat) n::nat. - <= N m & <= N n --> - <= (dist (m * nadd_rinv x n, n * nadd_rinv x m)) (B * (m + n)))" +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: "ALL x::nadd. - ~ nadd_eq x (nadd_of_num 0) --> - (ALL N::nat. - EX (A::nat) B::nat. - ALL (m::nat) n::nat. - <= m N --> - <= (dist (m * nadd_rinv x n, n * nadd_rinv x m)) (A * n + B))" +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: "ALL x::nadd. - ~ nadd_eq x (nadd_of_num 0) --> - (EX B::nat. - ALL (m::nat) n::nat. - <= (dist (m * nadd_rinv x n, n * nadd_rinv x m)) (B * (m + n)))" +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 +definition + nadd_inv :: "nadd => nadd" where "nadd_inv == -%u::nadd. - COND (nadd_eq u (nadd_of_num 0)) (nadd_of_num 0) (mk_nadd (nadd_rinv u))" +%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::nadd. - COND (nadd_eq u (nadd_of_num 0)) (nadd_of_num 0) - (mk_nadd (nadd_rinv u)))" +(%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: "ALL x::nadd. - dest_nadd (nadd_inv x) = - COND (nadd_eq x (nadd_of_num 0)) (%n::nat. 0) (nadd_rinv x)" +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: "ALL x::nadd. - ~ nadd_eq x (nadd_of_num 0) --> - nadd_eq (nadd_mul (nadd_inv x) x) (nadd_of_num (NUMERAL_BIT1 0))" +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: "ALL (x::nadd) y::nadd. nadd_eq x y --> nadd_eq (nadd_inv x) (nadd_inv y)" +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::nadd => bool. EX x::nadd. s = nadd_eq x}" morphisms "dest_hreal" "mk_hreal" - apply (rule light_ex_imp_nonempty[where t="nadd_eq (x::nadd)"]) +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 @@ -3500,328 +1746,285 @@ [where a="a :: hreal" and r=r , OF type_definition_hreal] -definition hreal_of_num :: "nat => hreal" where - "hreal_of_num == %m::nat. mk_hreal (nadd_eq (nadd_of_num m))" - -lemma DEF_hreal_of_num: "hreal_of_num = (%m::nat. mk_hreal (nadd_eq (nadd_of_num m)))" +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 +definition + hreal_add :: "hreal => hreal => hreal" where "hreal_add == -%(x::hreal) y::hreal. - mk_hreal - (%u::nadd. - EX (xa::nadd) ya::nadd. - nadd_eq (nadd_add xa ya) u & dest_hreal x xa & dest_hreal y ya)" +%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::hreal) y::hreal. - mk_hreal - (%u::nadd. - EX (xa::nadd) ya::nadd. - nadd_eq (nadd_add xa ya) u & dest_hreal x xa & dest_hreal y ya))" +(%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 +definition + hreal_mul :: "hreal => hreal => hreal" where "hreal_mul == -%(x::hreal) y::hreal. - mk_hreal - (%u::nadd. - EX (xa::nadd) ya::nadd. - nadd_eq (nadd_mul xa ya) u & dest_hreal x xa & dest_hreal y ya)" +%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::hreal) y::hreal. - mk_hreal - (%u::nadd. - EX (xa::nadd) ya::nadd. - nadd_eq (nadd_mul xa ya) u & dest_hreal x xa & dest_hreal y ya))" +(%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 +definition + hreal_le :: "hreal => hreal => bool" where "hreal_le == -%(x::hreal) y::hreal. - SOME u::bool. - EX (xa::nadd) ya::nadd. - nadd_le xa ya = u & dest_hreal x xa & dest_hreal y ya" +%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::hreal) y::hreal. - SOME u::bool. - EX (xa::nadd) ya::nadd. - nadd_le xa ya = u & dest_hreal x xa & dest_hreal y ya)" +(%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 +definition + hreal_inv :: "hreal => hreal" where "hreal_inv == -%x::hreal. - mk_hreal - (%u::nadd. EX xa::nadd. nadd_eq (nadd_inv xa) u & dest_hreal x xa)" +%x. mk_hreal (%u. EX xa. nadd_eq (nadd_inv xa) u & dest_hreal x xa)" lemma DEF_hreal_inv: "hreal_inv = -(%x::hreal. - mk_hreal - (%u::nadd. EX xa::nadd. nadd_eq (nadd_inv xa) u & dest_hreal x xa))" +(%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: "ALL (m::hreal) n::hreal. hreal_le m n = (EX d::hreal. n = hreal_add m d)" +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: "ALL (m::hreal) (n::hreal) p::hreal. - (hreal_add m n = hreal_add m p) = (n = p)" +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: "ALL (x::hreal) (xa::hreal) xb::hreal. - (hreal_add x xb = hreal_add xa xb) = (x = xa)" +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: "ALL (x::hreal) (xa::hreal) xb::hreal. - hreal_le (hreal_add x xa) (hreal_add x xb) = hreal_le xa xb" +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: "ALL (x::hreal) (xa::hreal) xb::hreal. - hreal_le (hreal_add x xb) (hreal_add xa xb) = hreal_le x xa" +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: "ALL x::hreal. hreal_add x (hreal_of_num 0) = x" +lemma HREAL_ADD_RID: "hreal_add x (hreal_of_num 0) = x" by (import hollight HREAL_ADD_RID) -lemma HREAL_ADD_RDISTRIB: "ALL (x::hreal) (xa::hreal) xb::hreal. - hreal_mul (hreal_add x xa) xb = - hreal_add (hreal_mul x xb) (hreal_mul xa xb)" +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: "ALL m::hreal. hreal_mul (hreal_of_num 0) m = hreal_of_num 0" +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: "ALL x::hreal. hreal_mul x (hreal_of_num 0) = hreal_of_num 0" +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::hreal) (n::hreal) = hreal_add n m & -hreal_add (hreal_add m n) (p::hreal) = hreal_add m (hreal_add n p) & +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: "ALL (a::hreal) (b::hreal) (c::hreal) d::hreal. - hreal_le a b & hreal_le c d --> hreal_le (hreal_add a c) (hreal_add b d)" +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: "ALL (a::hreal) (b::hreal) c::hreal. - hreal_le a b --> hreal_le (hreal_mul a c) (hreal_mul b c)" +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::nat. (hreal_of_num u, hreal_of_num 0)" - -lemma DEF_treal_of_num: "treal_of_num = (%u::nat. (hreal_of_num u, hreal_of_num 0))" +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::hreal * hreal. (snd u, fst u)" - -lemma DEF_treal_neg: "treal_neg = (%u::hreal * hreal. (snd u, fst u))" +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::hreal * hreal) ua::hreal * hreal. - (hreal_add (fst u) (fst ua), hreal_add (snd u) (snd ua))" +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::hreal * hreal) ua::hreal * hreal. - (hreal_add (fst u) (fst ua), hreal_add (snd u) (snd ua)))" +(%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 +definition + treal_mul :: "hreal * hreal => hreal * hreal => hreal * hreal" where "treal_mul == -%(u::hreal * hreal) ua::hreal * hreal. +%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::hreal * hreal) ua::hreal * hreal. +(%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 +definition + treal_le :: "hreal * hreal => hreal * hreal => bool" where "treal_le == -%(u::hreal * hreal) ua::hreal * hreal. - hreal_le (hreal_add (fst u) (snd ua)) (hreal_add (fst ua) (snd u))" +%u ua. hreal_le (hreal_add (fst u) (snd ua)) (hreal_add (fst ua) (snd u))" lemma DEF_treal_le: "treal_le = -(%(u::hreal * hreal) ua::hreal * hreal. - hreal_le (hreal_add (fst u) (snd ua)) (hreal_add (fst ua) (snd u)))" +(%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 +definition + treal_inv :: "hreal * hreal => hreal * hreal" where "treal_inv == -%u::hreal * hreal. - COND (fst u = snd u) (hreal_of_num 0, hreal_of_num 0) - (COND (hreal_le (snd u) (fst u)) - (hreal_inv (SOME d::hreal. fst u = hreal_add (snd u) d), - hreal_of_num 0) - (hreal_of_num 0, - hreal_inv (SOME d::hreal. snd u = hreal_add (fst u) d)))" +%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::hreal * hreal. - COND (fst u = snd u) (hreal_of_num 0, hreal_of_num 0) - (COND (hreal_le (snd u) (fst u)) - (hreal_inv (SOME d::hreal. fst u = hreal_add (snd u) d), - hreal_of_num 0) - (hreal_of_num 0, - hreal_inv (SOME d::hreal. snd u = hreal_add (fst u) d))))" +(%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::hreal * hreal) ua::hreal * hreal. - hreal_add (fst u) (snd ua) = hreal_add (fst ua) (snd u)" - -lemma DEF_treal_eq: "treal_eq = -(%(u::hreal * hreal) ua::hreal * hreal. - hreal_add (fst u) (snd ua) = hreal_add (fst ua) (snd u))" +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: "ALL x::hreal * hreal. treal_eq x x" +lemma TREAL_EQ_REFL: "treal_eq x x" by (import hollight TREAL_EQ_REFL) -lemma TREAL_EQ_SYM: "ALL (x::hreal * hreal) y::hreal * hreal. treal_eq x y = treal_eq y x" +lemma TREAL_EQ_SYM: "treal_eq x y = treal_eq y x" by (import hollight TREAL_EQ_SYM) -lemma TREAL_EQ_TRANS: "ALL (x::hreal * hreal) (y::hreal * hreal) z::hreal * hreal. - treal_eq x y & treal_eq y z --> treal_eq x z" +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: "ALL (x::hreal * hreal) y::hreal * hreal. x = y --> treal_eq x y" +lemma TREAL_EQ_AP: "x = xa ==> treal_eq x xa" by (import hollight TREAL_EQ_AP) -lemma TREAL_OF_NUM_EQ: "ALL (x::nat) xa::nat. treal_eq (treal_of_num x) (treal_of_num xa) = (x = xa)" +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: "ALL (x::nat) xa::nat. treal_le (treal_of_num x) (treal_of_num xa) = <= x xa" +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: "ALL (x::nat) xa::nat. - treal_eq (treal_add (treal_of_num x) (treal_of_num xa)) - (treal_of_num (x + xa))" +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: "ALL (x::nat) xa::nat. - treal_eq (treal_mul (treal_of_num x) (treal_of_num xa)) - (treal_of_num (x * xa))" +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: "ALL (x::hreal * hreal) y::hreal * hreal. treal_add x y = treal_add y x" +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: "ALL (x::hreal * hreal) y::hreal * hreal. treal_mul x y = treal_mul y x" +lemma TREAL_MUL_SYM_EQ: "treal_mul x y = treal_mul y x" by (import hollight TREAL_MUL_SYM_EQ) -lemma TREAL_ADD_SYM: "ALL (x::hreal * hreal) y::hreal * hreal. - treal_eq (treal_add x y) (treal_add y x)" +lemma TREAL_ADD_SYM: "treal_eq (treal_add x y) (treal_add y x)" by (import hollight TREAL_ADD_SYM) -lemma TREAL_ADD_ASSOC: "ALL (x::hreal * hreal) (y::hreal * hreal) z::hreal * hreal. - treal_eq (treal_add x (treal_add y z)) (treal_add (treal_add x y) z)" +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: "ALL x::hreal * hreal. treal_eq (treal_add (treal_of_num 0) x) x" +lemma TREAL_ADD_LID: "treal_eq (treal_add (treal_of_num 0) x) x" by (import hollight TREAL_ADD_LID) -lemma TREAL_ADD_LINV: "ALL x::hreal * hreal. treal_eq (treal_add (treal_neg x) x) (treal_of_num 0)" +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: "ALL (x::hreal * hreal) y::hreal * hreal. - treal_eq (treal_mul x y) (treal_mul y x)" +lemma TREAL_MUL_SYM: "treal_eq (treal_mul x xa) (treal_mul xa x)" by (import hollight TREAL_MUL_SYM) -lemma TREAL_MUL_ASSOC: "ALL (x::hreal * hreal) (y::hreal * hreal) z::hreal * hreal. - treal_eq (treal_mul x (treal_mul y z)) (treal_mul (treal_mul x y) z)" +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: "ALL x::hreal * hreal. - treal_eq (treal_mul (treal_of_num (NUMERAL_BIT1 0)) x) x" +lemma TREAL_MUL_LID: "treal_eq (treal_mul (treal_of_num 1) x) x" by (import hollight TREAL_MUL_LID) -lemma TREAL_ADD_LDISTRIB: "ALL (x::hreal * hreal) (y::hreal * hreal) z::hreal * hreal. - treal_eq (treal_mul x (treal_add y z)) - (treal_add (treal_mul x y) (treal_mul x z))" +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: "ALL x::hreal * hreal. treal_le x x" +lemma TREAL_LE_REFL: "treal_le x x" by (import hollight TREAL_LE_REFL) -lemma TREAL_LE_ANTISYM: "ALL (x::hreal * hreal) y::hreal * hreal. - (treal_le x y & treal_le y x) = treal_eq x y" +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: "ALL (x::hreal * hreal) (y::hreal * hreal) z::hreal * hreal. - treal_le x y & treal_le y z --> treal_le x z" +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: "ALL (x::hreal * hreal) y::hreal * hreal. treal_le x y | treal_le y x" +lemma TREAL_LE_TOTAL: "treal_le x y | treal_le y x" by (import hollight TREAL_LE_TOTAL) -lemma TREAL_LE_LADD_IMP: "ALL (x::hreal * hreal) (y::hreal * hreal) z::hreal * hreal. - treal_le y z --> treal_le (treal_add x y) (treal_add x z)" +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: "ALL (x::hreal * hreal) y::hreal * hreal. - treal_le (treal_of_num 0) x & treal_le (treal_of_num 0) y --> - treal_le (treal_of_num 0) (treal_mul x y)" +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: "ALL x::hreal * hreal. - ~ treal_eq x (treal_of_num 0) --> - treal_eq (treal_mul (treal_inv x) x) (treal_of_num (NUMERAL_BIT1 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: "ALL (m::nat) n::nat. m = n --> treal_eq (treal_of_num m) (treal_of_num n)" +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: "ALL (x1::hreal * hreal) x2::hreal * hreal. - treal_eq x1 x2 --> treal_eq (treal_neg x1) (treal_neg x2)" +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: "ALL (x1::hreal * hreal) (x2::hreal * hreal) y::hreal * hreal. - treal_eq x1 x2 --> treal_eq (treal_add x1 y) (treal_add x2 y)" +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: "ALL (x1::hreal * hreal) (x2::hreal * hreal) (y1::hreal * hreal) - y2::hreal * hreal. - treal_eq x1 x2 & treal_eq y1 y2 --> - treal_eq (treal_add x1 y1) (treal_add x2 y2)" +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: "ALL (x1::hreal * hreal) (x2::hreal * hreal) y::hreal * hreal. - treal_eq x1 x2 --> treal_eq (treal_mul x1 y) (treal_mul x2 y)" +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: "ALL (x1::hreal * hreal) (x2::hreal * hreal) (y1::hreal * hreal) - y2::hreal * hreal. - treal_eq x1 x2 & treal_eq y1 y2 --> - treal_eq (treal_mul x1 y1) (treal_mul x2 y2)" +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: "ALL (x::hreal * hreal) y::hreal * hreal. treal_eq x y --> treal_le x y" +lemma TREAL_EQ_IMP_LE: "treal_eq x y ==> treal_le x y" by (import hollight TREAL_EQ_IMP_LE) -lemma TREAL_LE_WELLDEF: "ALL (x1::hreal * hreal) (x2::hreal * hreal) (y1::hreal * hreal) - y2::hreal * hreal. - treal_eq x1 x2 & treal_eq y1 y2 --> treal_le x1 y1 = treal_le x2 y2" +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: "ALL (x::hreal * hreal) y::hreal * hreal. - treal_eq x y --> treal_eq (treal_inv x) (treal_inv y)" +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::hreal * hreal => bool. EX x::hreal * hreal. s = treal_eq x}" morphisms "dest_real" "mk_real" - apply (rule light_ex_imp_nonempty[where t="treal_eq (x::hreal * hreal)"]) +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 @@ -3834,3079 +2037,4606 @@ [where a="a :: hollight.real" and r=r , OF type_definition_real] -definition real_of_num :: "nat => hollight.real" where - "real_of_num == %m::nat. mk_real (treal_eq (treal_of_num m))" - -lemma DEF_real_of_num: "real_of_num = (%m::nat. mk_real (treal_eq (treal_of_num m)))" +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 +definition + real_neg :: "hollight.real => hollight.real" where "real_neg == -%x1::hollight.real. - mk_real - (%u::hreal * hreal. - EX x1a::hreal * hreal. - treal_eq (treal_neg x1a) u & dest_real x1 x1a)" +%x1. mk_real (%u. EX x1a. treal_eq (treal_neg x1a) u & dest_real x1 x1a)" lemma DEF_real_neg: "real_neg = -(%x1::hollight.real. - mk_real - (%u::hreal * hreal. - EX x1a::hreal * hreal. - treal_eq (treal_neg x1a) u & dest_real x1 x1a))" +(%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 +definition + real_add :: "hollight.real => hollight.real => hollight.real" where "real_add == -%(x1::hollight.real) y1::hollight.real. +%x1 y1. mk_real - (%u::hreal * hreal. - EX (x1a::hreal * hreal) y1a::hreal * hreal. - treal_eq (treal_add x1a y1a) u & - dest_real x1 x1a & dest_real y1 y1a)" + (%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::hollight.real) y1::hollight.real. +(%x1 y1. mk_real - (%u::hreal * hreal. - EX (x1a::hreal * hreal) y1a::hreal * hreal. - treal_eq (treal_add x1a y1a) u & - dest_real x1 x1a & dest_real y1 y1a))" + (%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 +definition + real_mul :: "hollight.real => hollight.real => hollight.real" where "real_mul == -%(x1::hollight.real) y1::hollight.real. +%x1 y1. mk_real - (%u::hreal * hreal. - EX (x1a::hreal * hreal) y1a::hreal * hreal. - treal_eq (treal_mul x1a y1a) u & - dest_real x1 x1a & dest_real y1 y1a)" + (%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::hollight.real) y1::hollight.real. +(%x1 y1. mk_real - (%u::hreal * hreal. - EX (x1a::hreal * hreal) y1a::hreal * hreal. - treal_eq (treal_mul x1a y1a) u & - dest_real x1 x1a & dest_real y1 y1a))" + (%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 +definition + real_le :: "hollight.real => hollight.real => bool" where "real_le == -%(x1::hollight.real) y1::hollight.real. - SOME u::bool. - EX (x1a::hreal * hreal) y1a::hreal * hreal. - treal_le x1a y1a = u & dest_real x1 x1a & dest_real y1 y1a" +%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::hollight.real) y1::hollight.real. - SOME u::bool. - EX (x1a::hreal * hreal) y1a::hreal * hreal. +(%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 +definition + real_inv :: "hollight.real => hollight.real" where "real_inv == -%x::hollight.real. - mk_real - (%u::hreal * hreal. - EX xa::hreal * hreal. treal_eq (treal_inv xa) u & dest_real x xa)" +%x. mk_real (%u. EX xa. treal_eq (treal_inv xa) u & dest_real x xa)" lemma DEF_real_inv: "real_inv = -(%x::hollight.real. - mk_real - (%u::hreal * hreal. - EX xa::hreal * hreal. treal_eq (treal_inv xa) u & dest_real x xa))" +(%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::hollight.real) ua::hollight.real. real_add u (real_neg ua)" - -lemma DEF_real_sub: "real_sub = (%(u::hollight.real) ua::hollight.real. real_add u (real_neg ua))" +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::hollight.real) ua::hollight.real. ~ real_le ua u" - -lemma DEF_real_lt: "real_lt = (%(u::hollight.real) ua::hollight.real. ~ real_le ua u)" +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) -consts - real_ge :: "hollight.real => hollight.real => bool" - -defs - real_ge_def: "hollight.real_ge == %(u::hollight.real) ua::hollight.real. real_le ua u" - -lemma DEF_real_ge: "hollight.real_ge = (%(u::hollight.real) ua::hollight.real. real_le ua u)" +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) -consts - real_gt :: "hollight.real => hollight.real => bool" - -defs - real_gt_def: "hollight.real_gt == %(u::hollight.real) ua::hollight.real. real_lt ua u" - -lemma DEF_real_gt: "hollight.real_gt = (%(u::hollight.real) ua::hollight.real. real_lt ua u)" +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::hollight.real. COND (real_le (real_of_num 0) u) u (real_neg u)" - -lemma DEF_real_abs: "real_abs = -(%u::hollight.real. COND (real_le (real_of_num 0) u) u (real_neg u))" +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 +definition + real_pow :: "hollight.real => nat => hollight.real" where "real_pow == -SOME real_pow::hollight.real => nat => hollight.real. - (ALL x::hollight.real. real_pow x 0 = real_of_num (NUMERAL_BIT1 0)) & - (ALL (x::hollight.real) n::nat. - real_pow x (Suc n) = real_mul x (real_pow x n))" +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::hollight.real => nat => hollight.real. - (ALL x::hollight.real. real_pow x 0 = real_of_num (NUMERAL_BIT1 0)) & - (ALL (x::hollight.real) n::nat. - real_pow x (Suc n) = real_mul x (real_pow x n)))" +(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::hollight.real) ua::hollight.real. real_mul u (real_inv ua)" - -lemma DEF_real_div: "real_div = (%(u::hollight.real) ua::hollight.real. real_mul u (real_inv ua))" +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::hollight.real) ua::hollight.real. COND (real_le u ua) ua u" - -lemma DEF_real_max: "real_max = (%(u::hollight.real) ua::hollight.real. COND (real_le u ua) ua u)" +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::hollight.real) ua::hollight.real. COND (real_le u ua) u ua" - -lemma DEF_real_min: "real_min = (%(u::hollight.real) ua::hollight.real. COND (real_le u ua) u ua)" +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::hreal => hollight.real. - (ALL xa::hollight.real. - real_le (real_of_num 0) xa = (EX y::hreal. xa = x y)) & - (ALL (y::hreal) z::hreal. hreal_le y z = real_le (x y) (x z))" +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::hollight.real => hreal) r::hreal => hollight.real. - (ALL xa::hreal. x (r xa) = xa) & - (ALL xa::hollight.real. real_le (real_of_num 0) xa --> r (x xa) = xa) & - (ALL x::hreal. real_le (real_of_num 0) (r x)) & - (ALL (x::hreal) y::hreal. hreal_le x y = real_le (r x) (r y))" +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: "ALL P::hollight.real => bool. - (EX x::hollight.real. P x & real_le (real_of_num 0) x) & - (EX M::hollight.real. ALL x::hollight.real. P x --> real_le x M) --> - (EX M::hollight.real. - (ALL x::hollight.real. P x --> real_le x M) & - (ALL M'::hollight.real. - (ALL x::hollight.real. P x --> real_le x M') --> real_le M M'))" +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: "ALL P::hollight.real => bool. - Ex P & - (EX M::hollight.real. ALL x::hollight.real. P x --> real_le x M) --> - (EX M::hollight.real. - (ALL x::hollight.real. P x --> real_le x M) & - (ALL M'::hollight.real. - (ALL x::hollight.real. P x --> real_le x M') --> real_le M M'))" +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::hollight.real) (n::hollight.real) = real_add n m & -real_add (real_add m n) (p::hollight.real) = real_add m (real_add n p) & +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: "ALL x::hollight.real. real_add x (real_neg x) = real_of_num 0" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - (real_add x y = real_add x z) = (y = z)" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - (real_add x z = real_add y z) = (x = y)" +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: "ALL x::hollight.real. real_mul x (real_of_num 0) = real_of_num 0" +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: "ALL x::hollight.real. real_mul (real_of_num 0) x = real_of_num 0" +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: "ALL x::hollight.real. real_neg (real_neg x) = x" +lemma REAL_NEG_NEG: "real_neg (real_neg x) = x" by (import hollight REAL_NEG_NEG) -lemma REAL_MUL_RNEG: "ALL (x::hollight.real) y::hollight.real. - real_mul x (real_neg y) = real_neg (real_mul x y)" +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: "ALL (x::hollight.real) y::hollight.real. - real_mul (real_neg x) y = real_neg (real_mul x y)" +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: "ALL (x::hollight.real) y::hollight.real. - real_neg (real_add x y) = real_add (real_neg x) (real_neg y)" +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: "ALL x::hollight.real. real_add x (real_of_num 0) = x" +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: "ALL (x::hollight.real) y::hollight.real. - real_le (real_neg x) y = real_le (real_of_num 0) (real_add x y)" +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: "ALL (x::hollight.real) y::hollight.real. - real_le (real_neg x) (real_neg y) = real_le y x" +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: "ALL (x::hollight.real) y::hollight.real. - real_le x (real_neg y) = real_le (real_add x y) (real_of_num 0)" +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: "ALL (x::nat) n::nat. real_pow (real_of_num x) n = real_of_num (EXP x n)" +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: "ALL (x::hollight.real) n::nat. - real_pow (real_neg x) n = - COND (EVEN n) (real_pow x n) (real_neg (real_pow x n))" +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: "ALL x::nat. real_abs (real_of_num x) = real_of_num x" +lemma REAL_ABS_NUM: "real_abs (real_of_num x) = real_of_num x" by (import hollight REAL_ABS_NUM) -lemma REAL_ABS_NEG: "ALL x::hollight.real. real_abs (real_neg x) = real_abs x" +lemma REAL_ABS_NEG: "real_abs (real_neg x) = real_abs x" by (import hollight REAL_ABS_NEG) -lemma REAL_LTE_TOTAL: "ALL (x::hollight.real) xa::hollight.real. real_lt x xa | real_le xa x" +lemma REAL_LTE_TOTAL: "real_lt x xa | real_le xa x" by (import hollight REAL_LTE_TOTAL) -lemma REAL_LET_TOTAL: "ALL (x::hollight.real) xa::hollight.real. real_le x xa | real_lt xa x" +lemma REAL_LET_TOTAL: "real_le x xa | real_lt xa x" by (import hollight REAL_LET_TOTAL) -lemma REAL_LT_IMP_LE: "ALL (x::hollight.real) y::hollight.real. real_lt x y --> real_le x y" +lemma REAL_LT_IMP_LE: "real_lt x y ==> real_le x y" by (import hollight REAL_LT_IMP_LE) -lemma REAL_LTE_TRANS: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_lt x y & real_le y z --> real_lt x z" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_le x y & real_lt y z --> real_lt x z" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_lt x y & real_lt y z --> real_lt x z" +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: "ALL (x::hollight.real) y::hollight.real. - real_le (real_of_num 0) x & real_le (real_of_num 0) y --> - real_le (real_of_num 0) (real_add x y)" +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: "ALL (x::hollight.real) y::hollight.real. ~ (real_lt x y & real_le y x)" +lemma REAL_LTE_ANTISYM: "~ (real_lt x y & real_le y x)" by (import hollight REAL_LTE_ANTISYM) -lemma REAL_SUB_LE: "ALL (x::hollight.real) xa::hollight.real. - real_le (real_of_num 0) (real_sub x xa) = real_le xa x" +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: "ALL (x::hollight.real) xa::hollight.real. - real_neg (real_sub x xa) = real_sub xa x" +lemma REAL_NEG_SUB: "real_neg (real_sub x xa) = real_sub xa x" by (import hollight REAL_NEG_SUB) -lemma REAL_LE_LT: "ALL (x::hollight.real) xa::hollight.real. - real_le x xa = (real_lt x xa | x = xa)" +lemma REAL_LE_LT: "real_le x xa = (real_lt x xa | x = xa)" by (import hollight REAL_LE_LT) -lemma REAL_SUB_LT: "ALL (x::hollight.real) xa::hollight.real. - real_lt (real_of_num 0) (real_sub x xa) = real_lt xa x" +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: "ALL (x::hollight.real) xa::hollight.real. (~ real_lt x xa) = real_le xa x" +lemma REAL_NOT_LT: "(~ real_lt x xa) = real_le xa x" by (import hollight REAL_NOT_LT) -lemma REAL_SUB_0: "ALL (x::hollight.real) y::hollight.real. - (real_sub x y = real_of_num 0) = (x = y)" +lemma REAL_SUB_0: "(real_sub x y = real_of_num 0) = (x = y)" by (import hollight REAL_SUB_0) -lemma REAL_LT_LE: "ALL (x::hollight.real) y::hollight.real. - real_lt x y = (real_le x y & x ~= y)" +lemma REAL_LT_LE: "real_lt x y = (real_le x y & x ~= y)" by (import hollight REAL_LT_LE) -lemma REAL_LT_REFL: "ALL x::hollight.real. ~ real_lt x x" +lemma REAL_LT_REFL: "~ real_lt x x" by (import hollight REAL_LT_REFL) -lemma REAL_LTE_ADD: "ALL (x::hollight.real) y::hollight.real. - real_lt (real_of_num 0) x & real_le (real_of_num 0) y --> - real_lt (real_of_num 0) (real_add x y)" +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: "ALL (x::hollight.real) y::hollight.real. - real_le (real_of_num 0) x & real_lt (real_of_num 0) y --> - real_lt (real_of_num 0) (real_add x y)" +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: "ALL (x::hollight.real) y::hollight.real. - real_lt (real_of_num 0) x & real_lt (real_of_num 0) y --> - real_lt (real_of_num 0) (real_add x y)" +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: "ALL (x::hollight.real) y::hollight.real. - (real_mul x y = real_of_num 0) = (x = real_of_num 0 | y = real_of_num 0)" +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: "ALL x::hollight.real. - real_le (real_of_num 0) x | real_le (real_of_num 0) (real_neg x)" +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: "ALL x::hollight.real. real_le (real_of_num 0) (real_mul x x)" +lemma REAL_LE_SQUARE: "real_le (real_of_num 0) (real_mul x x)" by (import hollight REAL_LE_SQUARE) -lemma REAL_MUL_RID: "ALL x::hollight.real. real_mul x (real_of_num (NUMERAL_BIT1 0)) = x" +lemma REAL_MUL_RID: "real_mul x (real_of_num 1) = x" by (import hollight REAL_MUL_RID) -lemma REAL_POW_2: "ALL x::hollight.real. - real_pow x (NUMERAL_BIT0 (NUMERAL_BIT1 0)) = real_mul x x" +lemma REAL_POW_2: "real_pow x 2 = real_mul x x" by (import hollight REAL_POW_2) -lemma REAL_POLY_CLAUSES: "(ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_add x (real_add y z) = real_add (real_add x y) z) & -(ALL (x::hollight.real) y::hollight.real. real_add x y = real_add y x) & -(ALL x::hollight.real. real_add (real_of_num 0) x = x) & -(ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_mul x (real_mul y z) = real_mul (real_mul x y) z) & -(ALL (x::hollight.real) y::hollight.real. real_mul x y = real_mul y x) & -(ALL x::hollight.real. real_mul (real_of_num (NUMERAL_BIT1 0)) x = x) & -(ALL x::hollight.real. real_mul (real_of_num 0) x = real_of_num 0) & -(ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real. +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::hollight.real. real_pow x 0 = real_of_num (NUMERAL_BIT1 0)) & -(ALL (x::hollight.real) xa::nat. - real_pow x (Suc xa) = real_mul x (real_pow x xa))" +(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::hollight.real. - real_neg x = real_mul (real_neg (real_of_num (NUMERAL_BIT1 0))) x) & -(ALL (x::hollight.real) xa::hollight.real. - real_sub x xa = - real_add x (real_mul (real_neg (real_of_num (NUMERAL_BIT1 0))) xa))" +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: "ALL x::nat. real_le (real_of_num 0) (real_of_num x)" +lemma REAL_POS: "real_le (real_of_num 0) (real_of_num x)" by (import hollight REAL_POS) -lemma REAL_OF_NUM_LT: "ALL (x::nat) xa::nat. real_lt (real_of_num x) (real_of_num xa) = < x xa" +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: "ALL (x::nat) xa::nat. - hollight.real_ge (real_of_num x) (real_of_num xa) = >= x xa" +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: "ALL (x::nat) xa::nat. - hollight.real_gt (real_of_num x) (real_of_num xa) = > x xa" +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_SUC: "ALL x::nat. - real_add (real_of_num x) (real_of_num (NUMERAL_BIT1 0)) = - real_of_num (Suc x)" +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: "ALL (m::nat) n::nat. - <= m n --> real_sub (real_of_num n) (real_of_num m) = real_of_num (n - m)" +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::hollight.real) (n::hollight.real) = real_mul n m & -real_mul (real_mul m n) (p::hollight.real) = real_mul m (real_mul n p) & +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_mul (real_add x y) z = real_add (real_mul x z) (real_mul y z)" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_lt y z --> real_lt (real_add x y) (real_add x z)" +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: "ALL (x::hollight.real) y::hollight.real. - real_lt (real_of_num 0) x & real_lt (real_of_num 0) y --> - real_lt (real_of_num 0) (real_mul x y)" +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: "ALL (x::hollight.real) y::hollight.real. - (real_add x y = x) = (y = real_of_num 0)" +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: "ALL (x::hollight.real) y::hollight.real. - (real_add x y = y) = (x = real_of_num 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: "ALL (x::hollight.real) y::hollight.real. - (real_add x y = real_of_num 0) = (x = real_neg y)" +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: "ALL (x::hollight.real) y::hollight.real. - (real_add x y = real_of_num 0) = (y = real_neg x)" +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: "ALL (x::hollight.real) y::hollight.real. - real_neg (real_mul x y) = real_mul (real_neg x) y" +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: "ALL (x::hollight.real) y::hollight.real. - real_neg (real_mul x y) = real_mul x (real_neg y)" +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: "ALL x::hollight.real. real_neg (real_neg x) = x" +lemma REAL_NEGNEG: "real_neg (real_neg x) = x" by (import hollight REAL_NEGNEG) -lemma REAL_NEG_MUL2: "ALL (x::hollight.real) y::hollight.real. - real_mul (real_neg x) (real_neg y) = real_mul x y" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_lt (real_add x y) (real_add x z) = real_lt y z" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_lt (real_add x z) (real_add y z) = real_lt x y" +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: "ALL (x::hollight.real) y::hollight.real. ~ (real_lt x y & real_lt y x)" +lemma REAL_LT_ANTISYM: "~ (real_lt x y & real_lt y x)" by (import hollight REAL_LT_ANTISYM) -lemma REAL_LT_GT: "ALL (x::hollight.real) y::hollight.real. real_lt x y --> ~ real_lt y x" +lemma REAL_LT_GT: "real_lt x y ==> ~ real_lt y x" by (import hollight REAL_LT_GT) -lemma REAL_NOT_EQ: "ALL (x::hollight.real) y::hollight.real. - (x ~= y) = (real_lt x y | real_lt y x)" +lemma REAL_NOT_EQ: "(x ~= y) = (real_lt x y | real_lt y x)" by (import hollight REAL_NOT_EQ) -lemma REAL_LE_TOTAL: "ALL (x::hollight.real) y::hollight.real. real_le x y | real_le y x" - by (import hollight REAL_LE_TOTAL) - -lemma REAL_LE_REFL: "ALL x::hollight.real. real_le x x" - by (import hollight REAL_LE_REFL) - -lemma REAL_LE_ANTISYM: "ALL (x::hollight.real) y::hollight.real. - (real_le x y & real_le y x) = (x = y)" - by (import hollight REAL_LE_ANTISYM) - -lemma REAL_LET_ANTISYM: "ALL (x::hollight.real) y::hollight.real. ~ (real_le x y & real_lt y x)" +lemma REAL_LET_ANTISYM: "~ (real_le x y & real_lt y x)" by (import hollight REAL_LET_ANTISYM) -lemma REAL_NEG_LT0: "ALL x::hollight.real. - real_lt (real_neg x) (real_of_num 0) = real_lt (real_of_num 0) x" +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: "ALL x::hollight.real. - real_lt (real_of_num 0) (real_neg x) = real_lt x (real_of_num 0)" +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: "ALL x::hollight.real. - real_le (real_neg x) (real_of_num 0) = real_le (real_of_num 0) x" +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: "ALL x::hollight.real. - real_le (real_of_num 0) (real_neg x) = real_le x (real_of_num 0)" +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: "ALL (x::hollight.real) y::hollight.real. x = y | real_lt x y | real_lt y x" +lemma REAL_LT_TOTAL: "x = y | real_lt x y | real_lt y x" by (import hollight REAL_LT_TOTAL) -lemma REAL_LT_NEGTOTAL: "ALL x::hollight.real. - x = real_of_num 0 | - real_lt (real_of_num 0) x | real_lt (real_of_num 0) (real_neg x)" +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 (NUMERAL_BIT1 0))" +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 (NUMERAL_BIT1 0))" +lemma REAL_LT_01: "real_lt (real_of_num 0) (real_of_num 1)" by (import hollight REAL_LT_01) -lemma REAL_LE_LADD: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_le (real_add x y) (real_add x z) = real_le y z" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_le (real_add x z) (real_add y z) = real_le x y" +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: "ALL (w::hollight.real) (x::hollight.real) (y::hollight.real) - z::hollight.real. - real_lt w x & real_lt y z --> real_lt (real_add w y) (real_add x z)" +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: "ALL (w::hollight.real) (x::hollight.real) (y::hollight.real) - z::hollight.real. - real_le w x & real_le y z --> real_le (real_add w y) (real_add x z)" +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: "ALL (x::hollight.real) xa::hollight.real. - real_lt (real_neg x) xa = real_lt (real_of_num 0) (real_add x xa)" +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: "ALL (x::hollight.real) xa::hollight.real. - real_lt x (real_neg xa) = real_lt (real_add x xa) (real_of_num 0)" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_lt y (real_add x (real_neg z)) = real_lt (real_add y z) x" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_lt (real_add x (real_neg y)) z = real_lt x (real_add z y)" +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: "ALL (x::hollight.real) y::hollight.real. - real_le x y --> real_lt x (real_add y (real_of_num (NUMERAL_BIT1 0)))" +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: "ALL (x::hollight.real) y::hollight.real. real_add (real_sub x y) y = x" +lemma REAL_SUB_ADD: "real_add (real_sub x y) y = x" by (import hollight REAL_SUB_ADD) -lemma REAL_SUB_ADD2: "ALL (x::hollight.real) y::hollight.real. real_add y (real_sub x y) = x" +lemma REAL_SUB_ADD2: "real_add y (real_sub x y) = x" by (import hollight REAL_SUB_ADD2) -lemma REAL_SUB_REFL: "ALL x::hollight.real. real_sub x x = real_of_num 0" +lemma REAL_SUB_REFL: "real_sub x x = real_of_num 0" by (import hollight REAL_SUB_REFL) -lemma REAL_LE_DOUBLE: "ALL x::hollight.real. - real_le (real_of_num 0) (real_add x x) = real_le (real_of_num 0) x" +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: "ALL x::hollight.real. real_le (real_neg x) x = real_le (real_of_num 0) x" +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: "ALL x::hollight.real. real_le x (real_neg x) = real_le x (real_of_num 0)" +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: "ALL x::hollight.real. (real_neg x = real_of_num 0) = (x = real_of_num 0)" +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: "ALL (x::hollight.real) y::hollight.real. real_sub (real_add x y) x = y" +lemma REAL_ADD_SUB: "real_sub (real_add x y) x = y" by (import hollight REAL_ADD_SUB) -lemma REAL_NEG_EQ: "ALL (x::hollight.real) y::hollight.real. (real_neg x = y) = (x = real_neg y)" +lemma REAL_NEG_EQ: "(real_neg x = y) = (x = real_neg y)" by (import hollight REAL_NEG_EQ) -lemma REAL_NEG_MINUS1: "ALL x::hollight.real. - real_neg x = real_mul (real_neg (real_of_num (NUMERAL_BIT1 0))) x" +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: "ALL (x::hollight.real) y::hollight.real. real_lt x y --> x ~= y" +lemma REAL_LT_IMP_NE: "real_lt x y ==> x ~= y" by (import hollight REAL_LT_IMP_NE) -lemma REAL_LE_ADDR: "ALL (x::hollight.real) y::hollight.real. - real_le x (real_add x y) = real_le (real_of_num 0) y" +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: "ALL (x::hollight.real) y::hollight.real. - real_le y (real_add x y) = real_le (real_of_num 0) x" +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: "ALL (x::hollight.real) y::hollight.real. - real_lt x (real_add x y) = real_lt (real_of_num 0) y" +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: "ALL (x::hollight.real) y::hollight.real. - real_lt y (real_add x y) = real_lt (real_of_num 0) x" +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: "ALL (x::hollight.real) y::hollight.real. - real_sub (real_sub x y) x = real_neg y" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_lt (real_add x y) z = real_lt x (real_sub z y)" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_lt (real_sub x y) z = real_lt x (real_add z y)" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_lt x (real_sub y z) = real_lt (real_add x z) y" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_le x (real_sub y z) = real_le (real_add x z) y" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_le (real_sub x y) z = real_le x (real_add z y)" +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: "ALL (x::hollight.real) y::hollight.real. - real_lt (real_neg x) (real_neg y) = real_lt y x" +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: "ALL (x::hollight.real) y::hollight.real. - real_le (real_neg x) (real_neg y) = real_le y x" +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: "ALL (a::hollight.real) (b::hollight.real) (c::hollight.real) - d::hollight.real. - real_sub (real_add a b) (real_add c d) = - real_add (real_sub a c) (real_sub b d)" +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: "ALL x::hollight.real. real_sub (real_of_num 0) x = real_neg x" +lemma REAL_SUB_LZERO: "real_sub (real_of_num 0) x = real_neg x" by (import hollight REAL_SUB_LZERO) -lemma REAL_SUB_RZERO: "ALL x::hollight.real. real_sub x (real_of_num 0) = x" +lemma REAL_SUB_RZERO: "real_sub x (real_of_num 0) = x" by (import hollight REAL_SUB_RZERO) -lemma REAL_LET_ADD2: "ALL (w::hollight.real) (x::hollight.real) (y::hollight.real) - z::hollight.real. - real_le w x & real_lt y z --> real_lt (real_add w y) (real_add x z)" +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: "ALL (w::hollight.real) (x::hollight.real) (y::hollight.real) - z::hollight.real. - real_lt w x & real_le y z --> real_lt (real_add w y) (real_add x z)" +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: "ALL (x::hollight.real) y::hollight.real. - real_sub (real_neg x) y = real_neg (real_add x y)" +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: "ALL (x::hollight.real) y::hollight.real. - real_sub x (real_neg y) = real_add x y" +lemma REAL_SUB_RNEG: "real_sub x (real_neg y) = real_add x y" by (import hollight REAL_SUB_RNEG) -lemma REAL_SUB_NEG2: "ALL (x::hollight.real) y::hollight.real. - real_sub (real_neg x) (real_neg y) = real_sub y x" +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: "ALL (a::hollight.real) (b::hollight.real) c::hollight.real. - real_add (real_sub a b) (real_sub b c) = real_sub a c" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - (x = real_sub y z) = (real_add x z = y)" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - (real_sub x y = z) = (x = real_add z y)" +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: "ALL (x::hollight.real) y::hollight.real. real_sub x (real_sub x y) = y" +lemma REAL_SUB_SUB2: "real_sub x (real_sub x y) = y" by (import hollight REAL_SUB_SUB2) -lemma REAL_ADD_SUB2: "ALL (x::hollight.real) y::hollight.real. - real_sub x (real_add x y) = real_neg y" +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: "ALL (x::hollight.real) y::hollight.real. x = y --> real_le x y" +lemma REAL_EQ_IMP_LE: "x = y ==> real_le x y" by (import hollight REAL_EQ_IMP_LE) -lemma REAL_POS_NZ: "ALL x::hollight.real. real_lt (real_of_num 0) x --> x ~= real_of_num 0" +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: "ALL (x::hollight.real) y::hollight.real. - real_mul (real_add x y) (real_sub x y) = - real_sub (real_mul x x) (real_mul y y)" +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: "ALL (x::hollight.real) y::hollight.real. (real_neg x = real_neg y) = (x = y)" +lemma REAL_EQ_NEG2: "(real_neg x = real_neg y) = (x = y)" by (import hollight REAL_EQ_NEG2) -lemma REAL_LT_NEG2: "ALL (x::hollight.real) y::hollight.real. - real_lt (real_neg x) (real_neg y) = real_lt y x" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_mul x (real_sub y z) = real_sub (real_mul x y) (real_mul x z)" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_mul (real_sub x y) z = real_sub (real_mul x z) (real_mul y z)" +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: "ALL x::hollight.real. (real_abs x = real_of_num 0) = (x = real_of_num 0)" +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 (NUMERAL_BIT1 0)) = real_of_num (NUMERAL_BIT1 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: "ALL (x::hollight.real) y::hollight.real. - real_le (real_abs (real_add x y)) (real_add (real_abs x) (real_abs y))" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_le (real_add (real_abs x) (real_abs (real_sub y x))) z --> - real_le (real_abs y) z" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_lt (real_add (real_abs x) (real_abs (real_sub y x))) z --> - real_lt (real_abs y) z" +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: "ALL x::hollight.real. real_le (real_of_num 0) (real_abs x)" +lemma REAL_ABS_POS: "real_le (real_of_num 0) (real_abs x)" by (import hollight REAL_ABS_POS) -lemma REAL_ABS_SUB: "ALL (x::hollight.real) y::hollight.real. - real_abs (real_sub x y) = real_abs (real_sub y x)" +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: "ALL x::hollight.real. - (x ~= real_of_num 0) = real_lt (real_of_num 0) (real_abs x)" +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: "ALL x::hollight.real. real_abs (real_abs x) = real_abs x" +lemma REAL_ABS_ABS: "real_abs (real_abs x) = real_abs x" by (import hollight REAL_ABS_ABS) -lemma REAL_ABS_LE: "ALL x::hollight.real. real_le x (real_abs x)" +lemma REAL_ABS_LE: "real_le x (real_abs x)" by (import hollight REAL_ABS_LE) -lemma REAL_ABS_REFL: "ALL x::hollight.real. (real_abs x = x) = real_le (real_of_num 0) x" +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: "ALL (x::hollight.real) (y::hollight.real) d::hollight.real. - (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" +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: "ALL (x::hollight.real) (y::hollight.real) d::hollight.real. - real_lt (real_abs (real_sub x y)) d --> real_lt y (real_add x d)" +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: "ALL (x::hollight.real) y::hollight.real. - real_lt (real_abs (real_sub x y)) (real_abs y) --> x ~= real_of_num 0" +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: "ALL x::hollight.real. - x = real_of_num 0 | real_lt (real_of_num 0) (real_abs x)" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_lt x z & real_lt (real_abs (real_sub y x)) (real_sub z x) --> - real_lt y z" +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: "ALL (x::hollight.real) y::hollight.real. - real_lt (real_abs (real_sub x y)) y --> real_lt (real_of_num 0) x" +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: "ALL (x::hollight.real) y::hollight.real. - real_lt (real_abs (real_sub x y)) (real_neg y) --> - real_lt x (real_of_num 0)" +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: "ALL (x::hollight.real) (y::hollight.real) h::hollight.real. - real_lt (real_abs h) (real_sub (real_abs y) (real_abs x)) --> - real_lt (real_abs (real_add x h)) (real_abs y)" +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: "ALL (x::hollight.real) y::hollight.real. - real_le (real_sub (real_abs x) (real_abs y)) (real_abs (real_sub x y))" +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: "ALL (x::hollight.real) y::hollight.real. - real_le (real_abs (real_sub (real_abs x) (real_abs y))) - (real_abs (real_sub x y))" +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: "ALL (x0::hollight.real) (x::hollight.real) (y0::hollight.real) - y::hollight.real. - real_lt x0 y0 & - real_lt - (real_mul (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 0))) - (real_abs (real_sub x x0))) - (real_sub y0 x0) & - real_lt - (real_mul (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 0))) - (real_abs (real_sub y y0))) - (real_sub y0 x0) --> - real_lt x y" +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: "ALL (x::hollight.real) k::hollight.real. - real_le (real_abs x) k = (real_le (real_neg k) x & real_le x k)" +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_MIN_MAX: "ALL (x::hollight.real) y::hollight.real. - real_min x y = real_neg (real_max (real_neg x) (real_neg y))" +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: "ALL (x::hollight.real) y::hollight.real. - real_max x y = real_neg (real_min (real_neg x) (real_neg y))" +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: "ALL (x::hollight.real) y::hollight.real. - real_le x (real_max x y) & real_le y (real_max x y)" +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: "ALL (x::hollight.real) y::hollight.real. - real_le (real_min x y) x & real_le (real_min x y) y" +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: "ALL (x::hollight.real) y::hollight.real. real_max x y = real_max y x" +lemma REAL_MAX_SYM: "real_max x y = real_max y x" by (import hollight REAL_MAX_SYM) -lemma REAL_MIN_SYM: "ALL (x::hollight.real) y::hollight.real. real_min x y = real_min y x" +lemma REAL_MIN_SYM: "real_min x y = real_min y x" by (import hollight REAL_MIN_SYM) -lemma REAL_LE_MAX: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_le z (real_max x y) = (real_le z x | real_le z y)" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_le z (real_min x y) = (real_le z x & real_le z y)" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_lt z (real_max x y) = (real_lt z x | real_lt z y)" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_lt z (real_min x y) = (real_lt z x & real_lt z y)" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_le (real_max x y) z = (real_le x z & real_le y z)" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_le (real_min x y) z = (real_le x z | real_le y z)" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_lt (real_max x y) z = (real_lt x z & real_lt y z)" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_lt (real_min x y) z = (real_lt x z | real_lt y z)" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_max x (real_max y z) = real_max (real_max x y) z" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_min x (real_min y z) = real_min (real_min x y) z" +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::hollight.real) (y::hollight.real) = real_max y x & -real_max (real_max x y) (z::hollight.real) = real_max x (real_max y z) & +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::hollight.real) (y::hollight.real) = real_min y x & -real_min (real_min x y) (z::hollight.real) = real_min x (real_min y z) & +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: "ALL (x::hollight.real) y::hollight.real. - real_abs (real_mul x y) = real_mul (real_abs x) (real_abs y)" +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: "ALL (x::hollight.real) n::nat. - real_le (real_of_num 0) x --> real_le (real_of_num 0) (real_pow x n)" +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: "ALL (x::hollight.real) n::nat. - real_lt (real_of_num 0) x --> real_lt (real_of_num 0) (real_pow x n)" +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: "ALL (x::hollight.real) n::nat. - real_abs (real_pow x n) = real_pow (real_abs x) n" +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: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real. - real_le (real_of_num 0) x & real_le xa xb --> - real_le (real_mul x xa) (real_mul x xb)" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_le x y & real_le (real_of_num 0) z --> - real_le (real_mul x z) (real_mul y z)" +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: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real. - real_lt (real_of_num 0) x & real_lt xa xb --> - real_lt (real_mul x xa) (real_mul x xb)" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_lt x y & real_lt (real_of_num 0) z --> - real_lt (real_mul x z) (real_mul y z)" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - (real_mul x y = real_mul x z) = (x = real_of_num 0 | y = z)" +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: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real. - (real_mul x xb = real_mul xa xb) = (x = xa | xb = real_of_num 0)" +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: "ALL (x::hollight.real) y::hollight.real. - real_mul x y = real_of_num (NUMERAL_BIT1 0) --> real_inv y = x" +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: "ALL (x::hollight.real) xa::hollight.real. - real_mul x xa = real_of_num (NUMERAL_BIT1 0) --> real_inv x = xa" +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: "ALL x::hollight.real. real_inv (real_inv x) = x" +lemma REAL_INV_INV: "real_inv (real_inv x) = x" by (import hollight REAL_INV_INV) -lemma REAL_INV_EQ_0: "ALL x::hollight.real. (real_inv x = real_of_num 0) = (x = real_of_num 0)" +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: "ALL x::hollight.real. - real_lt (real_of_num 0) x --> real_lt (real_of_num 0) (real_inv x)" +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: "ALL x::hollight.real. - real_lt (real_of_num 0) (real_inv x) = real_lt (real_of_num 0) x" +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: "ALL x::hollight.real. real_inv (real_neg x) = real_neg (real_inv x)" +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: "ALL x::hollight.real. - real_le (real_of_num 0) (real_inv x) = real_le (real_of_num 0) x" +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: "ALL x::hollight.real. - real_le (real_of_num 0) x --> real_le (real_of_num 0) (real_inv x)" +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: "ALL x::hollight.real. - x ~= real_of_num 0 --> - real_mul x (real_inv x) = real_of_num (NUMERAL_BIT1 0)" +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 (NUMERAL_BIT1 0)) = real_of_num (NUMERAL_BIT1 0)" +lemma REAL_INV_1: "real_inv (real_of_num 1) = real_of_num 1" by (import hollight REAL_INV_1) -lemma REAL_DIV_1: "ALL x::hollight.real. real_div x (real_of_num (NUMERAL_BIT1 0)) = x" +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: "ALL x::hollight.real. - x ~= real_of_num 0 --> real_div x x = real_of_num (NUMERAL_BIT1 0)" +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: "ALL (x::hollight.real) xa::hollight.real. - xa ~= real_of_num 0 --> real_mul (real_div x xa) xa = x" +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: "ALL (x::hollight.real) xa::hollight.real. - xa ~= real_of_num 0 --> real_mul xa (real_div x xa) = x" +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: "ALL x::hollight.real. real_abs (real_inv x) = real_inv (real_abs x)" +lemma REAL_ABS_INV: "real_abs (real_inv x) = real_inv (real_abs x)" by (import hollight REAL_ABS_INV) -lemma REAL_ABS_DIV: "ALL (x::hollight.real) xa::hollight.real. - real_abs (real_div x xa) = real_div (real_abs x) (real_abs xa)" +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: "ALL (x::hollight.real) y::hollight.real. - real_inv (real_mul x y) = real_mul (real_inv x) (real_inv y)" +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: "ALL (x::hollight.real) xa::hollight.real. - real_inv (real_div x xa) = real_div xa x" +lemma REAL_INV_DIV: "real_inv (real_div x xa) = real_div xa x" by (import hollight REAL_INV_DIV) -lemma REAL_POW_MUL: "ALL (x::hollight.real) (y::hollight.real) n::nat. - real_pow (real_mul x y) n = real_mul (real_pow x n) (real_pow y n)" +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: "ALL (x::hollight.real) n::nat. - real_pow (real_inv x) n = real_inv (real_pow x n)" +lemma REAL_POW_INV: "real_pow (real_inv x) n = real_inv (real_pow x n)" by (import hollight REAL_POW_INV) -lemma REAL_POW_DIV: "ALL (x::hollight.real) (xa::hollight.real) xb::nat. - real_pow (real_div x xa) xb = real_div (real_pow x xb) (real_pow xa xb)" +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: "ALL (x::hollight.real) (m::nat) n::nat. - real_pow x (m + n) = real_mul (real_pow x m) (real_pow x n)" +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: "ALL (x::hollight.real) n::nat. - x ~= real_of_num 0 --> real_pow x n ~= real_of_num 0" +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: "ALL (x::hollight.real) (m::nat) n::nat. - x ~= real_of_num 0 & <= m n --> - real_pow x (n - m) = real_div (real_pow x n) (real_pow x m)" +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: "ALL x::hollight.real. real_lt (real_of_num 0) x --> x ~= real_of_num 0" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_lt (real_of_num 0) x & real_lt (real_mul x y) (real_mul x z) --> - real_lt y z" +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: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real. - real_lt (real_of_num 0) xb & real_lt (real_mul x xb) (real_mul xa xb) --> - real_lt x xa" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_lt (real_of_num 0) x & real_le (real_mul x y) (real_mul x z) --> - real_le y z" +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: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real. - real_lt (real_of_num 0) xb & real_le (real_mul x xb) (real_mul xa xb) --> - real_le x xa" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_lt (real_of_num 0) z --> - real_le (real_mul x z) (real_mul y z) = real_le x y" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_lt (real_of_num 0) z --> - real_le (real_mul z x) (real_mul z y) = real_le x y" +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: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real. - real_lt (real_of_num 0) xb --> - real_lt (real_mul x xb) (real_mul xa xb) = real_lt x xa" +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: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real. - real_lt (real_of_num 0) xb --> - real_lt (real_mul xb x) (real_mul xb xa) = real_lt x xa" +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_RDIV_EQ: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_lt (real_of_num 0) z --> - real_le x (real_div y z) = real_le (real_mul x z) y" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - real_lt (real_of_num 0) z --> - real_le (real_div x z) y = real_le x (real_mul y z)" +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: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real. - real_lt (real_of_num 0) xb --> - real_lt x (real_div xa xb) = real_lt (real_mul x xb) xa" +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: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real. - real_lt (real_of_num 0) xb --> - real_lt (real_div x xb) xa = real_lt x (real_mul xa xb)" +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: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real. - real_lt (real_of_num 0) xb --> - (x = real_div xa xb) = (real_mul x xb = xa)" +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: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real. - real_lt (real_of_num 0) xb --> - (real_div x xb = xa) = (x = real_mul xa xb)" +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: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real. - real_lt (real_of_num 0) xb --> - real_lt (real_div x xb) (real_div xa xb) = real_lt x xa" +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: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real. - real_lt (real_of_num 0) xb --> - real_le (real_div x xb) (real_div xa xb) = real_le x xa" +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: "ALL x::hollight.real. - real_mul (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 0))) x = real_add x x" +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: "ALL (x::hollight.real) n::nat. - (real_pow x n = real_of_num 0) = (x = real_of_num 0 & n ~= 0)" +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: "ALL (w::hollight.real) (x::hollight.real) (y::hollight.real) - z::hollight.real. - 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)" +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: "ALL (w::hollight.real) (x::hollight.real) (y::hollight.real) - z::hollight.real. - 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)" +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: "ALL x::hollight.real. - real_lt (real_of_num 0) (real_mul x x) = (x ~= real_of_num 0)" +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_INV_LE_1: "ALL x::hollight.real. - real_le (real_of_num (NUMERAL_BIT1 0)) x --> - real_le (real_inv x) (real_of_num (NUMERAL_BIT1 0))" +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_POW_LE_1: "ALL (n::nat) x::hollight.real. - real_le (real_of_num (NUMERAL_BIT1 0)) x --> - real_le (real_of_num (NUMERAL_BIT1 0)) (real_pow x n)" - by (import hollight REAL_POW_LE_1) - -lemma REAL_POW_1_LE: "ALL (n::nat) x::hollight.real. - real_le (real_of_num 0) x & real_le x (real_of_num (NUMERAL_BIT1 0)) --> - real_le (real_pow x n) (real_of_num (NUMERAL_BIT1 0))" - by (import hollight REAL_POW_1_LE) - -lemma REAL_POW_1: "ALL x::hollight.real. real_pow x (NUMERAL_BIT1 0) = x" - by (import hollight REAL_POW_1) - -lemma REAL_POW_ONE: "ALL n::nat. - real_pow (real_of_num (NUMERAL_BIT1 0)) n = real_of_num (NUMERAL_BIT1 0)" - by (import hollight REAL_POW_ONE) - -lemma REAL_LT_INV2: "ALL (x::hollight.real) y::hollight.real. - 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: "ALL (x::hollight.real) y::hollight.real. - 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_INV_1_LE: "ALL x::hollight.real. - real_lt (real_of_num 0) x & real_le x (real_of_num (NUMERAL_BIT1 0)) --> - real_le (real_of_num (NUMERAL_BIT1 0)) (real_inv x)" +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_SUB_INV: "ALL (x::hollight.real) xa::hollight.real. - 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)" +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: "ALL d::hollight.real. - real_lt (real_of_num 0) d --> - (EX x::hollight.real. real_lt (real_of_num 0) x & real_lt x d)" +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: "ALL (d1::hollight.real) d2::hollight.real. - real_lt (real_of_num 0) d1 & real_lt (real_of_num 0) d2 --> - (EX e::hollight.real. - real_lt (real_of_num 0) e & real_lt e d1 & real_lt e d2)" +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: "ALL (n::nat) (x::hollight.real) y::hollight.real. - real_le (real_of_num 0) x & real_le x y --> - real_le (real_pow x n) (real_pow y n)" +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_MONO: "ALL (m::nat) (n::nat) x::hollight.real. - real_le (real_of_num (NUMERAL_BIT1 0)) x & <= m n --> - real_le (real_pow x m) (real_pow x n)" +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: "ALL (n::nat) (x::hollight.real) y::hollight.real. - n ~= 0 & real_le (real_of_num 0) x & real_lt x y --> - real_lt (real_pow x n) (real_pow y n)" +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_MONO_LT: "ALL (m::nat) (n::nat) x::hollight.real. - real_lt (real_of_num (NUMERAL_BIT1 0)) x & < m n --> - real_lt (real_pow x m) (real_pow x n)" +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: "ALL (x::hollight.real) (m::nat) n::nat. - real_pow (real_pow x m) n = real_pow x (m * n)" +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: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. - z ~= real_of_num 0 & real_mul x z = real_mul y z --> x = y" +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: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real. - xb ~= real_of_num 0 & real_mul xb x = real_mul xb xa --> x = xa" +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: "ALL (x::hollight.real) xa::hollight.real. - real_lt (real_of_num 0) x & real_lt (real_of_num 0) xa --> - real_lt (real_of_num 0) (real_div x xa)" +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: "ALL (x::hollight.real) xa::hollight.real. - real_le (real_of_num 0) x & real_le (real_of_num 0) xa --> - real_le (real_of_num 0) (real_div x xa)" +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: "ALL (x::hollight.real) (m::nat) n::nat. - x ~= real_of_num 0 --> - real_div (real_pow x m) (real_pow x n) = - COND (<= n m) (real_pow x (m - n)) (real_inv (real_pow x (n - m)))" +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: "ALL (x::hollight.real) (m::nat) n::nat. - x ~= real_of_num 0 --> - real_div (real_pow x m) (real_pow x n) = - COND (< n m) (real_pow x (m - n)) (real_inv (real_pow x (n - m)))" +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: "ALL x::nat. - real_lt (real_of_num 0) - (real_pow (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 0))) x)" +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: "ALL n::nat. - real_le (real_of_num (NUMERAL_BIT1 0)) - (real_pow (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 0))) n)" +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: "ALL x::hollight.real. - real_pow (real_abs x) (NUMERAL_BIT0 (NUMERAL_BIT1 0)) = - real_pow x (NUMERAL_BIT0 (NUMERAL_BIT1 0))" +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: "ALL (x::hollight.real) y::hollight.real. - real_le (real_abs x) (real_abs y) = - real_le (real_pow x (NUMERAL_BIT0 (NUMERAL_BIT1 0))) - (real_pow y (NUMERAL_BIT0 (NUMERAL_BIT1 0)))" +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_SOS_EQ_0: "ALL (x::hollight.real) y::hollight.real. - (real_add (real_pow x (NUMERAL_BIT0 (NUMERAL_BIT1 0))) - (real_pow y (NUMERAL_BIT0 (NUMERAL_BIT1 0))) = - real_of_num 0) = - (x = real_of_num 0 & y = real_of_num 0)" +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_WLOG_LE: "(ALL (x::hollight.real) y::hollight.real. - (P::hollight.real => hollight.real => bool) x y = P y x) & -(ALL (x::hollight.real) y::hollight.real. real_le x y --> P x y) --> -(ALL x::hollight.real. All (P x))" +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::hollight.real. (P::hollight.real => hollight.real => bool) x x) & -(ALL (x::hollight.real) y::hollight.real. P x y = P y x) & -(ALL (x::hollight.real) y::hollight.real. real_lt x y --> P x y) --> -(ALL x::hollight.real. All (P x))" +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 mod_real :: "hollight.real => hollight.real => hollight.real => bool" where - "mod_real == -%(u::hollight.real) (ua::hollight.real) ub::hollight.real. - EX q::hollight.real. real_sub ua ub = real_mul q u" - -lemma DEF_mod_real: "mod_real = -(%(u::hollight.real) (ua::hollight.real) ub::hollight.real. - EX q::hollight.real. real_sub ua ub = real_mul q u)" - by (import hollight DEF_mod_real) - -definition DECIMAL :: "nat => nat => hollight.real" where - "DECIMAL == %(u::nat) ua::nat. real_div (real_of_num u) (real_of_num ua)" - -lemma DEF_DECIMAL: "DECIMAL = (%(u::nat) ua::nat. real_div (real_of_num u) (real_of_num ua))" +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::hollight.real) ~= real_of_num 0 & -(y2::hollight.real) ~= real_of_num 0 --> -real_add (real_div (x1::hollight.real) y1) - (real_div (x2::hollight.real) y2) = -real_mul (real_add (real_mul x1 y2) (real_mul x2 y1)) - (real_mul (real_inv y1) (real_inv y2))" +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::hollight.real) & -real_lt (real_of_num 0) (y2::hollight.real) --> -real_add (real_div (x1::hollight.real) y1) - (real_div (x2::hollight.real) y2) = -real_mul (real_add (real_mul x1 y2) (real_mul x2 y1)) - (real_mul (real_inv y1) (real_inv y2))" +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::hollight.real) & -real_lt (real_of_num 0) (y2::hollight.real) --> -real_sub (real_div (x1::hollight.real) y1) - (real_div (x2::hollight.real) y2) = -real_mul (real_sub (real_mul x1 y2) (real_mul x2 y1)) - (real_mul (real_inv y1) (real_inv y2))" +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::hollight.real) & -real_lt (real_of_num 0) (y2::hollight.real) --> -real_le (real_div (x1::hollight.real) y1) - (real_div (x2::hollight.real) y2) = -real_le (real_mul x1 y2) (real_mul x2 y1)" +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::hollight.real) & -real_lt (real_of_num 0) (y2::hollight.real) --> -(real_div (x1::hollight.real) y1 = real_div (x2::hollight.real) y2) = -(real_mul x1 y2 = real_mul x2 y1)" +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) -definition is_int :: "hollight.real => bool" where - "is_int == -%u::hollight.real. - EX n::nat. u = real_of_num n | u = real_neg (real_of_num n)" - -lemma DEF_is_int: "is_int = -(%u::hollight.real. - EX n::nat. u = real_of_num n | u = real_neg (real_of_num n))" - by (import hollight DEF_is_int) - -typedef (open) int = "Collect is_int" morphisms "dest_int" "mk_int" - apply (rule light_ex_imp_nonempty[where t="real_of_num (NUMERAL 0)"]) +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 - dest_int :: _ + real_of_int :: _ syntax - mk_int :: _ + 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: "ALL x::hollight.int. - EX n::nat. - dest_int x = real_of_num n | dest_int x = real_neg (real_of_num n)" +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::hollight.int) ua::hollight.int. real_le (dest_int u) (dest_int ua)" - -lemma DEF_int_le: "int_le = -(%(u::hollight.int) ua::hollight.int. real_le (dest_int u) (dest_int ua))" +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::hollight.int) ua::hollight.int. real_lt (dest_int u) (dest_int ua)" - -lemma DEF_int_lt: "int_lt = -(%(u::hollight.int) ua::hollight.int. real_lt (dest_int u) (dest_int ua))" +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::hollight.int) ua::hollight.int. - hollight.real_ge (dest_int u) (dest_int ua)" - -lemma DEF_int_ge: "int_ge = -(%(u::hollight.int) ua::hollight.int. - hollight.real_ge (dest_int u) (dest_int ua))" +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::hollight.int) ua::hollight.int. - hollight.real_gt (dest_int u) (dest_int ua)" - -lemma DEF_int_gt: "int_gt = -(%(u::hollight.int) ua::hollight.int. - hollight.real_gt (dest_int u) (dest_int ua))" +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::nat. mk_int (real_of_num u)" - -lemma DEF_int_of_num: "int_of_num = (%u::nat. mk_int (real_of_num u))" +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: "ALL x::nat. dest_int (int_of_num x) = real_of_num x" +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::hollight.int. mk_int (real_neg (dest_int u))" - -lemma DEF_int_neg: "int_neg = (%u::hollight.int. mk_int (real_neg (dest_int u)))" +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: "ALL x::hollight.int. dest_int (int_neg x) = real_neg (dest_int x)" +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 +definition + int_add :: "hollight.int => hollight.int => hollight.int" where "int_add == -%(u::hollight.int) ua::hollight.int. - mk_int (real_add (dest_int u) (dest_int ua))" +%u ua. + int_of_real (real_add (hollight.real_of_int u) (hollight.real_of_int ua))" lemma DEF_int_add: "int_add = -(%(u::hollight.int) ua::hollight.int. - mk_int (real_add (dest_int u) (dest_int ua)))" +(%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: "ALL (x::hollight.int) xa::hollight.int. - dest_int (int_add x xa) = real_add (dest_int x) (dest_int xa)" +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 +definition + int_sub :: "hollight.int => hollight.int => hollight.int" where "int_sub == -%(u::hollight.int) ua::hollight.int. - mk_int (real_sub (dest_int u) (dest_int ua))" +%u ua. + int_of_real (real_sub (hollight.real_of_int u) (hollight.real_of_int ua))" lemma DEF_int_sub: "int_sub = -(%(u::hollight.int) ua::hollight.int. - mk_int (real_sub (dest_int u) (dest_int ua)))" +(%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: "ALL (x::hollight.int) xa::hollight.int. - dest_int (int_sub x xa) = real_sub (dest_int x) (dest_int xa)" +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 +definition + int_mul :: "hollight.int => hollight.int => hollight.int" where "int_mul == -%(u::hollight.int) ua::hollight.int. - mk_int (real_mul (dest_int u) (dest_int ua))" +%u ua. + int_of_real (real_mul (hollight.real_of_int u) (hollight.real_of_int ua))" lemma DEF_int_mul: "int_mul = -(%(u::hollight.int) ua::hollight.int. - mk_int (real_mul (dest_int u) (dest_int ua)))" +(%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: "ALL (x::hollight.int) y::hollight.int. - dest_int (int_mul x y) = real_mul (dest_int x) (dest_int y)" +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::hollight.int. mk_int (real_abs (dest_int u))" - -lemma DEF_int_abs: "int_abs = (%u::hollight.int. mk_int (real_abs (dest_int u)))" +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: "ALL x::hollight.int. dest_int (int_abs x) = real_abs (dest_int x)" +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_max :: "hollight.int => hollight.int => hollight.int" where +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::hollight.int) ua::hollight.int. - mk_int (real_max (dest_int u) (dest_int ua))" +%u ua. + int_of_real (real_max (hollight.real_of_int u) (hollight.real_of_int ua))" lemma DEF_int_max: "int_max = -(%(u::hollight.int) ua::hollight.int. - mk_int (real_max (dest_int u) (dest_int ua)))" +(%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: "ALL (x::hollight.int) y::hollight.int. - dest_int (int_max x y) = real_max (dest_int x) (dest_int y)" +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 +definition + int_min :: "hollight.int => hollight.int => hollight.int" where "int_min == -%(u::hollight.int) ua::hollight.int. - mk_int (real_min (dest_int u) (dest_int ua))" +%u ua. + int_of_real (real_min (hollight.real_of_int u) (hollight.real_of_int ua))" lemma DEF_int_min: "int_min = -(%(u::hollight.int) ua::hollight.int. - mk_int (real_min (dest_int u) (dest_int ua)))" +(%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: "ALL (x::hollight.int) y::hollight.int. - dest_int (int_min x y) = real_min (dest_int x) (dest_int y)" +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::hollight.int) ua::nat. mk_int (real_pow (dest_int u) ua)" - -lemma DEF_int_pow: "int_pow = (%(u::hollight.int) ua::nat. mk_int (real_pow (dest_int u) ua))" +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: "ALL (x::hollight.int) xa::nat. - dest_int (int_pow x xa) = real_pow (dest_int x) xa" +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: "ALL x::hollight.int. - (EX n::nat. x = int_of_num n) | (EX n::nat. x = int_neg (int_of_num n))" +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: "ALL (x::hollight.int) y::hollight.int. - int_lt x y = int_le (int_add x (int_of_num (NUMERAL_BIT1 0))) y" +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: "ALL (x::hollight.int) xa::hollight.int. - int_gt x xa = int_ge x (int_add xa (int_of_num (NUMERAL_BIT1 0)))" +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::nat. (P::hollight.int => bool) (int_of_num n)) = -(ALL i::hollight.int. int_le (int_of_num 0) i --> P i)" +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_ABS_MUL_1: "ALL (x::hollight.int) y::hollight.int. - (int_abs (int_mul x y) = int_of_num (NUMERAL_BIT1 0)) = - (int_abs x = int_of_num (NUMERAL_BIT1 0) & - int_abs y = int_of_num (NUMERAL_BIT1 0))" +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_POW: "int_pow (x::hollight.int) 0 = int_of_num (NUMERAL_BIT1 0) & -(ALL xa::nat. int_pow x (Suc xa) = int_mul x (int_pow x xa))" +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: "ALL x::hollight.int. - int_abs x = COND (int_le (int_of_num 0) x) x (int_neg x)" +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: "ALL (x::hollight.int) xa::hollight.int. int_ge x xa = int_le xa x" +lemma INT_GE: "int_ge x xa = int_le xa x" by (import hollight INT_GE) -lemma INT_GT: "ALL (x::hollight.int) xa::hollight.int. int_gt x xa = int_lt xa x" +lemma INT_GT: "int_gt x xa = int_lt xa x" by (import hollight INT_GT) -lemma INT_LT: "ALL (x::hollight.int) xa::hollight.int. int_lt x xa = (~ int_le xa x)" +lemma INT_LT: "int_lt x xa = (~ int_le xa x)" by (import hollight INT_LT) -lemma INT_ARCH: "ALL (x::hollight.int) d::hollight.int. - d ~= int_of_num 0 --> (EX c::hollight.int. int_lt x (int_mul c d))" +lemma INT_ARCH: "d ~= int_of_num 0 ==> EX c. int_lt x (int_mul c d)" by (import hollight INT_ARCH) -definition mod_int :: "hollight.int => hollight.int => hollight.int => bool" where - "mod_int == -%(u::hollight.int) (ua::hollight.int) ub::hollight.int. - EX q::hollight.int. int_sub ua ub = int_mul q u" - -lemma DEF_mod_int: "mod_int = -(%(u::hollight.int) (ua::hollight.int) ub::hollight.int. - EX q::hollight.int. int_sub ua ub = int_mul q u)" - by (import hollight DEF_mod_int) - -definition IN :: "'A => ('A => bool) => bool" where - "IN == %(u::'A::type) ua::'A::type => bool. ua u" - -lemma DEF_IN: "IN = (%(u::'A::type) ua::'A::type => bool. ua u)" - by (import hollight DEF_IN) - -lemma EXTENSION: "ALL (x::'A::type => bool) xa::'A::type => bool. - (x = xa) = (ALL xb::'A::type. IN xb x = IN xb xa)" - by (import hollight EXTENSION) - -definition GSPEC :: "('A => bool) => 'A => bool" where - "GSPEC == %u::'A::type => bool. u" - -lemma DEF_GSPEC: "GSPEC = (%u::'A::type => bool. u)" - by (import hollight DEF_GSPEC) - -definition SETSPEC :: "'q_37056 => bool => 'q_37056 => bool" where - "SETSPEC == %(u::'q_37056::type) (ua::bool) ub::'q_37056::type. ua & u = ub" - -lemma DEF_SETSPEC: "SETSPEC = (%(u::'q_37056::type) (ua::bool) ub::'q_37056::type. ua & u = ub)" - by (import hollight DEF_SETSPEC) - -lemma IN_ELIM_THM: "(ALL (P::(bool => 'q_37089::type => bool) => bool) x::'q_37089::type. - IN x (GSPEC (%v::'q_37089::type. P (SETSPEC v))) = - P (%(p::bool) t::'q_37089::type. p & x = t)) & -(ALL (p::'q_37120::type => bool) x::'q_37120::type. - IN x - (GSPEC (%v::'q_37120::type. EX y::'q_37120::type. SETSPEC v (p y) y)) = - p x) & -(ALL (P::(bool => 'q_37148::type => bool) => bool) x::'q_37148::type. - GSPEC (%v::'q_37148::type. P (SETSPEC v)) x = - P (%(p::bool) t::'q_37148::type. p & x = t)) & -(ALL (p::'q_37177::type => bool) x::'q_37177::type. - GSPEC (%v::'q_37177::type. EX y::'q_37177::type. SETSPEC v (p y) y) x = - p x) & -(ALL (p::'q_37194::type => bool) x::'q_37194::type. IN x p = p x)" - by (import hollight IN_ELIM_THM) - -definition EMPTY :: "'A => bool" where - "EMPTY == %x::'A::type. False" - -lemma DEF_EMPTY: "EMPTY = (%x::'A::type. False)" - by (import hollight DEF_EMPTY) - -definition INSERT :: "'A => ('A => bool) => 'A => bool" where - "INSERT == %(u::'A::type) (ua::'A::type => bool) y::'A::type. IN y ua | y = u" - -lemma DEF_INSERT: "INSERT = -(%(u::'A::type) (ua::'A::type => bool) y::'A::type. IN y ua | y = u)" - by (import hollight DEF_INSERT) +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 - UNIV :: "'A => bool" + div :: "hollight.int => hollight.int => hollight.int" ("div") defs - UNIV_def: "hollight.UNIV == %x::'A::type. True" - -lemma DEF_UNIV: "hollight.UNIV = (%x::'A::type. True)" - by (import hollight DEF_UNIV) + 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 - UNION :: "('A => bool) => ('A => bool) => 'A => bool" + 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 - UNION_def: "hollight.UNION == -%(u::'A::type => bool) ua::'A::type => bool. - GSPEC (%ub::'A::type. EX x::'A::type. SETSPEC ub (IN x u | IN x ua) x)" - -lemma DEF_UNION: "hollight.UNION = -(%(u::'A::type => bool) ua::'A::type => bool. - GSPEC (%ub::'A::type. EX x::'A::type. SETSPEC ub (IN x u | IN x ua) x))" - by (import hollight DEF_UNION) - -definition UNIONS :: "(('A => bool) => bool) => 'A => bool" where - "UNIONS == -%u::('A::type => bool) => bool. - GSPEC - (%ua::'A::type. - EX x::'A::type. - SETSPEC ua (EX ua::'A::type => bool. IN ua u & IN x ua) x)" - -lemma DEF_UNIONS: "UNIONS = -(%u::('A::type => bool) => bool. - GSPEC - (%ua::'A::type. - EX x::'A::type. - SETSPEC ua (EX ua::'A::type => bool. IN ua u & IN x ua) x))" - by (import hollight DEF_UNIONS) + 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 - INTER :: "('A => bool) => ('A => bool) => 'A => bool" + 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 - INTER_def: "hollight.INTER == -%(u::'A::type => bool) ua::'A::type => bool. - GSPEC (%ub::'A::type. EX x::'A::type. SETSPEC ub (IN x u & IN x ua) x)" - -lemma DEF_INTER: "hollight.INTER = -(%(u::'A::type => bool) ua::'A::type => bool. - GSPEC (%ub::'A::type. EX x::'A::type. SETSPEC ub (IN x u & IN x ua) x))" - by (import hollight DEF_INTER) - -definition INTERS :: "(('A => bool) => bool) => 'A => bool" where - "INTERS == -%u::('A::type => bool) => bool. - GSPEC - (%ua::'A::type. - EX x::'A::type. - SETSPEC ua (ALL ua::'A::type => bool. IN ua u --> IN x ua) x)" - -lemma DEF_INTERS: "INTERS = -(%u::('A::type => bool) => bool. - GSPEC - (%ua::'A::type. - EX x::'A::type. - SETSPEC ua (ALL ua::'A::type => bool. IN ua u --> IN x ua) x))" - by (import hollight DEF_INTERS) - -definition DIFF :: "('A => bool) => ('A => bool) => 'A => bool" where - "DIFF == -%(u::'A::type => bool) ua::'A::type => bool. - GSPEC (%ub::'A::type. EX x::'A::type. SETSPEC ub (IN x u & ~ IN x ua) x)" - -lemma DEF_DIFF: "DIFF = -(%(u::'A::type => bool) ua::'A::type => bool. - GSPEC - (%ub::'A::type. EX x::'A::type. SETSPEC ub (IN x u & ~ IN x ua) x))" - by (import hollight DEF_DIFF) - -lemma INSERT: "INSERT (x::'A::type) (s::'A::type => bool) = -GSPEC (%u::'A::type. EX y::'A::type. SETSPEC u (IN y s | y = x) y)" + 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. + {v::'q_43354. P (SETSPEC v)} x = + P (%(p::bool) t::'q_43354. p & x = t)) & +(ALL (p::'q_43383 => bool) x::'q_43383. + {v::'q_43383. EX y::'q_43383. p y & v = y} x = p x) & +(ALL (p::'q_43400 => bool) x::'q_43400. (x : p) = p x)" + by (import hollight IN_ELIM_THM) + +lemma INSERT: "insert (x::'A) (s::'A => bool) = {u::'A. EX y::'A. (y : s | y = x) & u = y}" by (import hollight INSERT) -definition DELETE :: "('A => bool) => 'A => 'A => bool" where - "DELETE == -%(u::'A::type => bool) ua::'A::type. - GSPEC (%ub::'A::type. EX y::'A::type. SETSPEC ub (IN y u & y ~= ua) y)" - -lemma DEF_DELETE: "DELETE = -(%(u::'A::type => bool) ua::'A::type. - GSPEC (%ub::'A::type. EX y::'A::type. SETSPEC ub (IN y u & y ~= ua) y))" - by (import hollight DEF_DELETE) - -definition SUBSET :: "('A => bool) => ('A => bool) => bool" where - "SUBSET == -%(u::'A::type => bool) ua::'A::type => bool. - ALL x::'A::type. IN x u --> IN x ua" - -lemma DEF_SUBSET: "SUBSET = -(%(u::'A::type => bool) ua::'A::type => bool. - ALL x::'A::type. IN x u --> IN x ua)" - by (import hollight DEF_SUBSET) - -definition PSUBSET :: "('A => bool) => ('A => bool) => bool" where - "PSUBSET == -%(u::'A::type => bool) ua::'A::type => bool. SUBSET u ua & u ~= ua" - -lemma DEF_PSUBSET: "PSUBSET = -(%(u::'A::type => bool) ua::'A::type => bool. SUBSET u ua & u ~= ua)" - by (import hollight DEF_PSUBSET) - -definition DISJOINT :: "('A => bool) => ('A => bool) => bool" where - "DISJOINT == -%(u::'A::type => bool) ua::'A::type => bool. hollight.INTER u ua = EMPTY" - -lemma DEF_DISJOINT: "DISJOINT = -(%(u::'A::type => bool) ua::'A::type => bool. hollight.INTER u ua = EMPTY)" - by (import hollight DEF_DISJOINT) - -definition SING :: "('A => bool) => bool" where - "SING == %u::'A::type => bool. EX x::'A::type. u = INSERT x EMPTY" - -lemma DEF_SING: "SING = (%u::'A::type => bool. EX x::'A::type. u = INSERT x EMPTY)" +definition + SING :: "('A => bool) => bool" where + "SING == %u::'A => bool. EX x::'A. u = {x}" + +lemma DEF_SING: "SING = (%u::'A => bool. EX x::'A. u = {x})" by (import hollight DEF_SING) -definition FINITE :: "('A => bool) => bool" where - "FINITE == -%a::'A::type => bool. - ALL FINITE'::('A::type => bool) => bool. - (ALL a::'A::type => bool. - a = EMPTY | - (EX (x::'A::type) s::'A::type => bool. - a = INSERT x s & FINITE' s) --> - FINITE' a) --> - FINITE' a" - -lemma DEF_FINITE: "FINITE = -(%a::'A::type => bool. - ALL FINITE'::('A::type => bool) => bool. - (ALL a::'A::type => bool. - a = EMPTY | - (EX (x::'A::type) s::'A::type => bool. - a = INSERT x s & FINITE' s) --> - FINITE' a) --> - FINITE' a)" - by (import hollight DEF_FINITE) - -definition INFINITE :: "('A => bool) => bool" where - "INFINITE == %u::'A::type => bool. ~ FINITE u" - -lemma DEF_INFINITE: "INFINITE = (%u::'A::type => bool. ~ FINITE u)" - by (import hollight DEF_INFINITE) - -definition IMAGE :: "('A => 'B) => ('A => bool) => 'B => bool" where - "IMAGE == -%(u::'A::type => 'B::type) ua::'A::type => bool. - GSPEC - (%ub::'B::type. - EX y::'B::type. SETSPEC ub (EX x::'A::type. IN x ua & y = u x) y)" - -lemma DEF_IMAGE: "IMAGE = -(%(u::'A::type => 'B::type) ua::'A::type => bool. - GSPEC - (%ub::'B::type. - EX y::'B::type. SETSPEC ub (EX x::'A::type. IN x ua & y = u x) y))" - by (import hollight DEF_IMAGE) - -definition INJ :: "('A => 'B) => ('A => bool) => ('B => bool) => bool" where +definition + INJ :: "('A => 'B) => ('A => bool) => ('B => bool) => bool" where "INJ == -%(u::'A::type => 'B::type) (ua::'A::type => bool) ub::'B::type => bool. - (ALL x::'A::type. IN x ua --> IN (u x) ub) & - (ALL (x::'A::type) y::'A::type. IN x ua & IN y ua & u x = u y --> x = y)" +%(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::type => 'B::type) (ua::'A::type => bool) ub::'B::type => bool. - (ALL x::'A::type. IN x ua --> IN (u x) ub) & - (ALL (x::'A::type) y::'A::type. - IN x ua & IN y ua & u x = u y --> x = y))" +(%(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 +definition + SURJ :: "('A => 'B) => ('A => bool) => ('B => bool) => bool" where "SURJ == -%(u::'A::type => 'B::type) (ua::'A::type => bool) ub::'B::type => bool. - (ALL x::'A::type. IN x ua --> IN (u x) ub) & - (ALL x::'B::type. IN x ub --> (EX y::'A::type. IN y ua & u y = x))" +%(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::type => 'B::type) (ua::'A::type => bool) ub::'B::type => bool. - (ALL x::'A::type. IN x ua --> IN (u x) ub) & - (ALL x::'B::type. IN x ub --> (EX y::'A::type. IN y ua & u y = x)))" +(%(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 +definition + BIJ :: "('A => 'B) => ('A => bool) => ('B => bool) => bool" where "BIJ == -%(u::'A::type => 'B::type) (ua::'A::type => bool) ub::'B::type => bool. - INJ u ua ub & SURJ u ua ub" +%(u::'A => 'B) (ua::'A => bool) ub::'B => bool. INJ u ua ub & SURJ u ua ub" lemma DEF_BIJ: "BIJ = -(%(u::'A::type => 'B::type) (ua::'A::type => bool) ub::'B::type => bool. - INJ u ua ub & SURJ u ua ub)" +(%(u::'A => 'B) (ua::'A => bool) ub::'B => bool. INJ u ua ub & SURJ u ua ub)" by (import hollight DEF_BIJ) -definition CHOICE :: "('A => bool) => 'A" where - "CHOICE == %u::'A::type => bool. SOME x::'A::type. IN x u" - -lemma DEF_CHOICE: "CHOICE = (%u::'A::type => bool. SOME x::'A::type. IN x u)" - by (import hollight DEF_CHOICE) - -definition REST :: "('A => bool) => 'A => bool" where - "REST == %u::'A::type => bool. DELETE u (CHOICE u)" - -lemma DEF_REST: "REST = (%u::'A::type => bool. DELETE u (CHOICE u))" +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) -definition CARD_GE :: "('q_37693 => bool) => ('q_37690 => bool) => bool" where - "CARD_GE == -%(u::'q_37693::type => bool) ua::'q_37690::type => bool. - EX f::'q_37693::type => 'q_37690::type. - ALL y::'q_37690::type. - IN y ua --> (EX x::'q_37693::type. IN x u & y = f x)" - -lemma DEF_CARD_GE: "CARD_GE = -(%(u::'q_37693::type => bool) ua::'q_37690::type => bool. - EX f::'q_37693::type => 'q_37690::type. - ALL y::'q_37690::type. - IN y ua --> (EX x::'q_37693::type. IN x u & y = f x))" - by (import hollight DEF_CARD_GE) - -definition CARD_LE :: "('q_37702 => bool) => ('q_37701 => bool) => bool" where - "CARD_LE == -%(u::'q_37702::type => bool) ua::'q_37701::type => bool. CARD_GE ua u" - -lemma DEF_CARD_LE: "CARD_LE = -(%(u::'q_37702::type => bool) ua::'q_37701::type => bool. CARD_GE ua u)" - by (import hollight DEF_CARD_LE) - -definition CARD_EQ :: "('q_37712 => bool) => ('q_37713 => bool) => bool" where - "CARD_EQ == -%(u::'q_37712::type => bool) ua::'q_37713::type => bool. - CARD_LE u ua & CARD_LE ua u" - -lemma DEF_CARD_EQ: "CARD_EQ = -(%(u::'q_37712::type => bool) ua::'q_37713::type => bool. - CARD_LE u ua & CARD_LE ua u)" - by (import hollight DEF_CARD_EQ) - -definition CARD_GT :: "('q_37727 => bool) => ('q_37728 => bool) => bool" where - "CARD_GT == -%(u::'q_37727::type => bool) ua::'q_37728::type => bool. - CARD_GE u ua & ~ CARD_GE ua u" - -lemma DEF_CARD_GT: "CARD_GT = -(%(u::'q_37727::type => bool) ua::'q_37728::type => bool. - CARD_GE u ua & ~ CARD_GE ua u)" - by (import hollight DEF_CARD_GT) - -definition CARD_LT :: "('q_37743 => bool) => ('q_37744 => bool) => bool" where - "CARD_LT == -%(u::'q_37743::type => bool) ua::'q_37744::type => bool. - CARD_LE u ua & ~ CARD_LE ua u" - -lemma DEF_CARD_LT: "CARD_LT = -(%(u::'q_37743::type => bool) ua::'q_37744::type => bool. - CARD_LE u ua & ~ CARD_LE ua u)" - by (import hollight DEF_CARD_LT) - -definition COUNTABLE :: "('q_37757 => bool) => bool" where - "(op ==::(('q_37757::type => bool) => bool) - => (('q_37757::type => bool) => bool) => prop) - (COUNTABLE::('q_37757::type => bool) => bool) - ((CARD_GE::(nat => bool) => ('q_37757::type => bool) => bool) - (hollight.UNIV::nat => bool))" - -lemma DEF_COUNTABLE: "(op =::(('q_37757::type => bool) => bool) - => (('q_37757::type => bool) => bool) => bool) - (COUNTABLE::('q_37757::type => bool) => bool) - ((CARD_GE::(nat => bool) => ('q_37757::type => bool) => bool) - (hollight.UNIV::nat => bool))" - by (import hollight DEF_COUNTABLE) - -lemma NOT_IN_EMPTY: "ALL x::'A::type. ~ IN x EMPTY" +lemma NOT_IN_EMPTY: "(x::'A) ~: {}" by (import hollight NOT_IN_EMPTY) -lemma IN_UNIV: "ALL x::'A::type. IN x hollight.UNIV" - by (import hollight IN_UNIV) - -lemma IN_UNION: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type. - IN xb (hollight.UNION x xa) = (IN xb x | IN xb xa)" - by (import hollight IN_UNION) - -lemma IN_UNIONS: "ALL (x::('A::type => bool) => bool) xa::'A::type. - IN xa (UNIONS x) = (EX t::'A::type => bool. IN t x & IN xa t)" +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_INTER: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type. - IN xb (hollight.INTER x xa) = (IN xb x & IN xb xa)" - by (import hollight IN_INTER) - -lemma IN_INTERS: "ALL (x::('A::type => bool) => bool) xa::'A::type. - IN xa (INTERS x) = (ALL t::'A::type => bool. IN t x --> IN xa t)" +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_DIFF: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type. - IN xb (DIFF x xa) = (IN xb x & ~ IN xb xa)" - by (import hollight IN_DIFF) - -lemma IN_INSERT: "ALL (x::'A::type) (xa::'A::type) xb::'A::type => bool. - IN x (INSERT xa xb) = (x = xa | IN x xb)" - by (import hollight IN_INSERT) - -lemma IN_DELETE: "ALL (x::'A::type => bool) (xa::'A::type) xb::'A::type. - IN xa (DELETE x xb) = (IN xa x & xa ~= xb)" +lemma IN_DELETE: "((xa::'A) : (x::'A => bool) - {xb::'A}) = (xa : x & xa ~= xb)" by (import hollight IN_DELETE) -lemma IN_SING: "ALL (x::'A::type) xa::'A::type. IN x (INSERT xa EMPTY) = (x = xa)" - by (import hollight IN_SING) - -lemma IN_IMAGE: "ALL (x::'B::type) (xa::'A::type => bool) xb::'A::type => 'B::type. - IN x (IMAGE xb xa) = (EX xc::'A::type. x = xb xc & IN xc xa)" +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: "ALL (x::'A::type) xa::'A::type => bool. - IN x (REST xa) = (IN x xa & x ~= CHOICE xa)" +lemma IN_REST: "((x::'A) : REST (xa::'A => bool)) = (x : xa & x ~= Eps xa)" by (import hollight IN_REST) -lemma CHOICE_DEF: "ALL x::'A::type => bool. x ~= EMPTY --> IN (CHOICE x) x" +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: "ALL (x::'A::type => bool) xa::'A::type => bool. - (x ~= xa) = (EX xb::'A::type. IN xb xa = (~ IN xb x))" +lemma NOT_EQUAL_SETS: "((x::'A => bool) ~= (xa::'A => bool)) = (EX xb::'A. (xb : xa) = (xb ~: x))" by (import hollight NOT_EQUAL_SETS) -lemma MEMBER_NOT_EMPTY: "ALL x::'A::type => bool. (EX xa::'A::type. IN xa x) = (x ~= EMPTY)" - by (import hollight MEMBER_NOT_EMPTY) - -lemma UNIV_NOT_EMPTY: "(Not::bool => bool) - ((op =::('A::type => bool) => ('A::type => bool) => bool) - (hollight.UNIV::'A::type => bool) (EMPTY::'A::type => bool))" - by (import hollight UNIV_NOT_EMPTY) - -lemma EMPTY_NOT_UNIV: "(Not::bool => bool) - ((op =::('A::type => bool) => ('A::type => bool) => bool) - (EMPTY::'A::type => bool) (hollight.UNIV::'A::type => bool))" +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::type. IN x (s::'A::type => bool)) = (s = hollight.UNIV)" +lemma EQ_UNIV: "(ALL x::'A. x : (s::'A => bool)) = (s = UNIV)" by (import hollight EQ_UNIV) -lemma SUBSET_TRANS: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type => bool. - SUBSET x xa & SUBSET xa xb --> SUBSET x xb" - by (import hollight SUBSET_TRANS) - -lemma SUBSET_REFL: "ALL x::'A::type => bool. SUBSET x x" - by (import hollight SUBSET_REFL) - -lemma SUBSET_ANTISYM: "ALL (x::'A::type => bool) xa::'A::type => bool. - SUBSET x xa & SUBSET xa x --> x = xa" - by (import hollight SUBSET_ANTISYM) - -lemma EMPTY_SUBSET: "(All::(('A::type => bool) => bool) => bool) - ((SUBSET::('A::type => bool) => ('A::type => bool) => bool) - (EMPTY::'A::type => bool))" - by (import hollight EMPTY_SUBSET) - -lemma SUBSET_EMPTY: "ALL x::'A::type => bool. SUBSET x EMPTY = (x = EMPTY)" - by (import hollight SUBSET_EMPTY) - -lemma SUBSET_UNIV: "ALL x::'A::type => bool. SUBSET x hollight.UNIV" - by (import hollight SUBSET_UNIV) - -lemma UNIV_SUBSET: "ALL x::'A::type => bool. SUBSET hollight.UNIV x = (x = hollight.UNIV)" +lemma UNIV_SUBSET: "(UNIV <= (x::'A => bool)) = (x = UNIV)" by (import hollight UNIV_SUBSET) -lemma PSUBSET_TRANS: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type => bool. - PSUBSET x xa & PSUBSET xa xb --> PSUBSET x xb" - by (import hollight PSUBSET_TRANS) - -lemma PSUBSET_SUBSET_TRANS: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type => bool. - PSUBSET x xa & SUBSET xa xb --> PSUBSET x xb" - by (import hollight PSUBSET_SUBSET_TRANS) - -lemma SUBSET_PSUBSET_TRANS: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type => bool. - SUBSET x xa & PSUBSET xa xb --> PSUBSET x xb" - by (import hollight SUBSET_PSUBSET_TRANS) - -lemma PSUBSET_IRREFL: "ALL x::'A::type => bool. ~ PSUBSET x x" - by (import hollight PSUBSET_IRREFL) - -lemma NOT_PSUBSET_EMPTY: "ALL x::'A::type => bool. ~ PSUBSET x EMPTY" - by (import hollight NOT_PSUBSET_EMPTY) - -lemma NOT_UNIV_PSUBSET: "ALL x::'A::type => bool. ~ PSUBSET hollight.UNIV x" - by (import hollight NOT_UNIV_PSUBSET) - -lemma PSUBSET_UNIV: "ALL x::'A::type => bool. - PSUBSET x hollight.UNIV = (EX xa::'A::type. ~ IN xa x)" +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 UNION_ASSOC: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type => bool. - hollight.UNION (hollight.UNION x xa) xb = - hollight.UNION x (hollight.UNION xa xb)" - by (import hollight UNION_ASSOC) - -lemma UNION_IDEMPOT: "ALL x::'A::type => bool. hollight.UNION x x = x" - by (import hollight UNION_IDEMPOT) - -lemma UNION_COMM: "ALL (x::'A::type => bool) xa::'A::type => bool. - hollight.UNION x xa = hollight.UNION xa x" - by (import hollight UNION_COMM) - -lemma SUBSET_UNION: "(ALL (x::'A::type => bool) xa::'A::type => bool. - SUBSET x (hollight.UNION x xa)) & -(ALL (x::'A::type => bool) xa::'A::type => bool. - SUBSET x (hollight.UNION xa x))" +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 SUBSET_UNION_ABSORPTION: "ALL (x::'A::type => bool) xa::'A::type => bool. - SUBSET x xa = (hollight.UNION x xa = xa)" - by (import hollight SUBSET_UNION_ABSORPTION) - -lemma UNION_EMPTY: "(ALL x::'A::type => bool. hollight.UNION EMPTY x = x) & -(ALL x::'A::type => bool. hollight.UNION x EMPTY = x)" +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::type => bool. hollight.UNION hollight.UNIV x = hollight.UNIV) & -(ALL x::'A::type => bool. hollight.UNION x hollight.UNIV = hollight.UNIV)" +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 EMPTY_UNION: "ALL (x::'A::type => bool) xa::'A::type => bool. - (hollight.UNION x xa = EMPTY) = (x = EMPTY & xa = EMPTY)" - by (import hollight EMPTY_UNION) - -lemma UNION_SUBSET: "ALL (x::'q_38594::type => bool) (xa::'q_38594::type => bool) - xb::'q_38594::type => bool. - SUBSET (hollight.UNION x xa) xb = (SUBSET x xb & SUBSET xa xb)" - by (import hollight UNION_SUBSET) - -lemma INTER_ASSOC: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type => bool. - hollight.INTER (hollight.INTER x xa) xb = - hollight.INTER x (hollight.INTER xa xb)" - by (import hollight INTER_ASSOC) - -lemma INTER_IDEMPOT: "ALL x::'A::type => bool. hollight.INTER x x = x" - by (import hollight INTER_IDEMPOT) - -lemma INTER_COMM: "ALL (x::'A::type => bool) xa::'A::type => bool. - hollight.INTER x xa = hollight.INTER xa x" - by (import hollight INTER_COMM) - -lemma INTER_SUBSET: "(ALL (x::'A::type => bool) xa::'A::type => bool. - SUBSET (hollight.INTER x xa) x) & -(ALL (x::'A::type => bool) xa::'A::type => bool. - SUBSET (hollight.INTER xa x) x)" +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 SUBSET_INTER_ABSORPTION: "ALL (x::'A::type => bool) xa::'A::type => bool. - SUBSET x xa = (hollight.INTER x xa = x)" - by (import hollight SUBSET_INTER_ABSORPTION) - -lemma INTER_EMPTY: "(ALL x::'A::type => bool. hollight.INTER EMPTY x = EMPTY) & -(ALL x::'A::type => bool. hollight.INTER x EMPTY = EMPTY)" +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::type => bool. hollight.INTER hollight.UNIV x = x) & -(ALL x::'A::type => bool. hollight.INTER x hollight.UNIV = x)" +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 UNION_OVER_INTER: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type => bool. - hollight.INTER x (hollight.UNION xa xb) = - hollight.UNION (hollight.INTER x xa) (hollight.INTER x xb)" - by (import hollight UNION_OVER_INTER) - -lemma INTER_OVER_UNION: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type => bool. - hollight.UNION x (hollight.INTER xa xb) = - hollight.INTER (hollight.UNION x xa) (hollight.UNION x xb)" - by (import hollight INTER_OVER_UNION) - -lemma IN_DISJOINT: "ALL (x::'A::type => bool) xa::'A::type => bool. - DISJOINT x xa = (~ (EX xb::'A::type. IN xb x & IN xb xa))" +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: "ALL (x::'A::type => bool) xa::'A::type => bool. - DISJOINT x xa = DISJOINT xa x" +lemma DISJOINT_SYM: "((x::'A => bool) Int (xa::'A => bool) = {}) = (xa Int x = {})" by (import hollight DISJOINT_SYM) -lemma DISJOINT_EMPTY: "ALL x::'A::type => bool. DISJOINT EMPTY x & DISJOINT x EMPTY" +lemma DISJOINT_EMPTY: "{} Int (x::'A => bool) = {} & x Int {} = {}" by (import hollight DISJOINT_EMPTY) -lemma DISJOINT_EMPTY_REFL: "ALL x::'A::type => bool. (x = EMPTY) = DISJOINT x x" +lemma DISJOINT_EMPTY_REFL: "((x::'A => bool) = {}) = (x Int x = {})" by (import hollight DISJOINT_EMPTY_REFL) -lemma DISJOINT_UNION: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type => bool. - DISJOINT (hollight.UNION x xa) xb = (DISJOINT x xb & DISJOINT xa xb)" +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 DIFF_EMPTY: "ALL x::'A::type => bool. DIFF x EMPTY = x" - by (import hollight DIFF_EMPTY) - -lemma EMPTY_DIFF: "ALL x::'A::type => bool. DIFF EMPTY x = EMPTY" - by (import hollight EMPTY_DIFF) - -lemma DIFF_UNIV: "ALL x::'A::type => bool. DIFF x hollight.UNIV = EMPTY" - by (import hollight DIFF_UNIV) - -lemma DIFF_DIFF: "ALL (x::'A::type => bool) xa::'A::type => bool. - DIFF (DIFF x xa) xa = DIFF x xa" - by (import hollight DIFF_DIFF) - -lemma DIFF_EQ_EMPTY: "ALL x::'A::type => bool. DIFF x x = EMPTY" - by (import hollight DIFF_EQ_EMPTY) - -lemma SUBSET_DIFF: "ALL (x::'q_39012::type => bool) xa::'q_39012::type => bool. - SUBSET (DIFF x xa) x" - by (import hollight SUBSET_DIFF) - -lemma COMPONENT: "ALL (x::'A::type) s::'A::type => bool. IN x (INSERT x s)" - by (import hollight COMPONENT) - -lemma DECOMPOSITION: "ALL (s::'A::type => bool) x::'A::type. - IN x s = (EX t::'A::type => bool. s = INSERT x t & ~ IN x t)" +lemma DECOMPOSITION: "((x::'A) : (s::'A => bool)) = (EX t::'A => bool. s = insert x t & x ~: t)" by (import hollight DECOMPOSITION) -lemma SET_CASES: "ALL s::'A::type => bool. - s = EMPTY | - (EX (x::'A::type) t::'A::type => bool. s = INSERT x t & ~ IN x t)" +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: "ALL (x::'A::type) xa::'A::type => bool. IN x xa = (INSERT x xa = xa)" +lemma ABSORPTION: "((x::'A) : (xa::'A => bool)) = (insert x xa = xa)" by (import hollight ABSORPTION) -lemma INSERT_INSERT: "ALL (x::'A::type) xa::'A::type => bool. INSERT x (INSERT x xa) = INSERT x xa" - by (import hollight INSERT_INSERT) - -lemma INSERT_COMM: "ALL (x::'A::type) (xa::'A::type) xb::'A::type => bool. - INSERT x (INSERT xa xb) = INSERT xa (INSERT x xb)" - by (import hollight INSERT_COMM) - -lemma INSERT_UNIV: "ALL x::'A::type. INSERT x hollight.UNIV = hollight.UNIV" +lemma INSERT_UNIV: "insert (x::'A) UNIV = UNIV" by (import hollight INSERT_UNIV) -lemma NOT_INSERT_EMPTY: "ALL (x::'A::type) xa::'A::type => bool. INSERT x xa ~= EMPTY" - by (import hollight NOT_INSERT_EMPTY) - -lemma NOT_EMPTY_INSERT: "ALL (x::'A::type) xa::'A::type => bool. EMPTY ~= INSERT x xa" - by (import hollight NOT_EMPTY_INSERT) - -lemma INSERT_UNION: "ALL (x::'A::type) (s::'A::type => bool) t::'A::type => bool. - hollight.UNION (INSERT x s) t = - COND (IN x t) (hollight.UNION s t) (INSERT x (hollight.UNION s t))" +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 INSERT_UNION_EQ: "ALL (x::'A::type) (xa::'A::type => bool) xb::'A::type => bool. - hollight.UNION (INSERT x xa) xb = INSERT x (hollight.UNION xa xb)" - by (import hollight INSERT_UNION_EQ) - -lemma INSERT_INTER: "ALL (x::'A::type) (s::'A::type => bool) t::'A::type => bool. - hollight.INTER (INSERT x s) t = - COND (IN x t) (INSERT x (hollight.INTER s t)) (hollight.INTER s t)" - by (import hollight INSERT_INTER) - -lemma DISJOINT_INSERT: "ALL (x::'A::type) (xa::'A::type => bool) xb::'A::type => bool. - DISJOINT (INSERT x xa) xb = (DISJOINT xa xb & ~ IN x xb)" +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_SUBSET: "ALL (x::'A::type) (xa::'A::type => bool) xb::'A::type => bool. - SUBSET (INSERT x xa) xb = (IN x xb & SUBSET xa xb)" - by (import hollight INSERT_SUBSET) - -lemma SUBSET_INSERT: "ALL (x::'A::type) xa::'A::type => bool. - ~ IN x xa --> - (ALL xb::'A::type => bool. SUBSET xa (INSERT x xb) = SUBSET xa xb)" - by (import hollight SUBSET_INSERT) - -lemma INSERT_DIFF: "ALL (s::'A::type => bool) (t::'A::type => bool) x::'A::type. - DIFF (INSERT x s) t = COND (IN x t) (DIFF s t) (INSERT x (DIFF s t))" - by (import hollight INSERT_DIFF) - -lemma INSERT_AC: "INSERT (x::'q_39468::type) - (INSERT (y::'q_39468::type) (s::'q_39468::type => bool)) = -INSERT y (INSERT x s) & -INSERT x (INSERT x s) = INSERT x s" +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: "hollight.INTER (p::'q_39535::type => bool) (q::'q_39535::type => bool) = -hollight.INTER q p & -hollight.INTER (hollight.INTER p q) (r::'q_39535::type => bool) = -hollight.INTER p (hollight.INTER q r) & -hollight.INTER p (hollight.INTER q r) = -hollight.INTER q (hollight.INTER p r) & -hollight.INTER p p = p & -hollight.INTER p (hollight.INTER p q) = hollight.INTER p q" +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: "hollight.UNION (p::'q_39601::type => bool) (q::'q_39601::type => bool) = -hollight.UNION q p & -hollight.UNION (hollight.UNION p q) (r::'q_39601::type => bool) = -hollight.UNION p (hollight.UNION q r) & -hollight.UNION p (hollight.UNION q r) = -hollight.UNION q (hollight.UNION p r) & -hollight.UNION p p = p & -hollight.UNION p (hollight.UNION p q) = hollight.UNION p q" +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: "ALL (x::'A::type) xa::'A::type => bool. (~ IN x xa) = (DELETE xa x = xa)" +lemma DELETE_NON_ELEMENT: "((x::'A) ~: (xa::'A => bool)) = (xa - {x} = xa)" by (import hollight DELETE_NON_ELEMENT) -lemma IN_DELETE_EQ: "ALL (s::'A::type => bool) (x::'A::type) x'::'A::type. - (IN x s = IN x' s) = (IN x (DELETE s x') = IN x' (DELETE s x))" +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: "ALL x::'A::type. DELETE EMPTY x = EMPTY" +lemma EMPTY_DELETE: "{} - {x::'A} = {}" by (import hollight EMPTY_DELETE) -lemma DELETE_DELETE: "ALL (x::'A::type) xa::'A::type => bool. DELETE (DELETE xa x) x = DELETE xa x" +lemma DELETE_DELETE: "(xa::'A => bool) - {x::'A} - {x} = xa - {x}" by (import hollight DELETE_DELETE) -lemma DELETE_COMM: "ALL (x::'A::type) (xa::'A::type) xb::'A::type => bool. - DELETE (DELETE xb x) xa = DELETE (DELETE xb xa) x" +lemma DELETE_COMM: "(xb::'A => bool) - {x::'A} - {xa::'A} = xb - {xa} - {x}" by (import hollight DELETE_COMM) -lemma DELETE_SUBSET: "ALL (x::'A::type) xa::'A::type => bool. SUBSET (DELETE xa x) xa" +lemma DELETE_SUBSET: "(xa::'A => bool) - {x::'A} <= xa" by (import hollight DELETE_SUBSET) -lemma SUBSET_DELETE: "ALL (x::'A::type) (xa::'A::type => bool) xb::'A::type => bool. - SUBSET xa (DELETE xb x) = (~ IN x xa & SUBSET xa xb)" +lemma SUBSET_DELETE: "((xa::'A => bool) <= (xb::'A => bool) - {x::'A}) = (x ~: xa & xa <= xb)" by (import hollight SUBSET_DELETE) -lemma SUBSET_INSERT_DELETE: "ALL (x::'A::type) (xa::'A::type => bool) xb::'A::type => bool. - SUBSET xa (INSERT x xb) = SUBSET (DELETE xa x) xb" +lemma SUBSET_INSERT_DELETE: "((xa::'A => bool) <= insert (x::'A) (xb::'A => bool)) = (xa - {x} <= xb)" by (import hollight SUBSET_INSERT_DELETE) -lemma DIFF_INSERT: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type. - DIFF x (INSERT xb xa) = DIFF (DELETE x xb) xa" - by (import hollight DIFF_INSERT) - -lemma PSUBSET_INSERT_SUBSET: "ALL (x::'A::type => bool) xa::'A::type => bool. - PSUBSET x xa = (EX xb::'A::type. ~ IN xb x & SUBSET (INSERT xb x) xa)" +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: "ALL (x::'A::type => bool) xa::'A::type => bool. - PSUBSET x xa = (SUBSET x xa & (EX y::'A::type. IN y xa & ~ IN y x))" +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: "ALL (x::'A::type) (y::'A::type) s::'A::type => bool. - DELETE (INSERT x s) y = COND (x = y) (DELETE s y) (INSERT x (DELETE s y))" +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 INSERT_DELETE: "ALL (x::'A::type) xa::'A::type => bool. - IN x xa --> INSERT x (DELETE xa x) = xa" - by (import hollight INSERT_DELETE) - -lemma DELETE_INTER: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type. - hollight.INTER (DELETE x xb) xa = DELETE (hollight.INTER x xa) xb" +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: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type. - DISJOINT (DELETE x xb) xa = DISJOINT (DELETE xa xb) x" +lemma DISJOINT_DELETE_SYM: "(((x::'A => bool) - {xb::'A}) Int (xa::'A => bool) = {}) = +((xa - {xb}) Int x = {})" by (import hollight DISJOINT_DELETE_SYM) -lemma UNIONS_0: "(op =::('q_40008::type => bool) => ('q_40008::type => bool) => bool) - ((UNIONS::(('q_40008::type => bool) => bool) => 'q_40008::type => bool) - (EMPTY::('q_40008::type => bool) => bool)) - (EMPTY::'q_40008::type => bool)" - by (import hollight UNIONS_0) - -lemma UNIONS_1: "UNIONS (INSERT (s::'q_40014::type => bool) EMPTY) = s" - by (import hollight UNIONS_1) - -lemma UNIONS_2: "UNIONS - (INSERT (s::'q_40034::type => bool) - (INSERT (t::'q_40034::type => bool) EMPTY)) = -hollight.UNION s t" - by (import hollight UNIONS_2) - -lemma UNIONS_INSERT: "UNIONS - (INSERT (s::'q_40048::type => bool) - (u::('q_40048::type => bool) => bool)) = -hollight.UNION s (UNIONS u)" - by (import hollight UNIONS_INSERT) - -lemma FORALL_IN_UNIONS: "ALL (x::'q_40090::type => bool) xa::('q_40090::type => bool) => bool. - (ALL xb::'q_40090::type. IN xb (UNIONS xa) --> x xb) = - (ALL (t::'q_40090::type => bool) xb::'q_40090::type. - IN t xa & IN xb t --> x xb)" +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 EMPTY_UNIONS: "ALL x::('q_40116::type => bool) => bool. - (UNIONS x = EMPTY) = - (ALL xa::'q_40116::type => bool. IN xa x --> xa = EMPTY)" +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 INTERS_0: "(op =::('q_40124::type => bool) => ('q_40124::type => bool) => bool) - ((INTERS::(('q_40124::type => bool) => bool) => 'q_40124::type => bool) - (EMPTY::('q_40124::type => bool) => bool)) - (hollight.UNIV::'q_40124::type => bool)" - by (import hollight INTERS_0) - -lemma INTERS_1: "INTERS (INSERT (s::'q_40130::type => bool) EMPTY) = s" - by (import hollight INTERS_1) - -lemma INTERS_2: "INTERS - (INSERT (s::'q_40150::type => bool) - (INSERT (t::'q_40150::type => bool) EMPTY)) = -hollight.INTER s t" - by (import hollight INTERS_2) - -lemma INTERS_INSERT: "INTERS - (INSERT (s::'q_40164::type => bool) - (u::('q_40164::type => bool) => bool)) = -hollight.INTER s (INTERS u)" - by (import hollight INTERS_INSERT) - -lemma IMAGE_CLAUSES: "IMAGE (f::'q_40190::type => 'q_40194::type) EMPTY = EMPTY & -IMAGE f (INSERT (x::'q_40190::type) (s::'q_40190::type => bool)) = -INSERT (f x) (IMAGE f s)" +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_UNION: "ALL (x::'q_40217::type => 'q_40228::type) (xa::'q_40217::type => bool) - xb::'q_40217::type => bool. - IMAGE x (hollight.UNION xa xb) = hollight.UNION (IMAGE x xa) (IMAGE x xb)" - by (import hollight IMAGE_UNION) - -lemma IMAGE_o: "ALL (x::'q_40261::type => 'q_40257::type) - (xa::'q_40252::type => 'q_40261::type) xb::'q_40252::type => bool. - IMAGE (x o xa) xb = IMAGE x (IMAGE xa xb)" - by (import hollight IMAGE_o) - -lemma IMAGE_SUBSET: "ALL (x::'q_40279::type => 'q_40290::type) (xa::'q_40279::type => bool) - xb::'q_40279::type => bool. - SUBSET xa xb --> SUBSET (IMAGE x xa) (IMAGE x xb)" - by (import hollight IMAGE_SUBSET) - -lemma IMAGE_DIFF_INJ: "(ALL (x::'q_40321::type) y::'q_40321::type. - (f::'q_40321::type => 'q_40332::type) x = f y --> x = y) --> -IMAGE f (DIFF (s::'q_40321::type => bool) (t::'q_40321::type => bool)) = -DIFF (IMAGE f s) (IMAGE f t)" +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: "(ALL x::'q_40367::type. - (f::'q_40367::type => 'q_40366::type) x = f (a::'q_40367::type) --> - x = a) --> -IMAGE f (DELETE (s::'q_40367::type => bool) a) = DELETE (IMAGE f s) (f a)" +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 IMAGE_EQ_EMPTY: "ALL (x::'q_40390::type => 'q_40386::type) xa::'q_40390::type => bool. - (IMAGE x xa = EMPTY) = (xa = EMPTY)" - by (import hollight IMAGE_EQ_EMPTY) - -lemma FORALL_IN_IMAGE: "ALL (x::'q_40426::type => 'q_40425::type) xa::'q_40426::type => bool. - (ALL xb::'q_40425::type. - IN xb (IMAGE x xa) --> (P::'q_40425::type => bool) xb) = - (ALL xb::'q_40426::type. IN xb xa --> P (x xb))" +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: "ALL (x::'q_40462::type => 'q_40461::type) xa::'q_40462::type => bool. - (EX xb::'q_40461::type. - IN xb (IMAGE x xa) & (P::'q_40461::type => bool) xb) = - (EX xb::'q_40462::type. IN xb xa & P (x xb))" +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 SUBSET_IMAGE: "ALL (f::'A::type => 'B::type) (s::'B::type => bool) t::'A::type => bool. - SUBSET s (IMAGE f t) = - (EX x::'A::type => bool. SUBSET x t & s = IMAGE f x)" - by (import hollight SUBSET_IMAGE) - -lemma IMAGE_CONST: "ALL (s::'q_40548::type => bool) c::'q_40553::type. - IMAGE (%x::'q_40548::type. c) s = COND (s = EMPTY) EMPTY (INSERT c EMPTY)" - by (import hollight IMAGE_CONST) - -lemma SIMPLE_IMAGE: "ALL (x::'q_40581::type => 'q_40585::type) xa::'q_40581::type => bool. - GSPEC - (%u::'q_40585::type. - EX xb::'q_40581::type. SETSPEC u (IN xb xa) (x xb)) = - IMAGE x xa" +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 EMPTY_GSPEC: "GSPEC (%u::'q_40602::type. Ex (SETSPEC u False)) = EMPTY" +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 IN_ELIM_PAIR_THM: "ALL (x::'q_40648::type => 'q_40647::type => bool) (xa::'q_40648::type) - xb::'q_40647::type. - IN (xa, xb) - (GSPEC - (%xa::'q_40648::type * 'q_40647::type. - EX (xb::'q_40648::type) y::'q_40647::type. - SETSPEC xa (x xb y) (xb, y))) = - x xa xb" +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 FINITE_INDUCT_STRONG: "ALL P::('A::type => bool) => bool. - P EMPTY & - (ALL (x::'A::type) s::'A::type => bool. - P s & ~ IN x s & FINITE s --> P (INSERT x s)) --> - (ALL s::'A::type => bool. FINITE s --> P s)" - by (import hollight FINITE_INDUCT_STRONG) - -lemma FINITE_SUBSET: "ALL (x::'A::type => bool) t::'A::type => bool. - FINITE t & SUBSET x t --> FINITE x" - by (import hollight FINITE_SUBSET) - -lemma FINITE_UNION_IMP: "ALL (x::'A::type => bool) xa::'A::type => bool. - FINITE x & FINITE xa --> FINITE (hollight.UNION x xa)" - by (import hollight FINITE_UNION_IMP) - -lemma FINITE_UNION: "ALL (s::'A::type => bool) t::'A::type => bool. - FINITE (hollight.UNION s t) = (FINITE s & FINITE t)" - by (import hollight FINITE_UNION) - -lemma FINITE_INTER: "ALL (s::'A::type => bool) t::'A::type => bool. - FINITE s | FINITE t --> FINITE (hollight.INTER s t)" - by (import hollight FINITE_INTER) - -lemma FINITE_INSERT: "ALL (s::'A::type => bool) x::'A::type. FINITE (INSERT x s) = FINITE s" - by (import hollight FINITE_INSERT) - -lemma FINITE_DELETE_IMP: "ALL (s::'A::type => bool) x::'A::type. FINITE s --> FINITE (DELETE s x)" +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: "ALL (s::'A::type => bool) x::'A::type. FINITE (DELETE s x) = FINITE s" +lemma FINITE_DELETE: "finite ((s::'A => bool) - {x::'A}) = finite s" by (import hollight FINITE_DELETE) -lemma FINITE_UNIONS: "ALL s::('q_40983::type => bool) => bool. - FINITE s --> - FINITE (UNIONS s) = (ALL t::'q_40983::type => bool. IN t s --> FINITE t)" - by (import hollight FINITE_UNIONS) - -lemma FINITE_IMAGE_EXPAND: "ALL (f::'A::type => 'B::type) s::'A::type => bool. - FINITE s --> - FINITE - (GSPEC - (%u::'B::type. - EX y::'B::type. SETSPEC u (EX x::'A::type. IN x s & y = f x) y))" +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: "ALL (x::'A::type => 'B::type) xa::'A::type => bool. - FINITE xa --> FINITE (IMAGE x xa)" - by (import hollight FINITE_IMAGE) - -lemma FINITE_IMAGE_INJ_GENERAL: "ALL (f::'A::type => 'B::type) (x::'B::type => bool) s::'A::type => bool. - (ALL (x::'A::type) y::'A::type. IN x s & IN y s & f x = f y --> x = y) & - FINITE x --> - FINITE - (GSPEC - (%u::'A::type. EX xa::'A::type. SETSPEC u (IN xa s & IN (f xa) x) xa))" +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_IMAGE_INJ: "ALL (f::'A::type => 'B::type) A::'B::type => bool. - (ALL (x::'A::type) y::'A::type. f x = f y --> x = y) & FINITE A --> - FINITE (GSPEC (%u::'A::type. EX x::'A::type. SETSPEC u (IN (f x) A) x))" +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: "ALL f::'A::type => 'B::type. - (ALL (x::'A::type) y::'A::type. f x = f y --> x = y) --> - (ALL s::'A::type => bool. INFINITE s --> INFINITE (IMAGE f s))" +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 INFINITE_NONEMPTY: "ALL s::'q_41466::type => bool. INFINITE s --> s ~= EMPTY" - by (import hollight INFINITE_NONEMPTY) - -lemma INFINITE_DIFF_FINITE: "ALL (s::'A::type => bool) t::'A::type => bool. - INFINITE s & FINITE t --> INFINITE (DIFF s t)" - by (import hollight INFINITE_DIFF_FINITE) - -lemma FINITE_SUBSET_IMAGE: "ALL (f::'A::type => 'B::type) (s::'A::type => bool) t::'B::type => bool. - (FINITE t & SUBSET t (IMAGE f s)) = - (EX x::'A::type => bool. FINITE x & SUBSET x s & t = IMAGE f x)" +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 FINITE_SUBSET_IMAGE_IMP: "ALL (f::'A::type => 'B::type) (s::'A::type => bool) t::'B::type => bool. - FINITE t & SUBSET t (IMAGE f s) --> - (EX s'::'A::type => bool. - FINITE s' & SUBSET s' s & SUBSET t (IMAGE f s'))" +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) -lemma FINITE_SUBSETS: "ALL s::'A::type => bool. - FINITE s --> - FINITE - (GSPEC - (%u::'A::type => bool. - EX t::'A::type => bool. SETSPEC u (SUBSET t s) t))" - by (import hollight FINITE_SUBSETS) - -lemma FINITE_DIFF: "ALL (s::'q_41764::type => bool) t::'q_41764::type => bool. - FINITE s --> FINITE (DIFF s t)" - by (import hollight FINITE_DIFF) - -definition FINREC :: "('q_41824 => 'q_41823 => 'q_41823) -=> 'q_41823 => ('q_41824 => bool) => 'q_41823 => nat => bool" where +definition + FINREC :: "('A => 'B => 'B) => 'B => ('A => bool) => 'B => nat => bool" where "FINREC == -SOME FINREC::('q_41824::type => 'q_41823::type => 'q_41823::type) - => 'q_41823::type - => ('q_41824::type => bool) - => 'q_41823::type => nat => bool. - (ALL (f::'q_41824::type => 'q_41823::type => 'q_41823::type) - (s::'q_41824::type => bool) (a::'q_41823::type) b::'q_41823::type. - FINREC f b s a 0 = (s = EMPTY & a = b)) & - (ALL (b::'q_41823::type) (s::'q_41824::type => bool) (n::nat) - (a::'q_41823::type) - f::'q_41824::type => 'q_41823::type => 'q_41823::type. +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::'q_41824::type) c::'q_41823::type. - IN x s & FINREC f b (DELETE s x) c n & a = f x c))" + (EX (x::'A) c::'B. x : s & FINREC f b (s - {x}) c n & a = f x c))" lemma DEF_FINREC: "FINREC = -(SOME FINREC::('q_41824::type => 'q_41823::type => 'q_41823::type) - => 'q_41823::type - => ('q_41824::type => bool) - => 'q_41823::type => nat => bool. - (ALL (f::'q_41824::type => 'q_41823::type => 'q_41823::type) - (s::'q_41824::type => bool) (a::'q_41823::type) b::'q_41823::type. - FINREC f b s a 0 = (s = EMPTY & a = b)) & - (ALL (b::'q_41823::type) (s::'q_41824::type => bool) (n::nat) - (a::'q_41823::type) - f::'q_41824::type => 'q_41823::type => 'q_41823::type. +(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::'q_41824::type) c::'q_41823::type. - IN x s & FINREC f b (DELETE s x) c n & a = f x c)))" + (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: "ALL (x::'q_41869::type => 'q_41868::type => 'q_41868::type) - (xa::'q_41868::type) (xb::'q_41869::type => bool) xc::'q_41868::type. - FINREC x xa xb xc (Suc 0) = - (EX xd::'q_41869::type. xb = INSERT xd EMPTY & xc = x xd xa)" +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: "ALL (f::'A::type => 'B::type => 'B::type) b::'B::type. - (ALL (x::'A::type) (y::'A::type) s::'B::type. - x ~= y --> f x (f y s) = f y (f x s)) --> - (ALL (n::nat) (s::'A::type => bool) z::'B::type. - FINREC f b s z (Suc n) --> - (ALL x::'A::type. - IN x s --> - (EX w::'B::type. FINREC f b (DELETE s x) w n & z = f x w)))" +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: "ALL (f::'A::type => 'B::type => 'B::type) b::'B::type. - (ALL (x::'A::type) (y::'A::type) s::'B::type. - x ~= y --> f x (f y s) = f y (f x s)) --> - (ALL (n1::nat) (n2::nat) (s::'A::type => bool) (a1::'B::type) - a2::'B::type. - FINREC f b s a1 n1 & FINREC f b s a2 n2 --> a1 = a2 & n1 = n2)" +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: "ALL (f::'A::type => 'B::type => 'B::type) (b::'B::type) s::'A::type => bool. - FINITE s --> (EX a::'B::type. Ex (FINREC f b s a))" +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 (P::'A::type => bool) R::'A::type => 'B::type => 'C::type => bool. - (ALL s::'A::type. P s --> (EX a::'B::type. Ex (R s a))) & - (ALL (n1::'C::type) (n2::'C::type) (s::'A::type) (a1::'B::type) - a2::'B::type. R s a1 n1 & R s a2 n2 --> a1 = a2 & n1 = n2) --> - (EX x::'A::type => 'B::type. - ALL (s::'A::type) a::'B::type. P s --> Ex (R s a) = (x s = a))" +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: "ALL (f::'A::type => 'B::type => 'B::type) b::'B::type. - (ALL (x::'A::type) (y::'A::type) s::'B::type. - x ~= y --> f x (f y s) = f y (f x s)) --> - (EX g::('A::type => bool) => 'B::type. - g EMPTY = b & - (ALL (s::'A::type => bool) x::'A::type. - FINITE s & IN x s --> g s = f x (g (DELETE s x))))" +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: "ALL (f::'A::type => 'B::type => 'B::type) b::'B::type. - (ALL (x::'A::type) (y::'A::type) s::'B::type. - x ~= y --> f x (f y s) = f y (f x s)) --> - (EX g::('A::type => bool) => 'B::type. - g EMPTY = b & - (ALL (x::'A::type) s::'A::type => bool. - FINITE s --> g (INSERT x s) = COND (IN x s) (g s) (f x (g s))))" +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_42525 => 'q_42524 => 'q_42524) -=> ('q_42525 => bool) => 'q_42524 => 'q_42524" where +definition + ITSET :: "('q_50575 => 'q_50574 => 'q_50574) +=> ('q_50575 => bool) => 'q_50574 => 'q_50574" where "ITSET == -%(u::'q_42525::type => 'q_42524::type => 'q_42524::type) - (ua::'q_42525::type => bool) ub::'q_42524::type. - (SOME g::('q_42525::type => bool) => 'q_42524::type. - g EMPTY = ub & - (ALL (x::'q_42525::type) s::'q_42525::type => bool. - FINITE s --> g (INSERT x s) = COND (IN x s) (g s) (u x (g s)))) +%(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_42525::type => 'q_42524::type => 'q_42524::type) - (ua::'q_42525::type => bool) ub::'q_42524::type. - (SOME g::('q_42525::type => bool) => 'q_42524::type. - g EMPTY = ub & - (ALL (x::'q_42525::type) s::'q_42525::type => bool. - FINITE s --> g (INSERT x s) = COND (IN x s) (g s) (u x (g s)))) +(%(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: "ALL (f::'A::type => 'B::type => 'B::type) b::'B::type. - (ALL (x::'A::type) (y::'A::type) s::'B::type. - x ~= y --> f x (f y s) = f y (f x s)) --> - ITSET f EMPTY b = b & - (ALL (x::'A::type) xa::'A::type => bool. - FINITE xa --> - ITSET f (INSERT x xa) b = - COND (IN x xa) (ITSET f xa b) (f x (ITSET f xa b)))" +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: "ALL (f::'A::type => 'B::type => 'B::type) b::'B::type. - (ALL (x::'A::type) (y::'A::type) s::'B::type. - x ~= y --> f x (f y s) = f y (f x s)) --> - ITSET f EMPTY b = b & - (ALL (x::'A::type) s::'A::type => bool. - FINITE s --> - ITSET f s b = - COND (IN x s) (f x (ITSET f (DELETE s x) b)) - (ITSET f (DELETE s x) b))" +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: "ALL (x::'q_42830::type => bool) - (xa::'q_42830::type => 'q_42831::type => 'q_42831::type) - (xb::'q_42830::type => 'q_42831::type => 'q_42831::type) - xc::'q_42831::type. - FINITE x & - (ALL xc::'q_42830::type. IN xc x --> xa xc = xb xc) & - (ALL (x::'q_42830::type) (y::'q_42830::type) s::'q_42831::type. - x ~= y --> xa x (xa y s) = xa y (xa x s)) & - (ALL (x::'q_42830::type) (y::'q_42830::type) s::'q_42831::type. - x ~= y --> xb x (xb y s) = xb y (xb x s)) --> - ITSET xa x xc = ITSET xb x xc" +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: "ALL (x::'q_42864::type => bool) xa::'q_42864::type => bool. - SUBSET - (GSPEC - (%u::'q_42864::type. - EX xb::'q_42864::type. SETSPEC u (IN xb x & xa xb) xb)) - x" +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: "ALL (s::'A::type => bool) p::'q_42882::type. - FINITE s --> - FINITE - (GSPEC - (%u::'A::type. - EX x::'A::type. SETSPEC u (IN x s & (P::'A::type => bool) x) x))" +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_42918 => bool) => nat" where - "CARD == %u::'q_42918::type => bool. ITSET (%x::'q_42918::type. Suc) u 0" - -lemma DEF_CARD: "CARD = (%u::'q_42918::type => bool. ITSET (%x::'q_42918::type. Suc) u 0)" +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: "(op &::bool => bool => bool) - ((op =::nat => nat => bool) - ((CARD::('A::type => bool) => nat) (EMPTY::'A::type => bool)) (0::nat)) - ((All::('A::type => bool) => bool) - (%x::'A::type. - (All::(('A::type => bool) => bool) => bool) - (%s::'A::type => bool. - (op -->::bool => bool => bool) - ((FINITE::('A::type => bool) => bool) s) - ((op =::nat => nat => bool) - ((CARD::('A::type => bool) => nat) - ((INSERT::'A::type - => ('A::type => bool) => 'A::type => bool) - x s)) - ((COND::bool => nat => nat => nat) - ((IN::'A::type => ('A::type => bool) => bool) x s) - ((CARD::('A::type => bool) => nat) s) - ((Suc::nat => nat) - ((CARD::('A::type => bool) => nat) s)))))))" +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: "ALL (x::'A::type => bool) xa::'A::type => bool. - FINITE x & FINITE xa & hollight.INTER x xa = EMPTY --> - CARD (hollight.UNION x xa) = CARD x + CARD xa" +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: "ALL (x::'A::type) s::'A::type => bool. - FINITE s --> - CARD (DELETE s x) = COND (IN x s) (CARD s - NUMERAL_BIT1 0) (CARD s)" +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: "ALL (s::'q_43163::type => bool) (t::'q_43163::type => bool) - u::'q_43163::type => bool. - FINITE u & hollight.INTER s t = EMPTY & hollight.UNION s t = u --> - CARD s + CARD t = CARD u" +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) -definition HAS_SIZE :: "('q_43199 => bool) => nat => bool" where - "HAS_SIZE == %(u::'q_43199::type => bool) ua::nat. FINITE u & CARD u = ua" - -lemma DEF_HAS_SIZE: "HAS_SIZE = (%(u::'q_43199::type => bool) ua::nat. FINITE u & CARD u = ua)" +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: "ALL (x::'q_43218::type => bool) xa::nat. HAS_SIZE x xa --> CARD x = xa" +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: "ALL (s::'A::type => bool) n::'q_43234::type. HAS_SIZE s 0 = (s = EMPTY)" +lemma HAS_SIZE_0: "HAS_SIZE (s::'A => bool) (0::nat) = (s = {})" by (import hollight HAS_SIZE_0) -lemma HAS_SIZE_SUC: "ALL (s::'A::type => bool) n::nat. - HAS_SIZE s (Suc n) = - (s ~= EMPTY & (ALL x::'A::type. IN x s --> HAS_SIZE (DELETE s x) n))" +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: "ALL (x::'q_43356::type => bool) (xa::'q_43356::type => bool) (xb::nat) - xc::nat. - HAS_SIZE x xb & HAS_SIZE xa xc & DISJOINT x xa --> - HAS_SIZE (hollight.UNION x xa) (xb + xc)" +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_UNIONS: "ALL (x::'A::type => bool) (xa::'A::type => 'B::type => bool) (xb::nat) - xc::nat. - HAS_SIZE x xb & - (ALL xb::'A::type. IN xb x --> HAS_SIZE (xa xb) xc) & - (ALL (xb::'A::type) y::'A::type. - IN xb x & IN y x & xb ~= y --> DISJOINT (xa xb) (xa y)) --> - HAS_SIZE - (UNIONS - (GSPEC - (%u::'B::type => bool. - EX xb::'A::type. SETSPEC u (IN xb x) (xa xb)))) - (xb * xc)" +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 HAS_SIZE_CLAUSES: "HAS_SIZE (s::'q_43604::type => bool) 0 = (s = EMPTY) & +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_43604::type) t::'q_43604::type => bool. - HAS_SIZE t n & ~ IN a t & s = INSERT a t)" +(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: "ALL (a::'A::type => bool) b::'A::type => bool. - FINITE b & SUBSET a b & CARD a = CARD b --> a = b" +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: "ALL (a::'A::type => bool) b::'A::type => bool. - SUBSET a b & FINITE b --> <= (CARD a) (CARD b)" +lemma CARD_SUBSET: "(a::'A => bool) <= (b::'A => bool) & finite b ==> CARD a <= CARD b" by (import hollight CARD_SUBSET) -lemma CARD_SUBSET_LE: "ALL (a::'A::type => bool) b::'A::type => bool. - FINITE b & SUBSET a b & <= (CARD b) (CARD a) --> a = b" +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 CARD_EQ_0: "ALL s::'q_43920::type => bool. FINITE s --> (CARD s = 0) = (s = EMPTY)" - by (import hollight CARD_EQ_0) - -lemma CARD_PSUBSET: "ALL (a::'A::type => bool) b::'A::type => bool. - PSUBSET a b & FINITE b --> < (CARD a) (CARD b)" +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: "ALL (s::'A::type => bool) t::'A::type => bool. - FINITE s & FINITE t --> <= (CARD (hollight.UNION s t)) (CARD s + CARD t)" +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: "ALL (x::'A::type => bool) (xa::'A::type => 'B::type => bool) (xb::nat) - xc::nat. - HAS_SIZE x xb & - (ALL xb::'A::type. IN xb x --> FINITE (xa xb) & <= (CARD (xa xb)) xc) --> - <= (CARD - (UNIONS - (GSPEC - (%u::'B::type => bool. - EX xb::'A::type. SETSPEC u (IN xb x) (xa xb))))) - (xb * xc)" +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_IMAGE_INJ: "ALL (f::'A::type => 'B::type) x::'A::type => bool. - (ALL (xa::'A::type) y::'A::type. - IN xa x & IN y x & f xa = f y --> xa = y) & - FINITE x --> - CARD (IMAGE f x) = CARD x" +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 (x::'A::type => 'B::type) (xa::'A::type => bool) xb::nat. - (ALL (xb::'A::type) y::'A::type. - IN xb xa & IN y xa & x xb = x y --> xb = y) & - HAS_SIZE xa xb --> - HAS_SIZE (IMAGE x xa) xb" +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: "ALL (f::'A::type => 'B::type) s::'A::type => bool. - FINITE s --> <= (CARD (IMAGE f s)) (CARD s)" +lemma CARD_IMAGE_LE: "finite (s::'A => bool) ==> CARD ((f::'A => 'B) ` s) <= CARD s" by (import hollight CARD_IMAGE_LE) -lemma CHOOSE_SUBSET: "ALL s::'A::type => bool. - FINITE s --> - (ALL n::nat. - <= n (CARD s) --> - (EX t::'A::type => bool. SUBSET t s & HAS_SIZE t n))" +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: "ALL (x::'A::type => bool) (xa::nat) (xb::'A::type => 'B::type => bool) - xc::nat. - HAS_SIZE x xa & (ALL xa::'A::type. IN xa x --> HAS_SIZE (xb xa) xc) --> - HAS_SIZE - (GSPEC - (%u::'A::type * 'B::type. - EX (xa::'A::type) y::'B::type. - SETSPEC u (IN xa x & IN y (xb xa)) (xa, y))) - (xa * xc)" +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: "ALL (x::'A::type => bool) xa::'A::type => 'B::type => bool. - FINITE x & (ALL xb::'A::type. IN xb x --> FINITE (xa xb)) --> - FINITE - (GSPEC - (%u::'A::type * 'B::type. - EX (xb::'A::type) y::'B::type. - SETSPEC u (IN xb x & IN y (xa xb)) (xb, y)))" +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: "ALL (x::'A::type => bool) xa::'B::type => bool. - FINITE x & FINITE xa --> - FINITE - (GSPEC - (%u::'A::type * 'B::type. - EX (xb::'A::type) y::'B::type. - SETSPEC u (IN xb x & IN y xa) (xb, y)))" +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: "ALL (s::'A::type => bool) t::'B::type => bool. - FINITE s & FINITE t --> - CARD - (GSPEC - (%u::'A::type * 'B::type. - EX (x::'A::type) y::'B::type. - SETSPEC u (IN x s & IN y t) (x, y))) = - CARD s * CARD t" +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: "ALL (x::'A::type => bool) (xa::nat) (xb::'B::type => bool) xc::nat. - HAS_SIZE x xa & HAS_SIZE xb xc --> - HAS_SIZE - (GSPEC - (%u::'A::type * 'B::type. - EX (xa::'A::type) y::'B::type. - SETSPEC u (IN xa x & IN y xb) (xa, y))) - (xa * xc)" +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) -lemma HAS_SIZE_FUNSPACE: "ALL (d::'B::type) (n::nat) (t::'B::type => bool) (m::nat) - s::'A::type => bool. - HAS_SIZE s m & HAS_SIZE t n --> - HAS_SIZE - (GSPEC - (%u::'A::type => 'B::type. - EX f::'A::type => 'B::type. - SETSPEC u - ((ALL x::'A::type. IN x s --> IN (f x) t) & - (ALL x::'A::type. ~ IN x s --> f x = d)) - f)) - (EXP n m)" +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: "ALL (s::'q_45275::type => bool) t::'q_45272::type => bool. - FINITE s & FINITE t --> - CARD - (GSPEC - (%u::'q_45275::type => 'q_45272::type. - EX f::'q_45275::type => 'q_45272::type. - SETSPEC u - ((ALL x::'q_45275::type. IN x s --> IN (f x) t) & - (ALL x::'q_45275::type. - ~ IN x s --> f x = (d::'q_45272::type))) - f)) = - EXP (CARD t) (CARD s)" +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: "ALL (s::'q_45341::type => bool) t::'q_45338::type => bool. - FINITE s & FINITE t --> - FINITE - (GSPEC - (%u::'q_45341::type => 'q_45338::type. - EX f::'q_45341::type => 'q_45338::type. - SETSPEC u - ((ALL x::'q_45341::type. IN x s --> IN (f x) t) & - (ALL x::'q_45341::type. - ~ IN x s --> f x = (d::'q_45338::type))) - f))" +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: "ALL (s::'A::type => bool) n::nat. - HAS_SIZE s n --> - HAS_SIZE - (GSPEC - (%u::'A::type => bool. - EX t::'A::type => bool. SETSPEC u (SUBSET t s) t)) - (EXP (NUMERAL_BIT0 (NUMERAL_BIT1 0)) n)" +lemma HAS_SIZE_POWERSET: "HAS_SIZE (s::'A => bool) (n::nat) +==> HAS_SIZE {u::'A => bool. EX t<=s. u = t} ((2::nat) ^ n)" by (import hollight HAS_SIZE_POWERSET) -lemma CARD_POWERSET: "ALL s::'A::type => bool. - FINITE s --> - CARD - (GSPEC - (%u::'A::type => bool. - EX t::'A::type => bool. SETSPEC u (SUBSET t s) t)) = - EXP (NUMERAL_BIT0 (NUMERAL_BIT1 0)) (CARD s)" +lemma CARD_POWERSET: "finite (s::'A => bool) +==> CARD {u::'A => bool. EX t<=s. u = t} = (2::nat) ^ CARD s" by (import hollight CARD_POWERSET) -lemma FINITE_POWERSET: "ALL s::'A::type => bool. - FINITE s --> - FINITE - (GSPEC - (%u::'A::type => bool. - EX t::'A::type => bool. SETSPEC u (SUBSET t s) t))" +lemma FINITE_POWERSET: "finite (s::'A => bool) ==> finite {u::'A => bool. EX t<=s. u = t}" by (import hollight FINITE_POWERSET) -lemma CARD_GE_REFL: "ALL s::'A::type => bool. CARD_GE s s" - by (import hollight CARD_GE_REFL) - -lemma CARD_GE_TRANS: "ALL (s::'A::type => bool) (t::'B::type => bool) u::'C::type => bool. - CARD_GE s t & CARD_GE t u --> CARD_GE s u" - by (import hollight CARD_GE_TRANS) - -lemma FINITE_HAS_SIZE_LEMMA: "ALL s::'A::type => bool. - FINITE s --> - (EX n::nat. CARD_GE (GSPEC (%u::nat. EX x::nat. SETSPEC u (< x n) x)) s)" - by (import hollight FINITE_HAS_SIZE_LEMMA) - -lemma HAS_SIZE_NUMSEG_LT: "ALL n::nat. HAS_SIZE (GSPEC (%u::nat. EX m::nat. SETSPEC u (< m n) m)) n" +lemma FINITE_UNIONS: "finite (Union (s::('A => bool) => bool)) = +(finite s & (ALL t::'A => bool. t : s --> finite t))" + by (import hollight FINITE_UNIONS) + +lemma POWERSET_CLAUSES: "{x::'q_54515 => bool. EX xa<={}. x = xa} = {{}} & +(ALL (x::'A) xa::'A => bool. + {xb::'A => bool. EX xc<=insert x xa. xb = xc} = + {u::'A => bool. EX s<=xa. u = s} Un + insert x ` {u::'A => bool. EX s<=xa. u = s})" + by (import hollight POWERSET_CLAUSES) + +lemma HAS_SIZE_NUMSEG_LT: "HAS_SIZE {u. EX m bool. FINITE s = (EX a::nat. ALL x::nat. IN x s --> <= x a)" +lemma num_FINITE: "finite (s::nat => bool) = (EX a::nat. ALL x::nat. x : s --> x <= a)" by (import hollight num_FINITE) -lemma num_FINITE_AVOID: "ALL s::nat => bool. FINITE s --> (EX a::nat. ~ IN a s)" +lemma num_FINITE_AVOID: "finite (s::nat => bool) ==> EX a::nat. a ~: s" by (import hollight num_FINITE_AVOID) -lemma num_INFINITE: "(INFINITE::(nat => bool) => bool) (hollight.UNIV::nat => bool)" - by (import hollight num_INFINITE) - -lemma HAS_SIZE_INDEX: "ALL (x::'A::type => bool) n::nat. - HAS_SIZE x n --> - (EX f::nat => 'A::type. - (ALL m::nat. < m n --> IN (f m) x) & - (ALL xa::'A::type. IN xa x --> (EX! m::nat. < m n & f m = xa)))" +lemma FINITE_REAL_INTERVAL: "(ALL a. infinite {u. EX x. real_lt a x & u = x}) & +(ALL a. infinite {u. EX x. real_le a x & u = x}) & +(ALL b. infinite {u. EX x. real_lt x b & u = x}) & +(ALL b. infinite {u. EX x. real_le x b & u = x}) & +(ALL x xa. + finite {u. EX xb. (real_lt x xb & real_lt xb xa) & u = xb} = + real_le xa x) & +(ALL a b. + finite {u. EX x. (real_le a x & real_lt x b) & u = x} = real_le b a) & +(ALL a b. + finite {u. EX x. (real_lt a x & real_le x b) & u = x} = real_le b a) & +(ALL a b. + finite {u. EX x. (real_le a x & real_le x b) & u = x} = real_le b a)" + by (import hollight FINITE_REAL_INTERVAL) + +lemma real_INFINITE: "(infinite::(hollight.real => bool) => bool) (UNIV::hollight.real => bool)" + by (import hollight real_INFINITE) + +lemma HAS_SIZE_INDEX: "HAS_SIZE (x::'A => bool) (n::nat) +==> EX f::nat => 'A. + (ALL m (EX! m::nat. m < n & f m = xa))" by (import hollight HAS_SIZE_INDEX) -definition set_of_list :: "'q_45968 hollight.list => 'q_45968 => bool" where - "set_of_list == -SOME set_of_list::'q_45968::type hollight.list => 'q_45968::type => bool. - set_of_list NIL = EMPTY & - (ALL (h::'q_45968::type) t::'q_45968::type hollight.list. - set_of_list (CONS h t) = INSERT h (set_of_list t))" - -lemma DEF_set_of_list: "set_of_list = -(SOME set_of_list::'q_45968::type hollight.list => 'q_45968::type => bool. - set_of_list NIL = EMPTY & - (ALL (h::'q_45968::type) t::'q_45968::type hollight.list. - set_of_list (CONS h t) = INSERT h (set_of_list t)))" - by (import hollight DEF_set_of_list) - -definition list_of_set :: "('q_45986 => bool) => 'q_45986 hollight.list" where - "list_of_set == -%u::'q_45986::type => bool. - SOME l::'q_45986::type hollight.list. - set_of_list l = u & LENGTH l = CARD u" - -lemma DEF_list_of_set: "list_of_set = -(%u::'q_45986::type => bool. - SOME l::'q_45986::type hollight.list. - set_of_list l = u & LENGTH l = CARD u)" - by (import hollight DEF_list_of_set) - -lemma LIST_OF_SET_PROPERTIES: "ALL x::'A::type => bool. - FINITE x --> - set_of_list (list_of_set x) = x & LENGTH (list_of_set x) = CARD x" - by (import hollight LIST_OF_SET_PROPERTIES) - -lemma SET_OF_LIST_OF_SET: "ALL s::'q_46035::type => bool. FINITE s --> set_of_list (list_of_set s) = s" - by (import hollight SET_OF_LIST_OF_SET) - -lemma LENGTH_LIST_OF_SET: "ALL s::'q_46051::type => bool. FINITE s --> LENGTH (list_of_set s) = CARD s" - by (import hollight LENGTH_LIST_OF_SET) - -lemma MEM_LIST_OF_SET: "ALL s::'A::type => bool. - FINITE s --> (ALL x::'A::type. MEM x (list_of_set s) = IN x s)" - by (import hollight MEM_LIST_OF_SET) - -lemma FINITE_SET_OF_LIST: "ALL l::'q_46096::type hollight.list. FINITE (set_of_list l)" - by (import hollight FINITE_SET_OF_LIST) - -lemma IN_SET_OF_LIST: "ALL (x::'q_46114::type) l::'q_46114::type hollight.list. - IN x (set_of_list l) = MEM x l" - by (import hollight IN_SET_OF_LIST) - -lemma SET_OF_LIST_APPEND: "ALL (x::'q_46139::type hollight.list) xa::'q_46139::type hollight.list. - set_of_list (APPEND x xa) = - hollight.UNION (set_of_list x) (set_of_list xa)" - by (import hollight SET_OF_LIST_APPEND) - -definition pairwise :: "('q_46198 => 'q_46198 => bool) => ('q_46198 => bool) => bool" where +definition + pairwise :: "('q_55938 => 'q_55938 => bool) => ('q_55938 => bool) => bool" where "pairwise == -%(u::'q_46198::type => 'q_46198::type => bool) ua::'q_46198::type => bool. - ALL (x::'q_46198::type) y::'q_46198::type. - IN x ua & IN y ua & x ~= y --> u x y" +%(u::'q_55938 => 'q_55938 => bool) ua::'q_55938 => bool. + ALL (x::'q_55938) y::'q_55938. x : ua & y : ua & x ~= y --> u x y" lemma DEF_pairwise: "pairwise = -(%(u::'q_46198::type => 'q_46198::type => bool) ua::'q_46198::type => bool. - ALL (x::'q_46198::type) y::'q_46198::type. - IN x ua & IN y ua & x ~= y --> u x y)" +(%(u::'q_55938 => 'q_55938 => bool) ua::'q_55938 => bool. + ALL (x::'q_55938) y::'q_55938. x : ua & y : ua & x ~= y --> u x y)" by (import hollight DEF_pairwise) -definition PAIRWISE :: "('q_46220 => 'q_46220 => bool) => 'q_46220 hollight.list => bool" where +definition + PAIRWISE :: "('A => 'A => bool) => 'A list => bool" where "PAIRWISE == -SOME PAIRWISE::('q_46220::type => 'q_46220::type => bool) - => 'q_46220::type hollight.list => bool. - (ALL r::'q_46220::type => 'q_46220::type => bool. - PAIRWISE r NIL = True) & - (ALL (h::'q_46220::type) (r::'q_46220::type => 'q_46220::type => bool) - t::'q_46220::type hollight.list. - PAIRWISE r (CONS h t) = (ALL_list (r h) t & PAIRWISE r t))" +SOME PAIRWISE::('A => 'A => bool) => 'A list => bool. + (ALL r::'A => 'A => bool. PAIRWISE r [] = True) & + (ALL (h::'A) (r::'A => 'A => bool) t::'A list. + PAIRWISE r (h # t) = (list_all (r h) t & PAIRWISE r t))" lemma DEF_PAIRWISE: "PAIRWISE = -(SOME PAIRWISE::('q_46220::type => 'q_46220::type => bool) - => 'q_46220::type hollight.list => bool. - (ALL r::'q_46220::type => 'q_46220::type => bool. - PAIRWISE r NIL = True) & - (ALL (h::'q_46220::type) (r::'q_46220::type => 'q_46220::type => bool) - t::'q_46220::type hollight.list. - PAIRWISE r (CONS h t) = (ALL_list (r h) t & PAIRWISE r t)))" +(SOME PAIRWISE::('A => 'A => bool) => 'A list => bool. + (ALL r::'A => 'A => bool. PAIRWISE r [] = True) & + (ALL (h::'A) (r::'A => 'A => bool) t::'A list. + PAIRWISE r (h # t) = (list_all (r h) t & PAIRWISE r t)))" by (import hollight DEF_PAIRWISE) -typedef (open) ('A) finite_image = "(Collect::('A::type => bool) => 'A::type set) - (%x::'A::type. - (op |::bool => bool => bool) - ((op =::'A::type => 'A::type => bool) x - ((Eps::('A::type => bool) => 'A::type) (%z::'A::type. True::bool))) - ((FINITE::('A::type => bool) => bool) - (hollight.UNIV::'A::type => bool)))" morphisms "dest_finite_image" "mk_finite_image" - apply (rule light_ex_imp_nonempty[where t="(Eps::('A::type => bool) => 'A::type) - (%x::'A::type. - (op |::bool => bool => bool) - ((op =::'A::type => 'A::type => bool) x - ((Eps::('A::type => bool) => 'A::type) (%z::'A::type. True::bool))) - ((FINITE::('A::type => bool) => bool) - (hollight.UNIV::'A::type => bool)))"]) +lemma PAIRWISE_EMPTY: "pairwise (x::'q_55973 => 'q_55973 => bool) {} = True" + by (import hollight PAIRWISE_EMPTY) + +lemma PAIRWISE_SING: "pairwise (x::'q_55991 => 'q_55991 => bool) {xa::'q_55991} = True" + by (import hollight PAIRWISE_SING) + +lemma PAIRWISE_MONO: "pairwise (x::'q_56011 => 'q_56011 => bool) (xa::'q_56011 => bool) & +(xb::'q_56011 => bool) <= xa +==> pairwise x xb" + by (import hollight PAIRWISE_MONO) + +lemma SURJECTIVE_IFF_INJECTIVE_GEN: "finite (s::'A => bool) & +finite (t::'B => bool) & CARD s = CARD t & (f::'A => 'B) ` s <= t +==> (ALL y::'B. y : t --> (EX x::'A. x : s & f x = y)) = + (ALL (x::'A) y::'A. x : s & y : s & f x = f y --> x = y)" + by (import hollight SURJECTIVE_IFF_INJECTIVE_GEN) + +lemma SURJECTIVE_IFF_INJECTIVE: "finite (x::'A => bool) & (xa::'A => 'A) ` x <= x +==> (ALL y::'A. y : x --> (EX xb::'A. xb : x & xa xb = y)) = + (ALL (xb::'A) y::'A. xb : x & y : x & xa xb = xa y --> xb = y)" + by (import hollight SURJECTIVE_IFF_INJECTIVE) + +lemma IMAGE_IMP_INJECTIVE_GEN: "[| finite (s::'A => bool) & + CARD s = CARD (t::'B => bool) & (f::'A => 'B) ` s = t; + (x::'A) : s & (y::'A) : s & f x = f y |] +==> x = y" + by (import hollight IMAGE_IMP_INJECTIVE_GEN) + +lemma IMAGE_IMP_INJECTIVE: "[| finite (s::'q_56387 => bool) & (f::'q_56387 => 'q_56387) ` s = s; + (x::'q_56387) : s & (y::'q_56387) : s & f x = f y |] +==> x = y" + by (import hollight IMAGE_IMP_INJECTIVE) + +lemma CARD_LE_INJ: "finite (x::'A => bool) & finite (xa::'B => bool) & CARD x <= CARD xa +==> EX f::'A => 'B. + f ` x <= xa & + (ALL (xa::'A) y::'A. xa : x & y : x & f xa = f y --> xa = y)" + by (import hollight CARD_LE_INJ) + +lemma FORALL_IN_CLAUSES: "(ALL x::'q_56493 => bool. (ALL xa::'q_56493. xa : {} --> x xa) = True) & +(ALL (x::'q_56533 => bool) (xa::'q_56533) xb::'q_56533 => bool. + (ALL xc::'q_56533. xc : insert xa xb --> x xc) = + (x xa & (ALL xa::'q_56533. xa : xb --> x xa)))" + by (import hollight FORALL_IN_CLAUSES) + +lemma EXISTS_IN_CLAUSES: "(ALL x::'q_56553 => bool. (EX xa::'q_56553. xa : {} & x xa) = False) & +(ALL (x::'q_56593 => bool) (xa::'q_56593) xb::'q_56593 => bool. + (EX xc::'q_56593. xc : insert xa xb & x xc) = + (x xa | (EX xa::'q_56593. xa : xb & x xa)))" + by (import hollight EXISTS_IN_CLAUSES) + +lemma SURJECTIVE_ON_RIGHT_INVERSE: "(ALL xb::'q_56650. + xb : (xa::'q_56650 => bool) --> + (EX xa::'q_56649. + xa : (s::'q_56649 => bool) & (x::'q_56649 => 'q_56650) xa = xb)) = +(EX g::'q_56650 => 'q_56649. + ALL y::'q_56650. y : xa --> g y : s & x (g y) = y)" + by (import hollight SURJECTIVE_ON_RIGHT_INVERSE) + +lemma INJECTIVE_ON_LEFT_INVERSE: "(ALL (xb::'q_56743) y::'q_56743. + xb : (xa::'q_56743 => bool) & + y : xa & (x::'q_56743 => 'q_56746) xb = x y --> + xb = y) = +(EX xb::'q_56746 => 'q_56743. ALL xc::'q_56743. xc : xa --> xb (x xc) = xc)" + by (import hollight INJECTIVE_ON_LEFT_INVERSE) + +lemma BIJECTIVE_ON_LEFT_RIGHT_INVERSE: "(!!x::'q_56878. + x : (s::'q_56878 => bool) + ==> (f::'q_56878 => 'q_56877) x : (t::'q_56877 => bool)) +==> ((ALL (x::'q_56878) y::'q_56878. x : s & y : s & f x = f y --> x = y) & + (ALL x::'q_56877. x : t --> (EX xa::'q_56878. xa : s & f xa = x))) = + (EX g::'q_56877 => 'q_56878. + (ALL y::'q_56877. y : t --> g y : s) & + (ALL y::'q_56877. y : t --> f (g y) = y) & + (ALL x::'q_56878. x : s --> g (f x) = x))" + by (import hollight BIJECTIVE_ON_LEFT_RIGHT_INVERSE) + +lemma SURJECTIVE_RIGHT_INVERSE: "(ALL y::'q_56902. EX x::'q_56905. (f::'q_56905 => 'q_56902) x = y) = +(EX g::'q_56902 => 'q_56905. ALL y::'q_56902. f (g y) = y)" + by (import hollight SURJECTIVE_RIGHT_INVERSE) + +lemma INJECTIVE_LEFT_INVERSE: "(ALL (x::'q_56939) xa::'q_56939. + (f::'q_56939 => 'q_56942) x = f xa --> x = xa) = +(EX g::'q_56942 => 'q_56939. ALL x::'q_56939. g (f x) = x)" + by (import hollight INJECTIVE_LEFT_INVERSE) + +lemma BIJECTIVE_LEFT_RIGHT_INVERSE: "((ALL (x::'A) y::'A. (f::'A => 'B) x = f y --> x = y) & + (ALL y::'B. EX x::'A. f x = y)) = +(EX g::'B => 'A. (ALL y::'B. f (g y) = y) & (ALL x::'A. g (f x) = x))" + by (import hollight BIJECTIVE_LEFT_RIGHT_INVERSE) + +lemma FUNCTION_FACTORS_RIGHT: "(ALL xb::'q_57046. + EX y::'q_57034. + (xa::'q_57034 => 'q_57047) y = (x::'q_57046 => 'q_57047) xb) = +(EX xb::'q_57046 => 'q_57034. x = xa o xb)" + by (import hollight FUNCTION_FACTORS_RIGHT) + +lemma FUNCTION_FACTORS_LEFT: "(ALL (xb::'q_57119) y::'q_57119. + (xa::'q_57119 => 'q_57099) xb = xa y --> + (x::'q_57119 => 'q_57120) xb = x y) = +(EX xb::'q_57099 => 'q_57120. x = xb o xa)" + by (import hollight FUNCTION_FACTORS_LEFT) + +lemma SURJECTIVE_FORALL_THM: "(ALL y::'B. EX x::'A. (f::'A => 'B) x = y) = +(ALL P::'B => bool. (ALL x::'A. P (f x)) = All P)" + by (import hollight SURJECTIVE_FORALL_THM) + +lemma SURJECTIVE_EXISTS_THM: "(ALL y::'B. EX x::'A. (f::'A => 'B) x = y) = +(ALL P::'B => bool. (EX x::'A. P (f x)) = Ex P)" + by (import hollight SURJECTIVE_EXISTS_THM) + +lemma SURJECTIVE_IMAGE_THM: "(ALL y::'B. EX x::'A. (f::'A => 'B) x = y) = +(ALL x::'B => bool. + f ` {u::'A. EX xa::'A. x (f xa) & u = xa} = + {u::'B. EX xa::'B. x xa & u = xa})" + by (import hollight SURJECTIVE_IMAGE_THM) + +lemma IMAGE_INJECTIVE_IMAGE_OF_SUBSET: "EX x<=s::'A => bool. + (f::'A => 'B) ` s = f ` x & + (ALL (xa::'A) y::'A. xa : x & y : x & f xa = f y --> xa = y)" + by (import hollight IMAGE_INJECTIVE_IMAGE_OF_SUBSET) + +lemma INJECTIVE_ON_IMAGE: "(ALL (s::'A => bool) t::'A => bool. + s <= (u::'A => bool) & t <= u & (f::'A => 'B) ` s = f ` t --> s = t) = +(ALL (x::'A) y::'A. x : u & y : u & f x = f y --> x = y)" + by (import hollight INJECTIVE_ON_IMAGE) + +lemma INJECTIVE_IMAGE: "(ALL (s::'A => bool) t::'A => bool. (f::'A => 'B) ` s = f ` t --> s = t) = +(ALL (x::'A) y::'A. f x = f y --> x = y)" + by (import hollight INJECTIVE_IMAGE) + +lemma SURJECTIVE_ON_IMAGE: "(ALL t<=v::'B => bool. EX s<=u::'A => bool. (f::'A => 'B) ` s = t) = +(ALL y::'B. y : v --> (EX x::'A. x : u & f x = y))" + by (import hollight SURJECTIVE_ON_IMAGE) + +lemma SURJECTIVE_IMAGE: "(ALL t::'B => bool. EX s::'A => bool. (f::'A => 'B) ` s = t) = +(ALL y::'B. EX x::'A. f x = y)" + by (import hollight SURJECTIVE_IMAGE) + +lemma CARD_EQ_BIJECTION: "finite (s::'A => bool) & finite (t::'B => bool) & CARD s = CARD t +==> EX f::'A => 'B. + (ALL x::'A. x : s --> f x : t) & + (ALL y::'B. y : t --> (EX x::'A. x : s & f x = y)) & + (ALL (x::'A) y::'A. x : s & y : s & f x = f y --> x = y)" + by (import hollight CARD_EQ_BIJECTION) + +lemma CARD_EQ_BIJECTIONS: "finite (s::'A => bool) & finite (t::'B => bool) & CARD s = CARD t +==> EX (f::'A => 'B) g::'B => 'A. + (ALL x::'A. x : s --> f x : t & g (f x) = x) & + (ALL y::'B. y : t --> g y : s & f (g y) = y)" + by (import hollight CARD_EQ_BIJECTIONS) + +lemma BIJECTIONS_HAS_SIZE: "(ALL x::'A. + x : (s::'A => bool) --> + (f::'A => 'B) x : (t::'B => bool) & (g::'B => 'A) (f x) = x) & +(ALL y::'B. y : t --> g y : s & f (g y) = y) & HAS_SIZE s (n::nat) +==> HAS_SIZE t n" + by (import hollight BIJECTIONS_HAS_SIZE) + +lemma BIJECTIONS_HAS_SIZE_EQ: "(ALL x::'A. + x : (s::'A => bool) --> + (f::'A => 'B) x : (t::'B => bool) & (g::'B => 'A) (f x) = x) & +(ALL y::'B. y : t --> g y : s & f (g y) = y) +==> HAS_SIZE s (n::nat) = HAS_SIZE t n" + by (import hollight BIJECTIONS_HAS_SIZE_EQ) + +lemma BIJECTIONS_CARD_EQ: "(finite (s::'A => bool) | finite (t::'B => bool)) & +(ALL x::'A. x : s --> (f::'A => 'B) x : t & (g::'B => 'A) (f x) = x) & +(ALL y::'B. y : t --> g y : s & f (g y) = y) +==> CARD s = CARD t" + by (import hollight BIJECTIONS_CARD_EQ) + +lemma WF_FINITE: "(ALL x::'A. ~ (u_556::'A => 'A => bool) x x) & +(ALL (x::'A) (y::'A) z::'A. u_556 x y & u_556 y z --> u_556 x z) & +(ALL x::'A. finite {u::'A. EX y::'A. u_556 y x & u = y}) +==> wfP u_556" + by (import hollight WF_FINITE) + +consts + "<=_c" :: "('q_58200 => bool) => ('q_58195 => bool) => bool" ("<='_c") + +defs + "<=_c_def": "<=_c == +%(u::'q_58200 => bool) ua::'q_58195 => bool. + EX f::'q_58200 => 'q_58195. + (ALL x::'q_58200. x : u --> f x : ua) & + (ALL (x::'q_58200) y::'q_58200. x : u & y : u & f x = f y --> x = y)" + +lemma DEF__lessthan__equal__c: "<=_c = +(%(u::'q_58200 => bool) ua::'q_58195 => bool. + EX f::'q_58200 => 'q_58195. + (ALL x::'q_58200. x : u --> f x : ua) & + (ALL (x::'q_58200) y::'q_58200. x : u & y : u & f x = f y --> x = y))" + by (import hollight DEF__lessthan__equal__c) + +consts + "<_c" :: "('q_58212 => bool) => ('q_58213 => bool) => bool" ("<'_c") + +defs + "<_c_def": "<_c == %(u::'q_58212 => bool) ua::'q_58213 => bool. <=_c u ua & ~ <=_c ua u" + +lemma DEF__lessthan__c: "<_c = (%(u::'q_58212 => bool) ua::'q_58213 => bool. <=_c u ua & ~ <=_c ua u)" + by (import hollight DEF__lessthan__c) + +consts + "=_c" :: "('q_58264 => bool) => ('q_58261 => bool) => bool" ("='_c") + +defs + "=_c_def": "=_c == +%(u::'q_58264 => bool) ua::'q_58261 => bool. + EX f::'q_58264 => 'q_58261. + (ALL x::'q_58264. x : u --> f x : ua) & + (ALL y::'q_58261. y : ua --> (EX! x::'q_58264. x : u & f x = y))" + +lemma DEF__equal__c: "=_c = +(%(u::'q_58264 => bool) ua::'q_58261 => bool. + EX f::'q_58264 => 'q_58261. + (ALL x::'q_58264. x : u --> f x : ua) & + (ALL y::'q_58261. y : ua --> (EX! x::'q_58264. x : u & f x = y)))" + by (import hollight DEF__equal__c) + +consts + ">=_c" :: "('q_58273 => bool) => ('q_58272 => bool) => bool" (">='_c") + +defs + ">=_c_def": ">=_c == %(u::'q_58273 => bool) ua::'q_58272 => bool. <=_c ua u" + +lemma DEF__greaterthan__equal__c: ">=_c = (%(u::'q_58273 => bool) ua::'q_58272 => bool. <=_c ua u)" + by (import hollight DEF__greaterthan__equal__c) + +consts + ">_c" :: "('q_58282 => bool) => ('q_58281 => bool) => bool" (">'_c") + +defs + ">_c_def": ">_c == %(u::'q_58282 => bool) ua::'q_58281 => bool. <_c ua u" + +lemma DEF__greaterthan__c: ">_c = (%(u::'q_58282 => bool) ua::'q_58281 => bool. <_c ua u)" + by (import hollight DEF__greaterthan__c) + +lemma LE_C: "<=_c (x::'q_58320 => bool) (xa::'q_58323 => bool) = +(EX xb::'q_58323 => 'q_58320. + ALL xc::'q_58320. xc : x --> (EX x::'q_58323. x : xa & xb x = xc))" + by (import hollight LE_C) + +lemma GE_C: ">=_c (x::'q_58364 => bool) (xa::'q_58361 => bool) = +(EX f::'q_58364 => 'q_58361. + ALL y::'q_58361. y : xa --> (EX xa::'q_58364. xa : x & y = f xa))" + by (import hollight GE_C) + +definition + COUNTABLE :: "('q_58372 => bool) => bool" where + "(op ==::(('q_58372::type => bool) => bool) + => (('q_58372::type => bool) => bool) => prop) + (COUNTABLE::('q_58372::type => bool) => bool) + ((>=_c::(nat => bool) => ('q_58372::type => bool) => bool) + (UNIV::nat => bool))" + +lemma DEF_COUNTABLE: "(op =::(('q_58372::type => bool) => bool) + => (('q_58372::type => bool) => bool) => bool) + (COUNTABLE::('q_58372::type => bool) => bool) + ((>=_c::(nat => bool) => ('q_58372::type => bool) => bool) + (UNIV::nat => bool))" + by (import hollight DEF_COUNTABLE) + +lemma NUMSEG_COMBINE_R: "(x::nat) <= (xa::nat) + (1::nat) & xa <= (xb::nat) +==> {x..xa} Un {xa + (1::nat)..xb} = {x..xb}" + by (import hollight NUMSEG_COMBINE_R) + +lemma NUMSEG_COMBINE_L: "(x::nat) <= (xa::nat) & xa <= (xb::nat) + (1::nat) +==> {x..xa - (1::nat)} Un {xa..xb} = {x..xb}" + by (import hollight NUMSEG_COMBINE_L) + +lemma NUMSEG_LREC: "(x::nat) <= (xa::nat) ==> insert x {x + (1::nat)..xa} = {x..xa}" + by (import hollight NUMSEG_LREC) + +lemma NUMSEG_RREC: "(x::nat) <= (xa::nat) ==> insert xa {x..xa - (1::nat)} = {x..xa}" + by (import hollight NUMSEG_RREC) + +lemma IN_NUMSEG_0: "((x::nat) : {0::nat..xa::nat}) = (x <= xa)" + by (import hollight IN_NUMSEG_0) + +lemma NUMSEG_EMPTY: "({x::nat..xa::nat} = {}) = (xa < x)" + by (import hollight NUMSEG_EMPTY) + +lemma CARD_NUMSEG_LEMMA: "CARD {m..m + d} = d + 1" + by (import hollight CARD_NUMSEG_LEMMA) + +lemma CARD_NUMSEG: "CARD {m..n} = n + 1 - m" + by (import hollight CARD_NUMSEG) + +lemma HAS_SIZE_NUMSEG: "HAS_SIZE {x..xa} (xa + 1 - x)" + by (import hollight HAS_SIZE_NUMSEG) + +lemma CARD_NUMSEG_1: "CARD {1..x} = x" + by (import hollight CARD_NUMSEG_1) + +lemma HAS_SIZE_NUMSEG_1: "HAS_SIZE {1..x} x" + by (import hollight HAS_SIZE_NUMSEG_1) + +lemma NUMSEG_CLAUSES: "(ALL m::nat. {m..0::nat} = (if m = (0::nat) then {0::nat} else {})) & +(ALL (m::nat) n::nat. + {m..Suc n} = (if m <= Suc n then insert (Suc n) {m..n} else {m..n}))" + by (import hollight NUMSEG_CLAUSES) + +lemma FINITE_INDEX_NUMSEG: "finite (s::'A => bool) = +(EX f::nat => 'A. + (ALL (i::nat) j::nat. + i : {1::nat..CARD s} & j : {1::nat..CARD s} & f i = f j --> i = j) & + s = f ` {1::nat..CARD s})" + by (import hollight FINITE_INDEX_NUMSEG) + +lemma FINITE_INDEX_NUMBERS: "finite (s::'A => bool) = +(EX (k::nat => bool) f::nat => 'A. + (ALL (i::nat) j::nat. i : k & j : k & f i = f j --> i = j) & + finite k & s = f ` k)" + by (import hollight FINITE_INDEX_NUMBERS) + +lemma DISJOINT_NUMSEG: "({x::nat..xa::nat} Int {xb::nat..xc::nat} = {}) = +(xa < xb | xc < x | xa < x | xc < xb)" + by (import hollight DISJOINT_NUMSEG) + +lemma NUMSEG_ADD_SPLIT: "(x::nat) <= (xa::nat) + (1::nat) +==> {x..xa + (xb::nat)} = {x..xa} Un {xa + (1::nat)..xa + xb}" + by (import hollight NUMSEG_ADD_SPLIT) + +lemma SUBSET_NUMSEG: "({m::nat..n::nat} <= {p::nat..q::nat}) = (n < m | p <= m & n <= q)" + by (import hollight SUBSET_NUMSEG) + +lemma NUMSEG_LE: "{u::nat. EX xa<=x::nat. u = xa} = {0::nat..x}" + by (import hollight NUMSEG_LE) + +lemma NUMSEG_LT: "{u::nat. EX x 'A => bool) x y & u_556 y x --> x = y) & + (ALL (x::'A) (y::'A) z::'A. u_556 x y & u_556 y z --> u_556 x z); + HAS_SIZE (s::'A => bool) (n::nat) |] +==> EX f::nat => 'A. + s = f ` {1::nat..n} & + (ALL (j::nat) k::nat. + j : {1::nat..n} & k : {1::nat..n} & j < k --> + ~ u_556 (f k) (f j))" + by (import hollight TOPOLOGICAL_SORT) + +lemma FINITE_INTSEG: "(ALL l r. finite {u. EX x. (int_le l x & int_le x r) & u = x}) & +(ALL l r. finite {u. EX x. (int_le l x & int_lt x r) & u = x}) & +(ALL l r. finite {u. EX x. (int_lt l x & int_le x r) & u = x}) & +(ALL l r. finite {u. EX x. (int_lt l x & int_lt x r) & u = x})" + by (import hollight FINITE_INTSEG) + +definition + neutral :: "('q_59899 => 'q_59899 => 'q_59899) => 'q_59899" where + "neutral == +%u::'q_59899 => 'q_59899 => 'q_59899. + SOME x::'q_59899. ALL y::'q_59899. u x y = y & u y x = y" + +lemma DEF_neutral: "neutral = +(%u::'q_59899 => 'q_59899 => 'q_59899. + SOME x::'q_59899. ALL y::'q_59899. u x y = y & u y x = y)" + by (import hollight DEF_neutral) + +definition + monoidal :: "('A => 'A => 'A) => bool" where + "monoidal == +%u::'A => 'A => 'A. + (ALL (x::'A) y::'A. u x y = u y x) & + (ALL (x::'A) (y::'A) z::'A. u x (u y z) = u (u x y) z) & + (ALL x::'A. u (neutral u) x = x)" + +lemma DEF_monoidal: "monoidal = +(%u::'A => 'A => 'A. + (ALL (x::'A) y::'A. u x y = u y x) & + (ALL (x::'A) (y::'A) z::'A. u x (u y z) = u (u x y) z) & + (ALL x::'A. u (neutral u) x = x))" + by (import hollight DEF_monoidal) + +lemma MONOIDAL_AC: "monoidal (x::'q_60055 => 'q_60055 => 'q_60055) +==> (ALL a::'q_60055. x (neutral x) a = a) & + (ALL a::'q_60055. x a (neutral x) = a) & + (ALL (a::'q_60055) b::'q_60055. x a b = x b a) & + (ALL (a::'q_60055) (b::'q_60055) c::'q_60055. + x (x a b) c = x a (x b c)) & + (ALL (a::'q_60055) (b::'q_60055) c::'q_60055. x a (x b c) = x b (x a c))" + by (import hollight MONOIDAL_AC) + +definition + support :: "('B => 'B => 'B) => ('A => 'B) => ('A => bool) => 'A => bool" where + "support == +%(u::'B => 'B => 'B) (ua::'A => 'B) ub::'A => bool. + {uc::'A. EX x::'A. (x : ub & ua x ~= neutral u) & uc = x}" + +lemma DEF_support: "support = +(%(u::'B => 'B => 'B) (ua::'A => 'B) ub::'A => bool. + {uc::'A. EX x::'A. (x : ub & ua x ~= neutral u) & uc = x})" + by (import hollight DEF_support) + +definition + iterate :: "('q_60113 => 'q_60113 => 'q_60113) +=> ('A => bool) => ('A => 'q_60113) => 'q_60113" where + "iterate == +%(u::'q_60113 => 'q_60113 => 'q_60113) (ua::'A => bool) ub::'A => 'q_60113. + if finite (support u ub ua) + then ITSET (%x::'A. u (ub x)) (support u ub ua) (neutral u) + else neutral u" + +lemma DEF_iterate: "iterate = +(%(u::'q_60113 => 'q_60113 => 'q_60113) (ua::'A => bool) ub::'A => 'q_60113. + if finite (support u ub ua) + then ITSET (%x::'A. u (ub x)) (support u ub ua) (neutral u) + else neutral u)" + by (import hollight DEF_iterate) + +lemma IN_SUPPORT: "((xb::'q_60163) + : support (x::'q_60160 => 'q_60160 => 'q_60160) (xa::'q_60163 => 'q_60160) + (xc::'q_60163 => bool)) = +(xb : xc & xa xb ~= neutral x)" + by (import hollight IN_SUPPORT) + +lemma SUPPORT_SUPPORT: "support (x::'q_60185 => 'q_60185 => 'q_60185) (xa::'q_60196 => 'q_60185) + (support x xa (xb::'q_60196 => bool)) = +support x xa xb" + by (import hollight SUPPORT_SUPPORT) + +lemma SUPPORT_EMPTY: "(ALL xc::'q_60235. + xc : (xb::'q_60235 => bool) --> + (xa::'q_60235 => 'q_60221) xc = + neutral (x::'q_60221 => 'q_60221 => 'q_60221)) = +(support x xa xb = {})" + by (import hollight SUPPORT_EMPTY) + +lemma SUPPORT_SUBSET: "support (x::'q_60255 => 'q_60255 => 'q_60255) (xa::'q_60256 => 'q_60255) + (xb::'q_60256 => bool) +<= xb" + by (import hollight SUPPORT_SUBSET) + +lemma FINITE_SUPPORT: "finite (s::'q_60273 => bool) +==> finite + (support (u::'q_60279 => 'q_60279 => 'q_60279) + (f::'q_60273 => 'q_60279) s)" + by (import hollight FINITE_SUPPORT) + +lemma SUPPORT_CLAUSES: "(ALL x::'q_60297 => 'q_60530. + support (u_4372::'q_60530 => 'q_60530 => 'q_60530) x {} = {}) & +(ALL (x::'q_60345 => 'q_60530) (xa::'q_60345) xb::'q_60345 => bool. + support u_4372 x (insert xa xb) = + (if x xa = neutral u_4372 then support u_4372 x xb + else insert xa (support u_4372 x xb))) & +(ALL (x::'q_60378 => 'q_60530) (xa::'q_60378) xb::'q_60378 => bool. + support u_4372 x (xb - {xa}) = support u_4372 x xb - {xa}) & +(ALL (x::'q_60416 => 'q_60530) (xa::'q_60416 => bool) xb::'q_60416 => bool. + support u_4372 x (xa Un xb) = + support u_4372 x xa Un support u_4372 x xb) & +(ALL (x::'q_60454 => 'q_60530) (xa::'q_60454 => bool) xb::'q_60454 => bool. + support u_4372 x (xa Int xb) = + support u_4372 x xa Int support u_4372 x xb) & +(ALL (x::'q_60492 => 'q_60530) (xa::'q_60492 => bool) xb::'q_60492 => bool. + support u_4372 x (xa - xb) = + support u_4372 x xa - support u_4372 x xb) & +(ALL (x::'q_60529 => 'q_60520) (xa::'q_60520 => 'q_60530) + xb::'q_60529 => bool. + support u_4372 xa (x ` xb) = x ` support u_4372 (xa o x) xb)" + by (import hollight SUPPORT_CLAUSES) + +lemma SUPPORT_DELTA: "support (x::'q_60556 => 'q_60556 => 'q_60556) + (%xa::'q_60584. + if xa = (xc::'q_60584) then (xb::'q_60584 => 'q_60556) xa + else neutral x) + (xa::'q_60584 => bool) = +(if xc : xa then support x xb {xc} else {})" + by (import hollight SUPPORT_DELTA) + +lemma FINITE_SUPPORT_DELTA: "finite + (support (x::'q_60605 => 'q_60605 => 'q_60605) + (%xc::'q_60614. + if xc = (xb::'q_60614) then (xa::'q_60614 => 'q_60605) xc + else neutral x) + (s::'q_60614 => bool))" + by (import hollight FINITE_SUPPORT_DELTA) + +lemma ITERATE_SUPPORT: "iterate (x::'q_60630 => 'q_60630 => 'q_60630) + (support x (xa::'q_60642 => 'q_60630) (xb::'q_60642 => bool)) xa = +iterate x xb xa" + by (import hollight ITERATE_SUPPORT) + +lemma ITERATE_EXPAND_CASES: "iterate (x::'q_60661 => 'q_60661 => 'q_60661) (xb::'q_60667 => bool) + (xa::'q_60667 => 'q_60661) = +(if finite (support x xa xb) then iterate x (support x xa xb) xa + else neutral x)" + by (import hollight ITERATE_EXPAND_CASES) + +lemma ITERATE_CLAUSES_GEN: "monoidal (u_4372::'B => 'B => 'B) +==> (ALL f::'A => 'B. iterate u_4372 {} f = neutral u_4372) & + (ALL (f::'A => 'B) (x::'A) s::'A => bool. + monoidal u_4372 & finite (support u_4372 f s) --> + iterate u_4372 (insert x s) f = + (if x : s then iterate u_4372 s f + else u_4372 (f x) (iterate u_4372 s f)))" + by (import hollight ITERATE_CLAUSES_GEN) + +lemma ITERATE_CLAUSES: "monoidal (x::'q_60857 => 'q_60857 => 'q_60857) +==> (ALL f::'q_60815 => 'q_60857. iterate x {} f = neutral x) & + (ALL (f::'q_60859 => 'q_60857) (xa::'q_60859) s::'q_60859 => bool. + finite s --> + iterate x (insert xa s) f = + (if xa : s then iterate x s f else x (f xa) (iterate x s f)))" + by (import hollight ITERATE_CLAUSES) + +lemma ITERATE_UNION: "[| monoidal (u_4372::'q_60945 => 'q_60945 => 'q_60945); + finite (s::'q_60930 => bool) & + finite (x::'q_60930 => bool) & s Int x = {} |] +==> iterate u_4372 (s Un x) (f::'q_60930 => 'q_60945) = + u_4372 (iterate u_4372 s f) (iterate u_4372 x f)" + by (import hollight ITERATE_UNION) + +lemma ITERATE_UNION_GEN: "[| monoidal (x::'B => 'B => 'B); + finite (support x (xa::'A => 'B) (xb::'A => bool)) & + finite (support x xa (xc::'A => bool)) & + support x xa xb Int support x xa xc = {} |] +==> iterate x (xb Un xc) xa = x (iterate x xb xa) (iterate x xc xa)" + by (import hollight ITERATE_UNION_GEN) + +lemma ITERATE_DIFF: "[| monoidal (u::'q_61087 => 'q_61087 => 'q_61087); + finite (s::'q_61083 => bool) & (t::'q_61083 => bool) <= s |] +==> u (iterate u (s - t) (f::'q_61083 => 'q_61087)) (iterate u t f) = + iterate u s f" + by (import hollight ITERATE_DIFF) + +lemma ITERATE_DIFF_GEN: "[| monoidal (x::'B => 'B => 'B); + finite (support x (xa::'A => 'B) (xb::'A => bool)) & + support x xa (xc::'A => bool) <= support x xa xb |] +==> x (iterate x (xb - xc) xa) (iterate x xc xa) = iterate x xb xa" + by (import hollight ITERATE_DIFF_GEN) + +lemma ITERATE_INCL_EXCL: "[| monoidal (u_4372::'q_61316 => 'q_61316 => 'q_61316); + finite (s::'q_61298 => bool) & finite (t::'q_61298 => bool) |] +==> u_4372 (iterate u_4372 s (f::'q_61298 => 'q_61316)) + (iterate u_4372 t f) = + u_4372 (iterate u_4372 (s Un t) f) (iterate u_4372 (s Int t) f)" + by (import hollight ITERATE_INCL_EXCL) + +lemma ITERATE_CLOSED: "[| monoidal (u_4372::'B => 'B => 'B); + (P::'B => bool) (neutral u_4372) & + (ALL (x::'B) y::'B. P x & P y --> P (u_4372 x y)); + !!x::'A. + x : (s::'A => bool) & (f::'A => 'B) x ~= neutral u_4372 ==> P (f x) |] +==> P (iterate u_4372 s f)" + by (import hollight ITERATE_CLOSED) + +lemma ITERATE_RELATED: "[| monoidal (u_4372::'B => 'B => 'B); + (R::'B => 'B => bool) (neutral u_4372) (neutral u_4372) & + (ALL (x1::'B) (y1::'B) (x2::'B) y2::'B. + R x1 x2 & R y1 y2 --> R (u_4372 x1 y1) (u_4372 x2 y2)); + finite (x::'A => bool) & + (ALL xa::'A. xa : x --> R ((f::'A => 'B) xa) ((g::'A => 'B) xa)) |] +==> R (iterate u_4372 x f) (iterate u_4372 x g)" + by (import hollight ITERATE_RELATED) + +lemma ITERATE_EQ_NEUTRAL: "[| monoidal (u_4372::'B => 'B => 'B); + !!x::'A. x : (s::'A => bool) ==> (f::'A => 'B) x = neutral u_4372 |] +==> iterate u_4372 s f = neutral u_4372" + by (import hollight ITERATE_EQ_NEUTRAL) + +lemma ITERATE_SING: "monoidal (x::'B => 'B => 'B) ==> iterate x {xa::'A} (f::'A => 'B) = f xa" + by (import hollight ITERATE_SING) + +lemma ITERATE_DELETE: "[| monoidal (u::'B => 'B => 'B); finite (s::'A => bool) & (a::'A) : s |] +==> u ((f::'A => 'B) a) (iterate u (s - {a}) f) = iterate u s f" + by (import hollight ITERATE_DELETE) + +lemma ITERATE_DELTA: "monoidal (u_4372::'q_61672 => 'q_61672 => 'q_61672) +==> iterate u_4372 (xb::'q_61691 => bool) + (%xb::'q_61691. + if xb = (xa::'q_61691) then (x::'q_61691 => 'q_61672) xb + else neutral u_4372) = + (if xa : xb then x xa else neutral u_4372)" + by (import hollight ITERATE_DELTA) + +lemma ITERATE_IMAGE: "[| monoidal (u_4372::'C => 'C => 'C); + !!(x::'A) y::'A. + x : (s::'A => bool) & y : s & (f::'A => 'B) x = f y ==> x = y |] +==> iterate u_4372 (f ` s) (g::'B => 'C) = iterate u_4372 s (g o f)" + by (import hollight ITERATE_IMAGE) + +lemma ITERATE_BIJECTION: "[| monoidal (u_4372::'B => 'B => 'B); + (ALL x::'A. x : (s::'A => bool) --> (p::'A => 'A) x : s) & + (ALL y::'A. y : s --> (EX! x::'A. x : s & p x = y)) |] +==> iterate u_4372 s (f::'A => 'B) = iterate u_4372 s (f o p)" + by (import hollight ITERATE_BIJECTION) + +lemma ITERATE_ITERATE_PRODUCT: "[| monoidal (u_4372::'C => 'C => 'C); + finite (x::'A => bool) & + (ALL i::'A. i : x --> finite ((xa::'A => 'B => bool) i)) |] +==> iterate u_4372 x + (%i::'A. iterate u_4372 (xa i) ((xb::'A => 'B => 'C) i)) = + iterate u_4372 + {u::'A * 'B. EX (i::'A) j::'B. (i : x & j : xa i) & u = (i, j)} + (SOME f::'A * 'B => 'C. ALL (i::'A) j::'B. f (i, j) = xb i j)" + by (import hollight ITERATE_ITERATE_PRODUCT) + +lemma ITERATE_EQ: "[| monoidal (u_4372::'B => 'B => 'B); + !!x::'A. x : (s::'A => bool) ==> (f::'A => 'B) x = (g::'A => 'B) x |] +==> iterate u_4372 s f = iterate u_4372 s g" + by (import hollight ITERATE_EQ) + +lemma ITERATE_EQ_GENERAL: "[| monoidal (u_4372::'C => 'C => 'C); + (ALL y::'B. + y : (t::'B => bool) --> + (EX! x::'A. x : (s::'A => bool) & (h::'A => 'B) x = y)) & + (ALL x::'A. x : s --> h x : t & (g::'B => 'C) (h x) = (f::'A => 'C) x) |] +==> iterate u_4372 s f = iterate u_4372 t g" + by (import hollight ITERATE_EQ_GENERAL) + +lemma ITERATE_EQ_GENERAL_INVERSES: "[| monoidal (u_4372::'C => 'C => 'C); + (ALL y::'B. + y : (t::'B => bool) --> + (k::'B => 'A) y : (s::'A => bool) & (h::'A => 'B) (k y) = y) & + (ALL x::'A. + x : s --> + h x : t & k (h x) = x & (g::'B => 'C) (h x) = (f::'A => 'C) x) |] +==> iterate u_4372 s f = iterate u_4372 t g" + by (import hollight ITERATE_EQ_GENERAL_INVERSES) + +lemma ITERATE_INJECTION: "[| monoidal (u_4372::'B => 'B => 'B); + finite (s::'A => bool) & + (ALL x::'A. x : s --> (p::'A => 'A) x : s) & + (ALL (x::'A) y::'A. x : s & y : s & p x = p y --> x = y) |] +==> iterate u_4372 s ((f::'A => 'B) o p) = iterate u_4372 s f" + by (import hollight ITERATE_INJECTION) + +lemma ITERATE_UNION_NONZERO: "[| monoidal (u_4372::'B => 'B => 'B); + finite (s::'A => bool) & + finite (t::'A => bool) & + (ALL x::'A. x : s Int t --> (f::'A => 'B) x = neutral u_4372) |] +==> iterate u_4372 (s Un t) f = + u_4372 (iterate u_4372 s f) (iterate u_4372 t f)" + by (import hollight ITERATE_UNION_NONZERO) + +lemma ITERATE_OP: "[| monoidal (u_4372::'q_62649 => 'q_62649 => 'q_62649); + finite (s::'q_62648 => bool) |] +==> iterate u_4372 s + (%x::'q_62648. + u_4372 ((f::'q_62648 => 'q_62649) x) + ((g::'q_62648 => 'q_62649) x)) = + u_4372 (iterate u_4372 s f) (iterate u_4372 s g)" + by (import hollight ITERATE_OP) + +lemma ITERATE_SUPERSET: "[| monoidal (u_4372::'B => 'B => 'B); + (u::'A => bool) <= (v::'A => bool) & + (ALL x::'A. x : v & x ~: u --> (f::'A => 'B) x = neutral u_4372) |] +==> iterate u_4372 v f = iterate u_4372 u f" + by (import hollight ITERATE_SUPERSET) + +lemma ITERATE_IMAGE_NONZERO: "[| monoidal (u_4372::'C => 'C => 'C); + finite (x::'A => bool) & + (ALL (xa::'A) y::'A. + xa : x & y : x & xa ~= y & (f::'A => 'B) xa = f y --> + (g::'B => 'C) (f xa) = neutral u_4372) |] +==> iterate u_4372 (f ` x) g = iterate u_4372 x (g o f)" + by (import hollight ITERATE_IMAGE_NONZERO) + +lemma ITERATE_CASES: "[| monoidal (u_4372::'B => 'B => 'B); finite (s::'A => bool) |] +==> iterate u_4372 s + (%x::'A. + if (P::'A => bool) x then (f::'A => 'B) x else (g::'A => 'B) x) = + u_4372 (iterate u_4372 {u::'A. EX x::'A. (x : s & P x) & u = x} f) + (iterate u_4372 {u::'A. EX x::'A. (x : s & ~ P x) & u = x} g)" + by (import hollight ITERATE_CASES) + +lemma ITERATE_OP_GEN: "[| monoidal (u_4372::'B => 'B => 'B); + finite (support u_4372 (f::'A => 'B) (s::'A => bool)) & + finite (support u_4372 (g::'A => 'B) s) |] +==> iterate u_4372 s (%x::'A. u_4372 (f x) (g x)) = + u_4372 (iterate u_4372 s f) (iterate u_4372 s g)" + by (import hollight ITERATE_OP_GEN) + +lemma ITERATE_CLAUSES_NUMSEG: "monoidal (x::'q_63246 => 'q_63246 => 'q_63246) +==> (ALL xa::nat. + iterate x {xa..0::nat} (f::nat => 'q_63246) = + (if xa = (0::nat) then f (0::nat) else neutral x)) & + (ALL (xa::nat) xb::nat. + iterate x {xa..Suc xb} f = + (if xa <= Suc xb then x (iterate x {xa..xb} f) (f (Suc xb)) + else iterate x {xa..xb} f))" + by (import hollight ITERATE_CLAUSES_NUMSEG) + +lemma ITERATE_PAIR: "monoidal (u_4372::'q_63421 => 'q_63421 => 'q_63421) +==> iterate u_4372 {(2::nat) * (m::nat)..(2::nat) * (n::nat) + (1::nat)} + (f::nat => 'q_63421) = + iterate u_4372 {m..n} + (%i::nat. u_4372 (f ((2::nat) * i)) (f ((2::nat) * i + (1::nat))))" + by (import hollight ITERATE_PAIR) + +definition + nsum :: "('q_63439 => bool) => ('q_63439 => nat) => nat" where + "(op ==::(('q_63439::type => bool) => ('q_63439::type => nat) => nat) + => (('q_63439::type => bool) => ('q_63439::type => nat) => nat) + => prop) + (nsum::('q_63439::type => bool) => ('q_63439::type => nat) => nat) + ((iterate::(nat => nat => nat) + => ('q_63439::type => bool) => ('q_63439::type => nat) => nat) + (op +::nat => nat => nat))" + +lemma DEF_nsum: "(op =::(('q_63439::type => bool) => ('q_63439::type => nat) => nat) + => (('q_63439::type => bool) => ('q_63439::type => nat) => nat) + => bool) + (nsum::('q_63439::type => bool) => ('q_63439::type => nat) => nat) + ((iterate::(nat => nat => nat) + => ('q_63439::type => bool) => ('q_63439::type => nat) => nat) + (op +::nat => nat => nat))" + by (import hollight DEF_nsum) + +lemma NEUTRAL_ADD: "neutral op + = (0::nat)" + by (import hollight NEUTRAL_ADD) + +lemma NEUTRAL_MUL: "neutral op * = (1::nat)" + by (import hollight NEUTRAL_MUL) + +lemma MONOIDAL_ADD: "(monoidal::(nat => nat => nat) => bool) (op +::nat => nat => nat)" + by (import hollight MONOIDAL_ADD) + +lemma MONOIDAL_MUL: "(monoidal::(nat => nat => nat) => bool) (op *::nat => nat => nat)" + by (import hollight MONOIDAL_MUL) + +lemma NSUM_CLAUSES: "(ALL x::'q_63477 => nat. nsum {} x = (0::nat)) & +(ALL (x::'q_63516) (xa::'q_63516 => nat) xb::'q_63516 => bool. + finite xb --> + nsum (insert x xb) xa = + (if x : xb then nsum xb xa else xa x + nsum xb xa))" + by (import hollight NSUM_CLAUSES) + +lemma NSUM_UNION: "finite (xa::'q_63542 => bool) & +finite (xb::'q_63542 => bool) & xa Int xb = {} +==> nsum (xa Un xb) (x::'q_63542 => nat) = nsum xa x + nsum xb x" + by (import hollight NSUM_UNION) + +lemma NSUM_DIFF: "finite (s::'q_63597 => bool) & (t::'q_63597 => bool) <= s +==> nsum (s - t) (f::'q_63597 => nat) = nsum s f - nsum t f" + by (import hollight NSUM_DIFF) + +lemma NSUM_INCL_EXCL: "finite (x::'A => bool) & finite (xa::'A => bool) +==> nsum x (xb::'A => nat) + nsum xa xb = + nsum (x Un xa) xb + nsum (x Int xa) xb" + by (import hollight NSUM_INCL_EXCL) + +lemma NSUM_SUPPORT: "nsum (support op + (x::'q_63686 => nat) (xa::'q_63686 => bool)) x = +nsum xa x" + by (import hollight NSUM_SUPPORT) + +lemma NSUM_ADD: "finite (xb::'q_63720 => bool) +==> nsum xb + (%xb::'q_63720. (x::'q_63720 => nat) xb + (xa::'q_63720 => nat) xb) = + nsum xb x + nsum xb xa" + by (import hollight NSUM_ADD) + +lemma NSUM_ADD_GEN: "finite + {xa::'q_63807. + EX xc::'q_63807. + (xc : (xb::'q_63807 => bool) & (x::'q_63807 => nat) xc ~= (0::nat)) & + xa = xc} & +finite + {x::'q_63807. + EX xc::'q_63807. + (xc : xb & (xa::'q_63807 => nat) xc ~= (0::nat)) & x = xc} +==> nsum xb (%xb::'q_63807. x xb + xa xb) = nsum xb x + nsum xb xa" + by (import hollight NSUM_ADD_GEN) + +lemma NSUM_EQ_0: "(!!xb::'A. xb : (xa::'A => bool) ==> (x::'A => nat) xb = (0::nat)) +==> nsum xa x = (0::nat)" + by (import hollight NSUM_EQ_0) + +lemma NSUM_0: "nsum (x::'A => bool) (%n::'A. 0::nat) = (0::nat)" + by (import hollight NSUM_0) + +lemma NSUM_LMUL: "nsum (s::'A => bool) (%x::'A. (c::nat) * (f::'A => nat) x) = c * nsum s f" + by (import hollight NSUM_LMUL) + +lemma NSUM_RMUL: "nsum (xb::'A => bool) (%xb::'A. (x::'A => nat) xb * (xa::nat)) = +nsum xb x * xa" + by (import hollight NSUM_RMUL) + +lemma NSUM_LE: "finite (xb::'q_63997 => bool) & +(ALL xc::'q_63997. + xc : xb --> (x::'q_63997 => nat) xc <= (xa::'q_63997 => nat) xc) +==> nsum xb x <= nsum xb xa" + by (import hollight NSUM_LE) + +lemma NSUM_LT: "finite (s::'A => bool) & +(ALL x::'A. x : s --> (f::'A => nat) x <= (g::'A => nat) x) & +(EX x::'A. x : s & f x < g x) +==> nsum s f < nsum s g" + by (import hollight NSUM_LT) + +lemma NSUM_LT_ALL: "finite (s::'q_64119 => bool) & +s ~= {} & +(ALL x::'q_64119. x : s --> (f::'q_64119 => nat) x < (g::'q_64119 => nat) x) +==> nsum s f < nsum s g" + by (import hollight NSUM_LT_ALL) + +lemma NSUM_EQ: "(!!xc::'q_64157. + xc : (xb::'q_64157 => bool) + ==> (x::'q_64157 => nat) xc = (xa::'q_64157 => nat) xc) +==> nsum xb x = nsum xb xa" + by (import hollight NSUM_EQ) + +lemma NSUM_CONST: "finite (s::'q_64187 => bool) ==> nsum s (%n::'q_64187. c::nat) = CARD s * c" + by (import hollight NSUM_CONST) + +lemma NSUM_POS_BOUND: "[| finite (x::'A => bool) & nsum x (f::'A => nat) <= (b::nat); + (xa::'A) : x |] +==> f xa <= b" + by (import hollight NSUM_POS_BOUND) + +lemma NSUM_EQ_0_IFF: "finite (s::'q_64296 => bool) +==> (nsum s (f::'q_64296 => nat) = (0::nat)) = + (ALL x::'q_64296. x : s --> f x = (0::nat))" + by (import hollight NSUM_EQ_0_IFF) + +lemma NSUM_DELETE: "finite (xa::'q_64325 => bool) & (xb::'q_64325) : xa +==> (x::'q_64325 => nat) xb + nsum (xa - {xb}) x = nsum xa x" + by (import hollight NSUM_DELETE) + +lemma NSUM_SING: "nsum {xa::'q_64354} (x::'q_64354 => nat) = x xa" + by (import hollight NSUM_SING) + +lemma NSUM_DELTA: "nsum (x::'A => bool) (%x::'A. if x = (xa::'A) then b::nat else (0::nat)) = +(if xa : x then b else (0::nat))" + by (import hollight NSUM_DELTA) + +lemma NSUM_SWAP: "finite (x::'A => bool) & finite (xa::'B => bool) +==> nsum x (%i::'A. nsum xa ((f::'A => 'B => nat) i)) = + nsum xa (%j::'B. nsum x (%i::'A. f i j))" + by (import hollight NSUM_SWAP) + +lemma NSUM_IMAGE: "(!!(xa::'q_64490) y::'q_64490. + xa : (xb::'q_64490 => bool) & + y : xb & (x::'q_64490 => 'q_64466) xa = x y + ==> xa = y) +==> nsum (x ` xb) (xa::'q_64466 => nat) = nsum xb (xa o x)" + by (import hollight NSUM_IMAGE) + +lemma NSUM_SUPERSET: "(xa::'A => bool) <= (xb::'A => bool) & +(ALL xc::'A. xc : xb & xc ~: xa --> (x::'A => nat) xc = (0::nat)) +==> nsum xb x = nsum xa x" + by (import hollight NSUM_SUPERSET) + +lemma NSUM_UNION_RZERO: "finite (u::'A => bool) & +(ALL x::'A. x : (v::'A => bool) & x ~: u --> (f::'A => nat) x = (0::nat)) +==> nsum (u Un v) f = nsum u f" + by (import hollight NSUM_UNION_RZERO) + +lemma NSUM_UNION_LZERO: "finite (v::'A => bool) & +(ALL x::'A. x : (u::'A => bool) & x ~: v --> (f::'A => nat) x = (0::nat)) +==> nsum (u Un v) f = nsum v f" + by (import hollight NSUM_UNION_LZERO) + +lemma NSUM_RESTRICT: "finite (s::'q_64681 => bool) +==> nsum s + (%x::'q_64681. if x : s then (f::'q_64681 => nat) x else (0::nat)) = + nsum s f" + by (import hollight NSUM_RESTRICT) + +lemma NSUM_BOUND: "finite (x::'A => bool) & +(ALL xc::'A. xc : x --> (xa::'A => nat) xc <= (xb::nat)) +==> nsum x xa <= CARD x * xb" + by (import hollight NSUM_BOUND) + +lemma NSUM_BOUND_GEN: "finite (x::'A => bool) & +x ~= {} & (ALL xa::'A. xa : x --> (f::'A => nat) xa <= (b::nat) div CARD x) +==> nsum x f <= b" + by (import hollight NSUM_BOUND_GEN) + +lemma NSUM_BOUND_LT: "finite (s::'A => bool) & +(ALL x::'A. x : s --> (f::'A => nat) x <= (b::nat)) & +(EX x::'A. x : s & f x < b) +==> nsum s f < CARD s * b" + by (import hollight NSUM_BOUND_LT) + +lemma NSUM_BOUND_LT_ALL: "finite (s::'q_64899 => bool) & +s ~= {} & (ALL x::'q_64899. x : s --> (f::'q_64899 => nat) x < (b::nat)) +==> nsum s f < CARD s * b" + by (import hollight NSUM_BOUND_LT_ALL) + +lemma NSUM_BOUND_LT_GEN: "finite (s::'A => bool) & +s ~= {} & (ALL x::'A. x : s --> (f::'A => nat) x < (b::nat) div CARD s) +==> nsum s f < b" + by (import hollight NSUM_BOUND_LT_GEN) + +lemma NSUM_UNION_EQ: "finite (u::'q_65000 => bool) & +(s::'q_65000 => bool) Int (t::'q_65000 => bool) = {} & s Un t = u +==> nsum s (f::'q_65000 => nat) + nsum t f = nsum u f" + by (import hollight NSUM_UNION_EQ) + +lemma NSUM_EQ_SUPERSET: "finite (t::'A => bool) & +t <= (s::'A => bool) & +(ALL x::'A. x : t --> (f::'A => nat) x = (g::'A => nat) x) & +(ALL x::'A. x : s & x ~: t --> f x = (0::nat)) +==> nsum s f = nsum t g" + by (import hollight NSUM_EQ_SUPERSET) + +lemma NSUM_RESTRICT_SET: "nsum + {u::'A. EX xb::'A. (xb : (xa::'A => bool) & (x::'A => bool) xb) & u = xb} + (xb::'A => nat) = +nsum xa (%xa::'A. if x xa then xb xa else (0::nat))" + by (import hollight NSUM_RESTRICT_SET) + +lemma NSUM_NSUM_RESTRICT: "finite (s::'q_65257 => bool) & finite (t::'q_65256 => bool) +==> nsum s + (%x::'q_65257. + nsum + {u::'q_65256. + EX y::'q_65256. + (y : t & (R::'q_65257 => 'q_65256 => bool) x y) & u = y} + ((f::'q_65257 => 'q_65256 => nat) x)) = + nsum t + (%y::'q_65256. + nsum {u::'q_65257. EX x::'q_65257. (x : s & R x y) & u = x} + (%x::'q_65257. f x y))" + by (import hollight NSUM_NSUM_RESTRICT) + +lemma CARD_EQ_NSUM: "finite (x::'q_65276 => bool) ==> CARD x = nsum x (%x::'q_65276. 1::nat)" + by (import hollight CARD_EQ_NSUM) + +lemma NSUM_MULTICOUNT_GEN: "finite (s::'A => bool) & +finite (t::'B => bool) & +(ALL j::'B. + j : t --> + CARD {u::'A. EX i::'A. (i : s & (R::'A => 'B => bool) i j) & u = i} = + (k::'B => nat) j) +==> nsum s (%i::'A. CARD {u::'B. EX j::'B. (j : t & R i j) & u = j}) = + nsum t k" + by (import hollight NSUM_MULTICOUNT_GEN) + +lemma NSUM_MULTICOUNT: "finite (s::'A => bool) & +finite (t::'B => bool) & +(ALL j::'B. + j : t --> + CARD {u::'A. EX i::'A. (i : s & (R::'A => 'B => bool) i j) & u = i} = + (k::nat)) +==> nsum s (%i::'A. CARD {u::'B. EX j::'B. (j : t & R i j) & u = j}) = + k * CARD t" + by (import hollight NSUM_MULTICOUNT) + +lemma NSUM_IMAGE_GEN: "finite (s::'A => bool) +==> nsum s (g::'A => nat) = + nsum ((f::'A => 'B) ` s) + (%y::'B. nsum {u::'A. EX x::'A. (x : s & f x = y) & u = x} g)" + by (import hollight NSUM_IMAGE_GEN) + +lemma NSUM_GROUP: "finite (s::'A => bool) & (f::'A => 'B) ` s <= (t::'B => bool) +==> nsum t + (%y::'B. + nsum {u::'A. EX x::'A. (x : s & f x = y) & u = x} (g::'A => nat)) = + nsum s g" + by (import hollight NSUM_GROUP) + +lemma NSUM_SUBSET: "finite (u::'A => bool) & +finite (v::'A => bool) & +(ALL x::'A. x : u - v --> (f::'A => nat) x = (0::nat)) +==> nsum u f <= nsum v f" + by (import hollight NSUM_SUBSET) + +lemma NSUM_SUBSET_SIMPLE: "finite (v::'q_65804 => bool) & (u::'q_65804 => bool) <= v +==> nsum u (f::'q_65804 => nat) <= nsum v f" + by (import hollight NSUM_SUBSET_SIMPLE) + +lemma NSUM_IMAGE_NONZERO: "finite (xb::'A => bool) & +(ALL (xc::'A) xd::'A. + xc : xb & xd : xb & xc ~= xd & (xa::'A => 'B) xc = xa xd --> + (x::'B => nat) (xa xc) = (0::nat)) +==> nsum (xa ` xb) x = nsum xb (x o xa)" + by (import hollight NSUM_IMAGE_NONZERO) + +lemma NSUM_BIJECTION: "(ALL x::'A. x : (xb::'A => bool) --> (xa::'A => 'A) x : xb) & +(ALL y::'A. y : xb --> (EX! x::'A. x : xb & xa x = y)) +==> nsum xb (x::'A => nat) = nsum xb (x o xa)" + by (import hollight NSUM_BIJECTION) + +lemma NSUM_NSUM_PRODUCT: "finite (x::'A => bool) & +(ALL i::'A. i : x --> finite ((xa::'A => 'B => bool) i)) +==> nsum x (%x::'A. nsum (xa x) ((xb::'A => 'B => nat) x)) = + nsum {u::'A * 'B. EX (i::'A) j::'B. (i : x & j : xa i) & u = (i, j)} + (SOME f::'A * 'B => nat. ALL (i::'A) j::'B. f (i, j) = xb i j)" + by (import hollight NSUM_NSUM_PRODUCT) + +lemma NSUM_EQ_GENERAL: "(ALL y::'B. + y : (xa::'B => bool) --> + (EX! xa::'A. xa : (x::'A => bool) & (xd::'A => 'B) xa = y)) & +(ALL xe::'A. + xe : x --> xd xe : xa & (xc::'B => nat) (xd xe) = (xb::'A => nat) xe) +==> nsum x xb = nsum xa xc" + by (import hollight NSUM_EQ_GENERAL) + +lemma NSUM_EQ_GENERAL_INVERSES: "(ALL y::'B. + y : (xa::'B => bool) --> + (xe::'B => 'A) y : (x::'A => bool) & (xd::'A => 'B) (xe y) = y) & +(ALL xf::'A. + xf : x --> + xd xf : xa & + xe (xd xf) = xf & (xc::'B => nat) (xd xf) = (xb::'A => nat) xf) +==> nsum x xb = nsum xa xc" + by (import hollight NSUM_EQ_GENERAL_INVERSES) + +lemma NSUM_INJECTION: "finite (xb::'q_66274 => bool) & +(ALL x::'q_66274. x : xb --> (xa::'q_66274 => 'q_66274) x : xb) & +(ALL (x::'q_66274) y::'q_66274. x : xb & y : xb & xa x = xa y --> x = y) +==> nsum xb ((x::'q_66274 => nat) o xa) = nsum xb x" + by (import hollight NSUM_INJECTION) + +lemma NSUM_UNION_NONZERO: "finite (xa::'q_66317 => bool) & +finite (xb::'q_66317 => bool) & +(ALL xc::'q_66317. xc : xa Int xb --> (x::'q_66317 => nat) xc = (0::nat)) +==> nsum (xa Un xb) x = nsum xa x + nsum xb x" + by (import hollight NSUM_UNION_NONZERO) + +lemma NSUM_UNIONS_NONZERO: "finite (x::('A => bool) => bool) & +(ALL t::'A => bool. t : x --> finite t) & +(ALL (t1::'A => bool) (t2::'A => bool) xa::'A. + t1 : x & t2 : x & t1 ~= t2 & xa : t1 & xa : t2 --> + (f::'A => nat) xa = (0::nat)) +==> nsum (Union x) f = nsum x (%t::'A => bool. nsum t f)" + by (import hollight NSUM_UNIONS_NONZERO) + +lemma NSUM_CASES: "finite (x::'A => bool) +==> nsum x + (%x::'A. + if (xa::'A => bool) x then (xb::'A => nat) x + else (xc::'A => nat) x) = + nsum {u::'A. EX xb::'A. (xb : x & xa xb) & u = xb} xb + + nsum {u::'A. EX xb::'A. (xb : x & ~ xa xb) & u = xb} xc" + by (import hollight NSUM_CASES) + +lemma NSUM_CLOSED: "(P::nat => bool) (0::nat) & +(ALL (x::nat) y::nat. P x & P y --> P (x + y)) & +(ALL a::'A. a : (s::'A => bool) --> P ((f::'A => nat) a)) +==> P (nsum s f)" + by (import hollight NSUM_CLOSED) + +lemma NSUM_ADD_NUMSEG: "nsum {xb::nat..xc::nat} (%i::nat. (x::nat => nat) i + (xa::nat => nat) i) = +nsum {xb..xc} x + nsum {xb..xc} xa" + by (import hollight NSUM_ADD_NUMSEG) + +lemma NSUM_LE_NUMSEG: "(!!i::nat. + (xb::nat) <= i & i <= (xc::nat) + ==> (x::nat => nat) i <= (xa::nat => nat) i) +==> nsum {xb..xc} x <= nsum {xb..xc} xa" + by (import hollight NSUM_LE_NUMSEG) + +lemma NSUM_EQ_NUMSEG: "(!!i::nat. + (m::nat) <= i & i <= (n::nat) ==> (f::nat => nat) i = (g::nat => nat) i) +==> nsum {m..n} f = nsum {m..n} g" + by (import hollight NSUM_EQ_NUMSEG) + +lemma NSUM_CONST_NUMSEG: "nsum {xa..xb} (%n. x) = (xb + 1 - xa) * x" + by (import hollight NSUM_CONST_NUMSEG) + +lemma NSUM_EQ_0_NUMSEG: "(!!i::nat. (m::nat) <= i & i <= (n::nat) ==> (x::nat => nat) i = (0::nat)) +==> nsum {m..n} x = (0::nat)" + by (import hollight NSUM_EQ_0_NUMSEG) + +lemma NSUM_EQ_0_IFF_NUMSEG: "(nsum {xa::nat..xb::nat} (x::nat => nat) = (0::nat)) = +(ALL i::nat. xa <= i & i <= xb --> x i = (0::nat))" + by (import hollight NSUM_EQ_0_IFF_NUMSEG) + +lemma NSUM_TRIV_NUMSEG: "(n::nat) < (m::nat) ==> nsum {m..n} (f::nat => nat) = (0::nat)" + by (import hollight NSUM_TRIV_NUMSEG) + +lemma NSUM_SING_NUMSEG: "nsum {xa::nat..xa} (x::nat => nat) = x xa" + by (import hollight NSUM_SING_NUMSEG) + +lemma NSUM_CLAUSES_NUMSEG: "(ALL m. nsum {m..0} f = (if m = 0 then f 0 else 0)) & +(ALL m n. + nsum {m..Suc n} f = + (if m <= Suc n then nsum {m..n} f + f (Suc n) else nsum {m..n} f))" + by (import hollight NSUM_CLAUSES_NUMSEG) + +lemma NSUM_SWAP_NUMSEG: "nsum {a::nat..b::nat} + (%i::nat. nsum {c::nat..d::nat} ((f::nat => nat => nat) i)) = +nsum {c..d} (%j::nat. nsum {a..b} (%i::nat. f i j))" + by (import hollight NSUM_SWAP_NUMSEG) + +lemma NSUM_ADD_SPLIT: "(xa::nat) <= (xb::nat) + (1::nat) +==> nsum {xa..xb + (xc::nat)} (x::nat => nat) = + nsum {xa..xb} x + nsum {xb + (1::nat)..xb + xc} x" + by (import hollight NSUM_ADD_SPLIT) + +lemma NSUM_OFFSET: "nsum {(xb::nat) + (x::nat)..(xc::nat) + x} (xa::nat => nat) = +nsum {xb..xc} (%i::nat. xa (i + x))" + by (import hollight NSUM_OFFSET) + +lemma NSUM_OFFSET_0: "(xa::nat) <= (xb::nat) +==> nsum {xa..xb} (x::nat => nat) = + nsum {0::nat..xb - xa} (%i::nat. x (i + xa))" + by (import hollight NSUM_OFFSET_0) + +lemma NSUM_CLAUSES_LEFT: "(xa::nat) <= (xb::nat) +==> nsum {xa..xb} (x::nat => nat) = x xa + nsum {xa + (1::nat)..xb} x" + by (import hollight NSUM_CLAUSES_LEFT) + +lemma NSUM_CLAUSES_RIGHT: "(0::nat) < (n::nat) & (m::nat) <= n +==> nsum {m..n} (f::nat => nat) = nsum {m..n - (1::nat)} f + f n" + by (import hollight NSUM_CLAUSES_RIGHT) + +lemma NSUM_PAIR: "nsum {(2::nat) * (m::nat)..(2::nat) * (n::nat) + (1::nat)} (f::nat => nat) = +nsum {m..n} (%i::nat. f ((2::nat) * i) + f ((2::nat) * i + (1::nat)))" + by (import hollight NSUM_PAIR) + +lemma CARD_UNIONS: "finite (x::('A => bool) => bool) & +(ALL t::'A => bool. t : x --> finite t) & +(ALL (t::'A => bool) u::'A => bool. t : x & u : x & t ~= u --> t Int u = {}) +==> CARD (Union x) = nsum x CARD" + by (import hollight CARD_UNIONS) + +consts + sum :: "('q_67488 => bool) => ('q_67488 => hollight.real) => hollight.real" + +defs + sum_def: "(op ==::(('q_67488::type => bool) + => ('q_67488::type => hollight.real) => hollight.real) + => (('q_67488::type => bool) + => ('q_67488::type => hollight.real) => hollight.real) + => prop) + (hollight.sum::('q_67488::type => bool) + => ('q_67488::type => hollight.real) => hollight.real) + ((iterate::(hollight.real => hollight.real => hollight.real) + => ('q_67488::type => bool) + => ('q_67488::type => hollight.real) => hollight.real) + (real_add::hollight.real => hollight.real => hollight.real))" + +lemma DEF_sum: "(op =::(('q_67488::type => bool) + => ('q_67488::type => hollight.real) => hollight.real) + => (('q_67488::type => bool) + => ('q_67488::type => hollight.real) => hollight.real) + => bool) + (hollight.sum::('q_67488::type => bool) + => ('q_67488::type => hollight.real) => hollight.real) + ((iterate::(hollight.real => hollight.real => hollight.real) + => ('q_67488::type => bool) + => ('q_67488::type => hollight.real) => hollight.real) + (real_add::hollight.real => hollight.real => hollight.real))" + by (import hollight DEF_sum) + +lemma NEUTRAL_REAL_ADD: "neutral real_add = real_of_num 0" + by (import hollight NEUTRAL_REAL_ADD) + +lemma NEUTRAL_REAL_MUL: "neutral real_mul = real_of_num 1" + by (import hollight NEUTRAL_REAL_MUL) + +lemma MONOIDAL_REAL_ADD: "monoidal real_add" + by (import hollight MONOIDAL_REAL_ADD) + +lemma MONOIDAL_REAL_MUL: "monoidal real_mul" + by (import hollight MONOIDAL_REAL_MUL) + +lemma SUM_CLAUSES: "(ALL x::'q_67530 => hollight.real. + hollight.sum {} x = real_of_num (0::nat)) & +(ALL (x::'q_67571) (xa::'q_67571 => hollight.real) xb::'q_67571 => bool. + finite xb --> + hollight.sum (insert x xb) xa = + (if x : xb then hollight.sum xb xa + else real_add (xa x) (hollight.sum xb xa)))" + by (import hollight SUM_CLAUSES) + +lemma SUM_UNION: "finite (xa::'q_67597 => bool) & +finite (xb::'q_67597 => bool) & xa Int xb = {} +==> hollight.sum (xa Un xb) (x::'q_67597 => hollight.real) = + real_add (hollight.sum xa x) (hollight.sum xb x)" + by (import hollight SUM_UNION) + +lemma SUM_DIFF: "finite (xa::'q_67637 => bool) & (xb::'q_67637 => bool) <= xa +==> hollight.sum (xa - xb) (x::'q_67637 => hollight.real) = + real_sub (hollight.sum xa x) (hollight.sum xb x)" + by (import hollight SUM_DIFF) + +lemma SUM_INCL_EXCL: "finite (x::'A => bool) & finite (xa::'A => bool) +==> real_add (hollight.sum x (xb::'A => hollight.real)) + (hollight.sum xa xb) = + real_add (hollight.sum (x Un xa) xb) (hollight.sum (x Int xa) xb)" + by (import hollight SUM_INCL_EXCL) + +lemma SUM_SUPPORT: "hollight.sum + (support real_add (x::'q_67726 => hollight.real) (xa::'q_67726 => bool)) + x = +hollight.sum xa x" + by (import hollight SUM_SUPPORT) + +lemma SUM_ADD: "finite (xb::'q_67760 => bool) +==> hollight.sum xb + (%xb::'q_67760. + real_add ((x::'q_67760 => hollight.real) xb) + ((xa::'q_67760 => hollight.real) xb)) = + real_add (hollight.sum xb x) (hollight.sum xb xa)" + by (import hollight SUM_ADD) + +lemma SUM_ADD_GEN: "finite + {xa::'q_67851. + EX xc::'q_67851. + (xc : (xb::'q_67851 => bool) & + (x::'q_67851 => hollight.real) xc ~= real_of_num (0::nat)) & + xa = xc} & +finite + {x::'q_67851. + EX xc::'q_67851. + (xc : xb & + (xa::'q_67851 => hollight.real) xc ~= real_of_num (0::nat)) & + x = xc} +==> hollight.sum xb (%xb::'q_67851. real_add (x xb) (xa xb)) = + real_add (hollight.sum xb x) (hollight.sum xb xa)" + by (import hollight SUM_ADD_GEN) + +lemma SUM_EQ_0: "(!!xb::'A. + xb : (xa::'A => bool) + ==> (x::'A => hollight.real) xb = real_of_num (0::nat)) +==> hollight.sum xa x = real_of_num (0::nat)" + by (import hollight SUM_EQ_0) + +lemma SUM_0: "hollight.sum (x::'A => bool) (%n::'A. real_of_num (0::nat)) = +real_of_num (0::nat)" + by (import hollight SUM_0) + +lemma SUM_LMUL: "hollight.sum (s::'A => bool) + (%x::'A. real_mul (c::hollight.real) ((f::'A => hollight.real) x)) = +real_mul c (hollight.sum s f)" + by (import hollight SUM_LMUL) + +lemma SUM_RMUL: "hollight.sum (xb::'A => bool) + (%xb::'A. real_mul ((x::'A => hollight.real) xb) (xa::hollight.real)) = +real_mul (hollight.sum xb x) xa" + by (import hollight SUM_RMUL) + +lemma SUM_NEG: "hollight.sum (xa::'q_68051 => bool) + (%xa::'q_68051. real_neg ((x::'q_68051 => hollight.real) xa)) = +real_neg (hollight.sum xa x)" + by (import hollight SUM_NEG) + +lemma SUM_SUB: "finite (xb::'q_68086 => bool) +==> hollight.sum xb + (%xb::'q_68086. + real_sub ((x::'q_68086 => hollight.real) xb) + ((xa::'q_68086 => hollight.real) xb)) = + real_sub (hollight.sum xb x) (hollight.sum xb xa)" + by (import hollight SUM_SUB) + +lemma SUM_LE: "finite (xb::'q_68128 => bool) & +(ALL xc::'q_68128. + xc : xb --> + real_le ((x::'q_68128 => hollight.real) xc) + ((xa::'q_68128 => hollight.real) xc)) +==> real_le (hollight.sum xb x) (hollight.sum xb xa)" + by (import hollight SUM_LE) + +lemma SUM_LT: "finite (s::'A => bool) & +(ALL x::'A. + x : s --> + real_le ((f::'A => hollight.real) x) ((g::'A => hollight.real) x)) & +(EX x::'A. x : s & real_lt (f x) (g x)) +==> real_lt (hollight.sum s f) (hollight.sum s g)" + by (import hollight SUM_LT) + +lemma SUM_LT_ALL: "finite (s::'q_68250 => bool) & +s ~= {} & +(ALL x::'q_68250. + x : s --> + real_lt ((f::'q_68250 => hollight.real) x) + ((g::'q_68250 => hollight.real) x)) +==> real_lt (hollight.sum s f) (hollight.sum s g)" + by (import hollight SUM_LT_ALL) + +lemma SUM_EQ: "(!!xc::'q_68288. + xc : (xb::'q_68288 => bool) + ==> (x::'q_68288 => hollight.real) xc = + (xa::'q_68288 => hollight.real) xc) +==> hollight.sum xb x = hollight.sum xb xa" + by (import hollight SUM_EQ) + +lemma SUM_ABS: "finite (s::'q_68347 => bool) +==> real_le (real_abs (hollight.sum s (f::'q_68347 => hollight.real))) + (hollight.sum s (%x::'q_68347. real_abs (f x)))" + by (import hollight SUM_ABS) + +lemma SUM_ABS_LE: "finite (s::'A => bool) & +(ALL x::'A. + x : s --> + real_le (real_abs ((f::'A => hollight.real) x)) + ((g::'A => hollight.real) x)) +==> real_le (real_abs (hollight.sum s f)) (hollight.sum s g)" + by (import hollight SUM_ABS_LE) + +lemma SUM_CONST: "finite (s::'q_68423 => bool) +==> hollight.sum s (%n::'q_68423. c::hollight.real) = + real_mul (real_of_num (CARD s)) c" + by (import hollight SUM_CONST) + +lemma SUM_POS_LE: "finite (xa::'q_68465 => bool) & +(ALL xb::'q_68465. + xb : xa --> + real_le (real_of_num (0::nat)) ((x::'q_68465 => hollight.real) xb)) +==> real_le (real_of_num (0::nat)) (hollight.sum xa x)" + by (import hollight SUM_POS_LE) + +lemma SUM_POS_BOUND: "[| finite (x::'A => bool) & + (ALL xa::'A. + xa : x --> + real_le (real_of_num (0::nat)) ((f::'A => hollight.real) xa)) & + real_le (hollight.sum x f) (b::hollight.real); + (xa::'A) : x |] +==> real_le (f xa) b" + by (import hollight SUM_POS_BOUND) + +lemma SUM_POS_EQ_0: "[| finite (xa::'q_68612 => bool) & + (ALL xb::'q_68612. + xb : xa --> + real_le (real_of_num (0::nat)) ((x::'q_68612 => hollight.real) xb)) & + hollight.sum xa x = real_of_num (0::nat); + (xb::'q_68612) : xa |] +==> x xb = real_of_num (0::nat)" + by (import hollight SUM_POS_EQ_0) + +lemma SUM_ZERO_EXISTS: "finite (s::'A => bool) & +hollight.sum s (u::'A => hollight.real) = real_of_num (0::nat) +==> (ALL i::'A. i : s --> u i = real_of_num (0::nat)) | + (EX (j::'A) k::'A. + j : s & + real_lt (u j) (real_of_num (0::nat)) & + k : s & real_gt (u k) (real_of_num (0::nat)))" + by (import hollight SUM_ZERO_EXISTS) + +lemma SUM_DELETE: "finite (xa::'q_68854 => bool) & (xb::'q_68854) : xa +==> hollight.sum (xa - {xb}) (x::'q_68854 => hollight.real) = + real_sub (hollight.sum xa x) (x xb)" + by (import hollight SUM_DELETE) + +lemma SUM_DELETE_CASES: "finite (s::'q_68907 => bool) +==> hollight.sum (s - {a::'q_68907}) (f::'q_68907 => hollight.real) = + (if a : s then real_sub (hollight.sum s f) (f a) else hollight.sum s f)" + by (import hollight SUM_DELETE_CASES) + +lemma SUM_SING: "hollight.sum {xa::'q_68930} (x::'q_68930 => hollight.real) = x xa" + by (import hollight SUM_SING) + +lemma SUM_DELTA: "hollight.sum (x::'A => bool) + (%x::'A. if x = (xa::'A) then b::hollight.real else real_of_num (0::nat)) = +(if xa : x then b else real_of_num (0::nat))" + by (import hollight SUM_DELTA) + +lemma SUM_SWAP: "finite (x::'A => bool) & finite (xa::'B => bool) +==> hollight.sum x + (%i::'A. hollight.sum xa ((f::'A => 'B => hollight.real) i)) = + hollight.sum xa (%j::'B. hollight.sum x (%i::'A. f i j))" + by (import hollight SUM_SWAP) + +lemma SUM_IMAGE: "(!!(xa::'q_69070) y::'q_69070. + xa : (xb::'q_69070 => bool) & + y : xb & (x::'q_69070 => 'q_69046) xa = x y + ==> xa = y) +==> hollight.sum (x ` xb) (xa::'q_69046 => hollight.real) = + hollight.sum xb (xa o x)" + by (import hollight SUM_IMAGE) + +lemma SUM_SUPERSET: "(xa::'A => bool) <= (xb::'A => bool) & +(ALL xc::'A. + xc : xb & xc ~: xa --> + (x::'A => hollight.real) xc = real_of_num (0::nat)) +==> hollight.sum xb x = hollight.sum xa x" + by (import hollight SUM_SUPERSET) + +lemma SUM_UNION_RZERO: "finite (u::'A => bool) & +(ALL x::'A. + x : (v::'A => bool) & x ~: u --> + (f::'A => hollight.real) x = real_of_num (0::nat)) +==> hollight.sum (u Un v) f = hollight.sum u f" + by (import hollight SUM_UNION_RZERO) + +lemma SUM_UNION_LZERO: "finite (v::'A => bool) & +(ALL x::'A. + x : (u::'A => bool) & x ~: v --> + (f::'A => hollight.real) x = real_of_num (0::nat)) +==> hollight.sum (u Un v) f = hollight.sum v f" + by (import hollight SUM_UNION_LZERO) + +lemma SUM_RESTRICT: "finite (s::'q_69267 => bool) +==> hollight.sum s + (%x::'q_69267. + if x : s then (f::'q_69267 => hollight.real) x + else real_of_num (0::nat)) = + hollight.sum s f" + by (import hollight SUM_RESTRICT) + +lemma SUM_BOUND: "finite (x::'A => bool) & +(ALL xc::'A. + xc : x --> real_le ((xa::'A => hollight.real) xc) (xb::hollight.real)) +==> real_le (hollight.sum x xa) (real_mul (real_of_num (CARD x)) xb)" + by (import hollight SUM_BOUND) + +lemma SUM_BOUND_GEN: "finite (s::'A => bool) & +s ~= {} & +(ALL x::'A. + x : s --> + real_le ((f::'A => hollight.real) x) + (real_div (b::hollight.real) (real_of_num (CARD s)))) +==> real_le (hollight.sum s f) b" + by (import hollight SUM_BOUND_GEN) + +lemma SUM_ABS_BOUND: "finite (s::'A => bool) & +(ALL x::'A. + x : s --> + real_le (real_abs ((f::'A => hollight.real) x)) (b::hollight.real)) +==> real_le (real_abs (hollight.sum s f)) + (real_mul (real_of_num (CARD s)) b)" + by (import hollight SUM_ABS_BOUND) + +lemma SUM_BOUND_LT: "finite (s::'A => bool) & +(ALL x::'A. + x : s --> real_le ((f::'A => hollight.real) x) (b::hollight.real)) & +(EX x::'A. x : s & real_lt (f x) b) +==> real_lt (hollight.sum s f) (real_mul (real_of_num (CARD s)) b)" + by (import hollight SUM_BOUND_LT) + +lemma SUM_BOUND_LT_ALL: "finite (s::'q_69531 => bool) & +s ~= {} & +(ALL x::'q_69531. + x : s --> real_lt ((f::'q_69531 => hollight.real) x) (b::hollight.real)) +==> real_lt (hollight.sum s f) (real_mul (real_of_num (CARD s)) b)" + by (import hollight SUM_BOUND_LT_ALL) + +lemma SUM_BOUND_LT_GEN: "finite (s::'A => bool) & +s ~= {} & +(ALL x::'A. + x : s --> + real_lt ((f::'A => hollight.real) x) + (real_div (b::hollight.real) (real_of_num (CARD s)))) +==> real_lt (hollight.sum s f) b" + by (import hollight SUM_BOUND_LT_GEN) + +lemma SUM_UNION_EQ: "finite (u::'q_69614 => bool) & +(s::'q_69614 => bool) Int (t::'q_69614 => bool) = {} & s Un t = u +==> real_add (hollight.sum s (f::'q_69614 => hollight.real)) + (hollight.sum t f) = + hollight.sum u f" + by (import hollight SUM_UNION_EQ) + +lemma SUM_EQ_SUPERSET: "finite (t::'A => bool) & +t <= (s::'A => bool) & +(ALL x::'A. + x : t --> (f::'A => hollight.real) x = (g::'A => hollight.real) x) & +(ALL x::'A. x : s & x ~: t --> f x = real_of_num (0::nat)) +==> hollight.sum s f = hollight.sum t g" + by (import hollight SUM_EQ_SUPERSET) + +lemma SUM_RESTRICT_SET: "hollight.sum + {u::'q_69783. + EX xb::'q_69783. + (xb : (xa::'q_69783 => bool) & (x::'q_69783 => bool) xb) & u = xb} + (xb::'q_69783 => hollight.real) = +hollight.sum xa + (%xa::'q_69783. if x xa then xb xa else real_of_num (0::nat))" + by (import hollight SUM_RESTRICT_SET) + +lemma SUM_SUM_RESTRICT: "finite (s::'q_69875 => bool) & finite (t::'q_69874 => bool) +==> hollight.sum s + (%x::'q_69875. + hollight.sum + {u::'q_69874. + EX y::'q_69874. + (y : t & (R::'q_69875 => 'q_69874 => bool) x y) & u = y} + ((f::'q_69875 => 'q_69874 => hollight.real) x)) = + hollight.sum t + (%y::'q_69874. + hollight.sum {u::'q_69875. EX x::'q_69875. (x : s & R x y) & u = x} + (%x::'q_69875. f x y))" + by (import hollight SUM_SUM_RESTRICT) + +lemma CARD_EQ_SUM: "finite (x::'q_69896 => bool) +==> real_of_num (CARD x) = + hollight.sum x (%x::'q_69896. real_of_num (1::nat))" + by (import hollight CARD_EQ_SUM) + +lemma SUM_MULTICOUNT_GEN: "finite (s::'A => bool) & +finite (t::'B => bool) & +(ALL j::'B. + j : t --> + CARD {u::'A. EX i::'A. (i : s & (R::'A => 'B => bool) i j) & u = i} = + (k::'B => nat) j) +==> hollight.sum s + (%i::'A. + real_of_num (CARD {u::'B. EX j::'B. (j : t & R i j) & u = j})) = + hollight.sum t (%i::'B. real_of_num (k i))" + by (import hollight SUM_MULTICOUNT_GEN) + +lemma SUM_MULTICOUNT: "finite (s::'A => bool) & +finite (t::'B => bool) & +(ALL j::'B. + j : t --> + CARD {u::'A. EX i::'A. (i : s & (R::'A => 'B => bool) i j) & u = i} = + (k::nat)) +==> hollight.sum s + (%i::'A. + real_of_num (CARD {u::'B. EX j::'B. (j : t & R i j) & u = j})) = + real_of_num (k * CARD t)" + by (import hollight SUM_MULTICOUNT) + +lemma SUM_IMAGE_GEN: "finite (s::'A => bool) +==> hollight.sum s (g::'A => hollight.real) = + hollight.sum ((f::'A => 'B) ` s) + (%y::'B. hollight.sum {u::'A. EX x::'A. (x : s & f x = y) & u = x} g)" + by (import hollight SUM_IMAGE_GEN) + +lemma SUM_GROUP: "finite (s::'A => bool) & (f::'A => 'B) ` s <= (t::'B => bool) +==> hollight.sum t + (%y::'B. + hollight.sum {u::'A. EX x::'A. (x : s & f x = y) & u = x} + (g::'A => hollight.real)) = + hollight.sum s g" + by (import hollight SUM_GROUP) + +lemma REAL_OF_NUM_SUM: "finite (s::'q_70361 => bool) +==> real_of_num (nsum s (f::'q_70361 => nat)) = + hollight.sum s (%x::'q_70361. real_of_num (f x))" + by (import hollight REAL_OF_NUM_SUM) + +lemma SUM_SUBSET: "finite (u::'A => bool) & +finite (v::'A => bool) & +(ALL x::'A. + x : u - v --> + real_le ((f::'A => hollight.real) x) (real_of_num (0::nat))) & +(ALL x::'A. x : v - u --> real_le (real_of_num (0::nat)) (f x)) +==> real_le (hollight.sum u f) (hollight.sum v f)" + by (import hollight SUM_SUBSET) + +lemma SUM_SUBSET_SIMPLE: "finite (v::'A => bool) & +(u::'A => bool) <= v & +(ALL x::'A. + x : v - u --> + real_le (real_of_num (0::nat)) ((f::'A => hollight.real) x)) +==> real_le (hollight.sum u f) (hollight.sum v f)" + by (import hollight SUM_SUBSET_SIMPLE) + +lemma SUM_IMAGE_NONZERO: "finite (xb::'A => bool) & +(ALL (xc::'A) xd::'A. + xc : xb & xd : xb & xc ~= xd & (xa::'A => 'B) xc = xa xd --> + (x::'B => hollight.real) (xa xc) = real_of_num (0::nat)) +==> hollight.sum (xa ` xb) x = hollight.sum xb (x o xa)" + by (import hollight SUM_IMAGE_NONZERO) + +lemma SUM_BIJECTION: "(ALL x::'A. x : (xb::'A => bool) --> (xa::'A => 'A) x : xb) & +(ALL y::'A. y : xb --> (EX! x::'A. x : xb & xa x = y)) +==> hollight.sum xb (x::'A => hollight.real) = hollight.sum xb (x o xa)" + by (import hollight SUM_BIJECTION) + +lemma SUM_SUM_PRODUCT: "finite (x::'A => bool) & +(ALL i::'A. i : x --> finite ((xa::'A => 'B => bool) i)) +==> hollight.sum x + (%x::'A. hollight.sum (xa x) ((xb::'A => 'B => hollight.real) x)) = + hollight.sum + {u::'A * 'B. EX (i::'A) j::'B. (i : x & j : xa i) & u = (i, j)} + (SOME f::'A * 'B => hollight.real. + ALL (i::'A) j::'B. f (i, j) = xb i j)" + by (import hollight SUM_SUM_PRODUCT) + +lemma SUM_EQ_GENERAL: "(ALL y::'B. + y : (xa::'B => bool) --> + (EX! xa::'A. xa : (x::'A => bool) & (xd::'A => 'B) xa = y)) & +(ALL xe::'A. + xe : x --> + xd xe : xa & + (xc::'B => hollight.real) (xd xe) = (xb::'A => hollight.real) xe) +==> hollight.sum x xb = hollight.sum xa xc" + by (import hollight SUM_EQ_GENERAL) + +lemma SUM_EQ_GENERAL_INVERSES: "(ALL y::'B. + y : (xa::'B => bool) --> + (xe::'B => 'A) y : (x::'A => bool) & (xd::'A => 'B) (xe y) = y) & +(ALL xf::'A. + xf : x --> + xd xf : xa & + xe (xd xf) = xf & + (xc::'B => hollight.real) (xd xf) = (xb::'A => hollight.real) xf) +==> hollight.sum x xb = hollight.sum xa xc" + by (import hollight SUM_EQ_GENERAL_INVERSES) + +lemma SUM_INJECTION: "finite (xb::'q_71007 => bool) & +(ALL x::'q_71007. x : xb --> (xa::'q_71007 => 'q_71007) x : xb) & +(ALL (x::'q_71007) y::'q_71007. x : xb & y : xb & xa x = xa y --> x = y) +==> hollight.sum xb ((x::'q_71007 => hollight.real) o xa) = + hollight.sum xb x" + by (import hollight SUM_INJECTION) + +lemma SUM_UNION_NONZERO: "finite (xa::'q_71050 => bool) & +finite (xb::'q_71050 => bool) & +(ALL xc::'q_71050. + xc : xa Int xb --> + (x::'q_71050 => hollight.real) xc = real_of_num (0::nat)) +==> hollight.sum (xa Un xb) x = + real_add (hollight.sum xa x) (hollight.sum xb x)" + by (import hollight SUM_UNION_NONZERO) + +lemma SUM_UNIONS_NONZERO: "finite (x::('A => bool) => bool) & +(ALL t::'A => bool. t : x --> finite t) & +(ALL (t1::'A => bool) (t2::'A => bool) xa::'A. + t1 : x & t2 : x & t1 ~= t2 & xa : t1 & xa : t2 --> + (f::'A => hollight.real) xa = real_of_num (0::nat)) +==> hollight.sum (Union x) f = + hollight.sum x (%t::'A => bool. hollight.sum t f)" + by (import hollight SUM_UNIONS_NONZERO) + +lemma SUM_CASES: "finite (x::'A => bool) +==> hollight.sum x + (%x::'A. + if (xa::'A => bool) x then (xb::'A => hollight.real) x + else (xc::'A => hollight.real) x) = + real_add (hollight.sum {u::'A. EX xb::'A. (xb : x & xa xb) & u = xb} xb) + (hollight.sum {u::'A. EX xb::'A. (xb : x & ~ xa xb) & u = xb} xc)" + by (import hollight SUM_CASES) + +lemma SUM_CASES_1: "finite (s::'q_71319 => bool) & (a::'q_71319) : s +==> hollight.sum s + (%x::'q_71319. + if x = a then y::hollight.real + else (f::'q_71319 => hollight.real) x) = + real_add (hollight.sum s f) (real_sub y (f a))" + by (import hollight SUM_CASES_1) + +lemma SUM_LE_INCLUDED: "finite (s::'A => bool) & +finite (t::'B => bool) & +(ALL y::'B. + y : t --> real_le (real_of_num (0::nat)) ((g::'B => hollight.real) y)) & +(ALL x::'A. + x : s --> + (EX y::'B. + y : t & + (i::'B => 'A) y = x & real_le ((f::'A => hollight.real) x) (g y))) +==> real_le (hollight.sum s f) (hollight.sum t g)" + by (import hollight SUM_LE_INCLUDED) + +lemma SUM_IMAGE_LE: "finite (s::'A => bool) & +(ALL x::'A. + x : s --> + real_le (real_of_num (0::nat)) + ((g::'B => hollight.real) ((f::'A => 'B) x))) +==> real_le (hollight.sum (f ` s) g) (hollight.sum s (g o f))" + by (import hollight SUM_IMAGE_LE) + +lemma SUM_CLOSED: "(P::hollight.real => bool) (real_of_num (0::nat)) & +(ALL (x::hollight.real) y::hollight.real. P x & P y --> P (real_add x y)) & +(ALL a::'A. a : (s::'A => bool) --> P ((f::'A => hollight.real) a)) +==> P (hollight.sum s f)" + by (import hollight SUM_CLOSED) + +lemma SUM_ADD_NUMSEG: "hollight.sum {xb::nat..xc::nat} + (%i::nat. + real_add ((x::nat => hollight.real) i) + ((xa::nat => hollight.real) i)) = +real_add (hollight.sum {xb..xc} x) (hollight.sum {xb..xc} xa)" + by (import hollight SUM_ADD_NUMSEG) + +lemma SUM_SUB_NUMSEG: "hollight.sum {xb::nat..xc::nat} + (%i::nat. + real_sub ((x::nat => hollight.real) i) + ((xa::nat => hollight.real) i)) = +real_sub (hollight.sum {xb..xc} x) (hollight.sum {xb..xc} xa)" + by (import hollight SUM_SUB_NUMSEG) + +lemma SUM_LE_NUMSEG: "(!!i::nat. + (xb::nat) <= i & i <= (xc::nat) + ==> real_le ((x::nat => hollight.real) i) + ((xa::nat => hollight.real) i)) +==> real_le (hollight.sum {xb..xc} x) (hollight.sum {xb..xc} xa)" + by (import hollight SUM_LE_NUMSEG) + +lemma SUM_EQ_NUMSEG: "(!!i::nat. + (m::nat) <= i & i <= (n::nat) + ==> (f::nat => hollight.real) i = (g::nat => hollight.real) i) +==> hollight.sum {m..n} f = hollight.sum {m..n} g" + by (import hollight SUM_EQ_NUMSEG) + +lemma SUM_ABS_NUMSEG: "real_le + (real_abs (hollight.sum {xa::nat..xb::nat} (x::nat => hollight.real))) + (hollight.sum {xa..xb} (%i::nat. real_abs (x i)))" + by (import hollight SUM_ABS_NUMSEG) + +lemma SUM_CONST_NUMSEG: "hollight.sum {xa..xb} (%n. x) = real_mul (real_of_num (xb + 1 - xa)) x" + by (import hollight SUM_CONST_NUMSEG) + +lemma SUM_EQ_0_NUMSEG: "(!!i::nat. + (m::nat) <= i & i <= (n::nat) + ==> (x::nat => hollight.real) i = real_of_num (0::nat)) +==> hollight.sum {m..n} x = real_of_num (0::nat)" + by (import hollight SUM_EQ_0_NUMSEG) + +lemma SUM_TRIV_NUMSEG: "(n::nat) < (m::nat) +==> hollight.sum {m..n} (f::nat => hollight.real) = real_of_num (0::nat)" + by (import hollight SUM_TRIV_NUMSEG) + +lemma SUM_POS_LE_NUMSEG: "(!!p::nat. + (x::nat) <= p & p <= (xa::nat) + ==> real_le (real_of_num (0::nat)) ((xb::nat => hollight.real) p)) +==> real_le (real_of_num (0::nat)) (hollight.sum {x..xa} xb)" + by (import hollight SUM_POS_LE_NUMSEG) + +lemma SUM_POS_EQ_0_NUMSEG: "[| (ALL p::nat. + (m::nat) <= p & p <= (n::nat) --> + real_le (real_of_num (0::nat)) ((f::nat => hollight.real) p)) & + hollight.sum {m..n} f = real_of_num (0::nat); + m <= (p::nat) & p <= n |] +==> f p = real_of_num (0::nat)" + by (import hollight SUM_POS_EQ_0_NUMSEG) + +lemma SUM_SING_NUMSEG: "hollight.sum {xa::nat..xa} (x::nat => hollight.real) = x xa" + by (import hollight SUM_SING_NUMSEG) + +lemma SUM_CLAUSES_NUMSEG: "(ALL m. hollight.sum {m..0} f = (if m = 0 then f 0 else real_of_num 0)) & +(ALL m n. + hollight.sum {m..Suc n} f = + (if m <= Suc n then real_add (hollight.sum {m..n} f) (f (Suc n)) + else hollight.sum {m..n} f))" + by (import hollight SUM_CLAUSES_NUMSEG) + +lemma SUM_SWAP_NUMSEG: "hollight.sum {a::nat..b::nat} + (%i::nat. + hollight.sum {c::nat..d::nat} ((f::nat => nat => hollight.real) i)) = +hollight.sum {c..d} (%j::nat. hollight.sum {a..b} (%i::nat. f i j))" + by (import hollight SUM_SWAP_NUMSEG) + +lemma SUM_ADD_SPLIT: "(xa::nat) <= (xb::nat) + (1::nat) +==> hollight.sum {xa..xb + (xc::nat)} (x::nat => hollight.real) = + real_add (hollight.sum {xa..xb} x) + (hollight.sum {xb + (1::nat)..xb + xc} x)" + by (import hollight SUM_ADD_SPLIT) + +lemma SUM_OFFSET: "hollight.sum {(xb::nat) + (x::nat)..(xc::nat) + x} + (xa::nat => hollight.real) = +hollight.sum {xb..xc} (%i::nat. xa (i + x))" + by (import hollight SUM_OFFSET) + +lemma SUM_OFFSET_0: "(xa::nat) <= (xb::nat) +==> hollight.sum {xa..xb} (x::nat => hollight.real) = + hollight.sum {0::nat..xb - xa} (%i::nat. x (i + xa))" + by (import hollight SUM_OFFSET_0) + +lemma SUM_CLAUSES_LEFT: "(xa::nat) <= (xb::nat) +==> hollight.sum {xa..xb} (x::nat => hollight.real) = + real_add (x xa) (hollight.sum {xa + (1::nat)..xb} x)" + by (import hollight SUM_CLAUSES_LEFT) + +lemma SUM_CLAUSES_RIGHT: "(0::nat) < (n::nat) & (m::nat) <= n +==> hollight.sum {m..n} (f::nat => hollight.real) = + real_add (hollight.sum {m..n - (1::nat)} f) (f n)" + by (import hollight SUM_CLAUSES_RIGHT) + +lemma SUM_PAIR: "hollight.sum {(2::nat) * (m::nat)..(2::nat) * (n::nat) + (1::nat)} + (f::nat => hollight.real) = +hollight.sum {m..n} + (%i::nat. real_add (f ((2::nat) * i)) (f ((2::nat) * i + (1::nat))))" + by (import hollight SUM_PAIR) + +lemma REAL_OF_NUM_SUM_NUMSEG: "real_of_num (nsum {xa::nat..xb::nat} (x::nat => nat)) = +hollight.sum {xa..xb} (%i::nat. real_of_num (x i))" + by (import hollight REAL_OF_NUM_SUM_NUMSEG) + +lemma SUM_PARTIAL_SUC: "hollight.sum {m::nat..n::nat} + (%k::nat. + real_mul ((f::nat => hollight.real) k) + (real_sub ((g::nat => hollight.real) (k + (1::nat))) (g k))) = +(if m <= n + then real_sub + (real_sub (real_mul (f (n + (1::nat))) (g (n + (1::nat)))) + (real_mul (f m) (g m))) + (hollight.sum {m..n} + (%k::nat. + real_mul (g (k + (1::nat))) + (real_sub (f (k + (1::nat))) (f k)))) + else real_of_num (0::nat))" + by (import hollight SUM_PARTIAL_SUC) + +lemma SUM_PARTIAL_PRE: "hollight.sum {m::nat..n::nat} + (%k::nat. + real_mul ((f::nat => hollight.real) k) + (real_sub ((g::nat => hollight.real) k) (g (k - (1::nat))))) = +(if m <= n + then real_sub + (real_sub (real_mul (f (n + (1::nat))) (g n)) + (real_mul (f m) (g (m - (1::nat))))) + (hollight.sum {m..n} + (%k::nat. real_mul (g k) (real_sub (f (k + (1::nat))) (f k)))) + else real_of_num (0::nat))" + by (import hollight SUM_PARTIAL_PRE) + +lemma SUM_DIFFS: "hollight.sum {x::nat..xa::nat} + (%x::nat. real_sub ((f::nat => hollight.real) x) (f (x + (1::nat)))) = +(if x <= xa then real_sub (f x) (f (xa + (1::nat))) + else real_of_num (0::nat))" + by (import hollight SUM_DIFFS) + +lemma SUM_DIFFS_ALT: "hollight.sum {m::nat..n::nat} + (%x::nat. real_sub ((f::nat => hollight.real) (x + (1::nat))) (f x)) = +(if m <= n then real_sub (f (n + (1::nat))) (f m) else real_of_num (0::nat))" + by (import hollight SUM_DIFFS_ALT) + +lemma SUM_COMBINE_R: "(m::nat) <= (n::nat) + (1::nat) & n <= (p::nat) +==> real_add (hollight.sum {m..n} (f::nat => hollight.real)) + (hollight.sum {n + (1::nat)..p} f) = + hollight.sum {m..p} f" + by (import hollight SUM_COMBINE_R) + +lemma SUM_COMBINE_L: "(0::nat) < (n::nat) & (m::nat) <= n & n <= (p::nat) + (1::nat) +==> real_add (hollight.sum {m..n - (1::nat)} (f::nat => hollight.real)) + (hollight.sum {n..p} f) = + hollight.sum {m..p} f" + by (import hollight SUM_COMBINE_L) + +lemma REAL_SUB_POW: "1 <= xb +==> real_sub (real_pow x xb) (real_pow xa xb) = + real_mul (real_sub x xa) + (hollight.sum {0..xb - 1} + (%i. real_mul (real_pow x i) (real_pow xa (xb - 1 - i))))" + by (import hollight REAL_SUB_POW) + +lemma REAL_SUB_POW_R1: "1 <= n +==> real_sub (real_pow x n) (real_of_num 1) = + real_mul (real_sub x (real_of_num 1)) + (hollight.sum {0..n - 1} (real_pow x))" + by (import hollight REAL_SUB_POW_R1) + +lemma REAL_SUB_POW_L1: "1 <= xa +==> real_sub (real_of_num 1) (real_pow x xa) = + real_mul (real_sub (real_of_num 1) x) + (hollight.sum {0..xa - 1} (real_pow x))" + by (import hollight REAL_SUB_POW_L1) + +definition + dimindex :: "('A => bool) => nat" where + "(op ==::(('A::type => bool) => nat) => (('A::type => bool) => nat) => prop) + (dimindex::('A::type => bool) => nat) + (%u::'A::type => bool. + (If::bool => nat => nat => nat) + ((finite::('A::type => bool) => bool) (UNIV::'A::type => bool)) + ((CARD::('A::type => bool) => nat) (UNIV::'A::type => bool)) (1::nat))" + +lemma DEF_dimindex: "(op =::(('A::type => bool) => nat) => (('A::type => bool) => nat) => bool) + (dimindex::('A::type => bool) => nat) + (%u::'A::type => bool. + (If::bool => nat => nat => nat) + ((finite::('A::type => bool) => bool) (UNIV::'A::type => bool)) + ((CARD::('A::type => bool) => nat) (UNIV::'A::type => bool)) (1::nat))" + by (import hollight DEF_dimindex) + +lemma DIMINDEX_NONZERO: "dimindex (s::'A => bool) ~= (0::nat)" + by (import hollight DIMINDEX_NONZERO) + +lemma DIMINDEX_GE_1: "(1::nat) <= dimindex (x::'A => bool)" + by (import hollight DIMINDEX_GE_1) + +lemma DIMINDEX_UNIV: "(op =::nat => nat => bool) + ((dimindex::('A::type => bool) => nat) (x::'A::type => bool)) + ((dimindex::('A::type => bool) => nat) (UNIV::'A::type => bool))" + by (import hollight DIMINDEX_UNIV) + +lemma DIMINDEX_UNIQUE: "(op ==>::prop => prop => prop) + ((Trueprop::bool => prop) + ((HAS_SIZE::('A::type => bool) => nat => bool) (UNIV::'A::type => bool) + (n::nat))) + ((Trueprop::bool => prop) + ((op =::nat => nat => bool) + ((dimindex::('A::type => bool) => nat) (UNIV::'A::type => bool)) n))" + by (import hollight DIMINDEX_UNIQUE) + +typedef (open) ('A) finite_image = "{x::nat. x : dotdot (NUMERAL (NUMERAL_BIT1 (0::nat))) (dimindex UNIV)}" morphisms "dest_finite_image" "finite_index" + apply (rule light_ex_imp_nonempty[where t="SOME x::nat. x : dotdot (NUMERAL (NUMERAL_BIT1 (0::nat))) (dimindex UNIV)"]) by (import hollight TYDEF_finite_image) syntax dest_finite_image :: _ syntax - mk_finite_image :: _ + finite_index :: _ lemmas "TYDEF_finite_image_@intern" = typedef_hol2hollight [where a="a :: 'A finite_image" and r=r , @@ -6914,160 +6644,61 @@ lemma FINITE_IMAGE_IMAGE: "(op =::('A::type finite_image => bool) => ('A::type finite_image => bool) => bool) - (hollight.UNIV::'A::type finite_image => bool) - ((IMAGE::('A::type => 'A::type finite_image) - => ('A::type => bool) => 'A::type finite_image => bool) - (mk_finite_image::'A::type => 'A::type finite_image) - ((COND::bool - => ('A::type => bool) => ('A::type => bool) => 'A::type => bool) - ((FINITE::('A::type => bool) => bool) - (hollight.UNIV::'A::type => bool)) - (hollight.UNIV::'A::type => bool) - ((INSERT::'A::type => ('A::type => bool) => 'A::type => bool) - ((Eps::('A::type => bool) => 'A::type) (%z::'A::type. True::bool)) - (EMPTY::'A::type => bool))))" + (UNIV::'A::type finite_image => bool) + ((op `::(nat => 'A::type finite_image) + => (nat => bool) => 'A::type finite_image => bool) + (finite_index::nat => 'A::type finite_image) + ((atLeastAtMost::nat => nat => nat => bool) (1::nat) + ((dimindex::('A::type => bool) => nat) (UNIV::'A::type => bool))))" by (import hollight FINITE_IMAGE_IMAGE) -definition dimindex :: "('A => bool) => nat" where - "(op ==::(('A::type => bool) => nat) => (('A::type => bool) => nat) => prop) - (dimindex::('A::type => bool) => nat) - (%u::'A::type => bool. - (COND::bool => nat => nat => nat) - ((FINITE::('A::type => bool) => bool) - (hollight.UNIV::'A::type => bool)) - ((CARD::('A::type => bool) => nat) (hollight.UNIV::'A::type => bool)) - ((NUMERAL_BIT1::nat => nat) (0::nat)))" - -lemma DEF_dimindex: "(op =::(('A::type => bool) => nat) => (('A::type => bool) => nat) => bool) - (dimindex::('A::type => bool) => nat) - (%u::'A::type => bool. - (COND::bool => nat => nat => nat) - ((FINITE::('A::type => bool) => bool) - (hollight.UNIV::'A::type => bool)) - ((CARD::('A::type => bool) => nat) (hollight.UNIV::'A::type => bool)) - ((NUMERAL_BIT1::nat => nat) (0::nat)))" - by (import hollight DEF_dimindex) - -lemma HAS_SIZE_FINITE_IMAGE: "(All::(('A::type => bool) => bool) => bool) - (%s::'A::type => bool. - (HAS_SIZE::('A::type finite_image => bool) => nat => bool) - (hollight.UNIV::'A::type finite_image => bool) - ((dimindex::('A::type => bool) => nat) s))" +lemma HAS_SIZE_FINITE_IMAGE: "(HAS_SIZE::('A::type finite_image => bool) => nat => bool) + (UNIV::'A::type finite_image => bool) + ((dimindex::('A::type => bool) => nat) (s::'A::type => bool))" by (import hollight HAS_SIZE_FINITE_IMAGE) -lemma CARD_FINITE_IMAGE: "(All::(('A::type => bool) => bool) => bool) - (%s::'A::type => bool. - (op =::nat => nat => bool) - ((CARD::('A::type finite_image => bool) => nat) - (hollight.UNIV::'A::type finite_image => bool)) - ((dimindex::('A::type => bool) => nat) s))" +lemma CARD_FINITE_IMAGE: "(op =::nat => nat => bool) + ((CARD::('A::type finite_image => bool) => nat) + (UNIV::'A::type finite_image => bool)) + ((dimindex::('A::type => bool) => nat) (s::'A::type => bool))" by (import hollight CARD_FINITE_IMAGE) -lemma FINITE_FINITE_IMAGE: "(FINITE::('A::type finite_image => bool) => bool) - (hollight.UNIV::'A::type finite_image => bool)" +lemma FINITE_FINITE_IMAGE: "(finite::('A::type finite_image => bool) => bool) + (UNIV::'A::type finite_image => bool)" by (import hollight FINITE_FINITE_IMAGE) -lemma DIMINDEX_NONZERO: "ALL s::'A::type => bool. dimindex s ~= 0" - by (import hollight DIMINDEX_NONZERO) - -lemma DIMINDEX_GE_1: "ALL x::'A::type => bool. <= (NUMERAL_BIT1 0) (dimindex x)" - by (import hollight DIMINDEX_GE_1) - -lemma DIMINDEX_FINITE_IMAGE: "ALL (s::'A::type finite_image => bool) t::'A::type => bool. - dimindex s = dimindex t" +lemma DIMINDEX_FINITE_IMAGE: "dimindex (s::'A finite_image => bool) = dimindex (t::'A => bool)" by (import hollight DIMINDEX_FINITE_IMAGE) -definition finite_index :: "nat => 'A" where - "(op ==::(nat => 'A::type) => (nat => 'A::type) => prop) - (finite_index::nat => 'A::type) - ((Eps::((nat => 'A::type) => bool) => nat => 'A::type) - (%f::nat => 'A::type. - (All::('A::type => bool) => bool) - (%x::'A::type. - (Ex1::(nat => bool) => bool) - (%n::nat. - (op &::bool => bool => bool) - ((<=::nat => nat => bool) - ((NUMERAL_BIT1::nat => nat) (0::nat)) n) - ((op &::bool => bool => bool) - ((<=::nat => nat => bool) n - ((dimindex::('A::type => bool) => nat) - (hollight.UNIV::'A::type => bool))) - ((op =::'A::type => 'A::type => bool) (f n) x))))))" - -lemma DEF_finite_index: "(op =::(nat => 'A::type) => (nat => 'A::type) => bool) - (finite_index::nat => 'A::type) - ((Eps::((nat => 'A::type) => bool) => nat => 'A::type) - (%f::nat => 'A::type. - (All::('A::type => bool) => bool) - (%x::'A::type. - (Ex1::(nat => bool) => bool) - (%n::nat. - (op &::bool => bool => bool) - ((<=::nat => nat => bool) - ((NUMERAL_BIT1::nat => nat) (0::nat)) n) - ((op &::bool => bool => bool) - ((<=::nat => nat => bool) n - ((dimindex::('A::type => bool) => nat) - (hollight.UNIV::'A::type => bool))) - ((op =::'A::type => 'A::type => bool) (f n) x))))))" - by (import hollight DEF_finite_index) - -lemma FINITE_INDEX_WORKS_FINITE: "(op -->::bool => bool => bool) - ((FINITE::('N::type => bool) => bool) (hollight.UNIV::'N::type => bool)) - ((All::('N::type => bool) => bool) - (%x::'N::type. - (Ex1::(nat => bool) => bool) - (%xa::nat. - (op &::bool => bool => bool) - ((<=::nat => nat => bool) ((NUMERAL_BIT1::nat => nat) (0::nat)) - xa) - ((op &::bool => bool => bool) - ((<=::nat => nat => bool) xa - ((dimindex::('N::type => bool) => nat) - (hollight.UNIV::'N::type => bool))) - ((op =::'N::type => 'N::type => bool) - ((finite_index::nat => 'N::type) xa) x)))))" - by (import hollight FINITE_INDEX_WORKS_FINITE) - -lemma FINITE_INDEX_WORKS: "(All::('A::type finite_image => bool) => bool) - (%i::'A::type finite_image. - (Ex1::(nat => bool) => bool) - (%n::nat. - (op &::bool => bool => bool) - ((<=::nat => nat => bool) ((NUMERAL_BIT1::nat => nat) (0::nat)) - n) - ((op &::bool => bool => bool) - ((<=::nat => nat => bool) n - ((dimindex::('A::type => bool) => nat) - (hollight.UNIV::'A::type => bool))) - ((op =::'A::type finite_image => 'A::type finite_image => bool) - ((finite_index::nat => 'A::type finite_image) n) i))))" +lemma FINITE_INDEX_WORKS: "(Ex1::(nat => bool) => bool) + (%xa::nat. + (op &::bool => bool => bool) ((op <=::nat => nat => bool) (1::nat) xa) + ((op &::bool => bool => bool) + ((op <=::nat => nat => bool) xa + ((dimindex::('A::type => bool) => nat) (UNIV::'A::type => bool))) + ((op =::'A::type finite_image => 'A::type finite_image => bool) + ((finite_index::nat => 'A::type finite_image) xa) + (x::'A::type finite_image))))" by (import hollight FINITE_INDEX_WORKS) -lemma FINITE_INDEX_INJ: "(All::(nat => bool) => bool) - (%x::nat. - (All::(nat => bool) => bool) - (%xa::nat. - (op -->::bool => bool => bool) - ((op &::bool => bool => bool) - ((<=::nat => nat => bool) ((NUMERAL_BIT1::nat => nat) (0::nat)) - x) - ((op &::bool => bool => bool) - ((<=::nat => nat => bool) x - ((dimindex::('A::type => bool) => nat) - (hollight.UNIV::'A::type => bool))) - ((op &::bool => bool => bool) - ((<=::nat => nat => bool) - ((NUMERAL_BIT1::nat => nat) (0::nat)) xa) - ((<=::nat => nat => bool) xa - ((dimindex::('A::type => bool) => nat) - (hollight.UNIV::'A::type => bool)))))) - ((op =::bool => bool => bool) - ((op =::'A::type => 'A::type => bool) - ((finite_index::nat => 'A::type) x) - ((finite_index::nat => 'A::type) xa)) - ((op =::nat => nat => bool) x xa))))" +lemma FINITE_INDEX_INJ: "(op ==>::prop => prop => prop) + ((Trueprop::bool => prop) + ((op &::bool => bool => bool) + ((op <=::nat => nat => bool) (1::nat) (i::nat)) + ((op &::bool => bool => bool) + ((op <=::nat => nat => bool) i + ((dimindex::('A::type => bool) => nat) (UNIV::'A::type => bool))) + ((op &::bool => bool => bool) + ((op <=::nat => nat => bool) (1::nat) (j::nat)) + ((op <=::nat => nat => bool) j + ((dimindex::('A::type => bool) => nat) + (UNIV::'A::type => bool))))))) + ((Trueprop::bool => prop) + ((op =::bool => bool => bool) + ((op =::'A::type finite_image => 'A::type finite_image => bool) + ((finite_index::nat => 'A::type finite_image) i) + ((finite_index::nat => 'A::type finite_image) j)) + ((op =::nat => nat => bool) i j)))" by (import hollight FINITE_INDEX_INJ) lemma FORALL_FINITE_INDEX: "(op =::bool => bool => bool) @@ -7077,19 +6708,15 @@ (%i::nat. (op -->::bool => bool => bool) ((op &::bool => bool => bool) - ((<=::nat => nat => bool) ((NUMERAL_BIT1::nat => nat) (0::nat)) i) - ((<=::nat => nat => bool) i + ((op <=::nat => nat => bool) (1::nat) i) + ((op <=::nat => nat => bool) i ((dimindex::('N::type => bool) => nat) - (hollight.UNIV::'N::type => bool)))) + (UNIV::'N::type => bool)))) (P ((finite_index::nat => 'N::type finite_image) i))))" by (import hollight FORALL_FINITE_INDEX) -typedef (open) ('A, 'B) cart = "(Collect::(('B::type finite_image => 'A::type) => bool) - => ('B::type finite_image => 'A::type) set) - (%f::'B::type finite_image => 'A::type. True::bool)" morphisms "dest_cart" "mk_cart" - apply (rule light_ex_imp_nonempty[where t="(Eps::(('B::type finite_image => 'A::type) => bool) - => 'B::type finite_image => 'A::type) - (%f::'B::type finite_image => 'A::type. True::bool)"]) +typedef (open) ('A, 'B) cart = "{f. True}" morphisms "dest_cart" "mk_cart" + apply (rule light_ex_imp_nonempty[where t="SOME f. True"]) by (import hollight TYDEF_cart) syntax @@ -7103,42 +6730,32 @@ OF type_definition_cart] consts - "$" :: "('q_46627, 'q_46634) cart => nat => 'q_46627" ("$") + "$" :: "('q_73536, 'q_73546) cart => nat => 'q_73536" ("$") defs - "$_def": "$ == -%(u::('q_46627::type, 'q_46634::type) cart) ua::nat. - dest_cart u (finite_index ua)" - -lemma "DEF_$": "$ = -(%(u::('q_46627::type, 'q_46634::type) cart) ua::nat. - dest_cart u (finite_index ua))" + "$_def": "$ == %(u::('q_73536, 'q_73546) cart) ua::nat. dest_cart u (finite_index ua)" + +lemma "DEF_$": "$ = (%(u::('q_73536, 'q_73546) cart) ua::nat. dest_cart u (finite_index ua))" by (import hollight "DEF_$") -lemma CART_EQ: "(All::(('A::type, 'B::type) cart => bool) => bool) - (%x::('A::type, 'B::type) cart. - (All::(('A::type, 'B::type) cart => bool) => bool) - (%y::('A::type, 'B::type) cart. - (op =::bool => bool => bool) - ((op =::('A::type, 'B::type) cart - => ('A::type, 'B::type) cart => bool) - x y) - ((All::(nat => bool) => bool) - (%xa::nat. - (op -->::bool => bool => bool) - ((op &::bool => bool => bool) - ((<=::nat => nat => bool) - ((NUMERAL_BIT1::nat => nat) (0::nat)) xa) - ((<=::nat => nat => bool) xa - ((dimindex::('B::type => bool) => nat) - (hollight.UNIV::'B::type => bool)))) - ((op =::'A::type => 'A::type => bool) - (($::('A::type, 'B::type) cart => nat => 'A::type) x xa) - (($::('A::type, 'B::type) cart => nat => 'A::type) y - xa))))))" +lemma CART_EQ: "(op =::bool => bool => bool) + ((op =::('A::type, 'B::type) cart => ('A::type, 'B::type) cart => bool) + (x::('A::type, 'B::type) cart) (y::('A::type, 'B::type) cart)) + ((All::(nat => bool) => bool) + (%xa::nat. + (op -->::bool => bool => bool) + ((op &::bool => bool => bool) + ((op <=::nat => nat => bool) (1::nat) xa) + ((op <=::nat => nat => bool) xa + ((dimindex::('B::type => bool) => nat) + (UNIV::'B::type => bool)))) + ((op =::'A::type => 'A::type => bool) + (($::('A::type, 'B::type) cart => nat => 'A::type) x xa) + (($::('A::type, 'B::type) cart => nat => 'A::type) y xa))))" by (import hollight CART_EQ) -definition lambda :: "(nat => 'A) => ('A, 'B) cart" where +definition + lambda :: "(nat => 'A) => ('A, 'B) cart" where "(op ==::((nat => 'A::type) => ('A::type, 'B::type) cart) => ((nat => 'A::type) => ('A::type, 'B::type) cart) => prop) (lambda::(nat => 'A::type) => ('A::type, 'B::type) cart) @@ -7149,11 +6766,10 @@ (%i::nat. (op -->::bool => bool => bool) ((op &::bool => bool => bool) - ((<=::nat => nat => bool) - ((NUMERAL_BIT1::nat => nat) (0::nat)) i) - ((<=::nat => nat => bool) i + ((op <=::nat => nat => bool) (1::nat) i) + ((op <=::nat => nat => bool) i ((dimindex::('B::type => bool) => nat) - (hollight.UNIV::'B::type => bool)))) + (UNIV::'B::type => bool)))) ((op =::'A::type => 'A::type => bool) (($::('A::type, 'B::type) cart => nat => 'A::type) f i) (u i)))))" @@ -7168,64 +6784,73 @@ (%i::nat. (op -->::bool => bool => bool) ((op &::bool => bool => bool) - ((<=::nat => nat => bool) - ((NUMERAL_BIT1::nat => nat) (0::nat)) i) - ((<=::nat => nat => bool) i + ((op <=::nat => nat => bool) (1::nat) i) + ((op <=::nat => nat => bool) i ((dimindex::('B::type => bool) => nat) - (hollight.UNIV::'B::type => bool)))) + (UNIV::'B::type => bool)))) ((op =::'A::type => 'A::type => bool) (($::('A::type, 'B::type) cart => nat => 'A::type) f i) (u i)))))" by (import hollight DEF_lambda) -lemma LAMBDA_BETA: "(All::(nat => bool) => bool) - (%x::nat. - (op -->::bool => bool => bool) - ((op &::bool => bool => bool) - ((<=::nat => nat => bool) ((NUMERAL_BIT1::nat => nat) (0::nat)) x) - ((<=::nat => nat => bool) x - ((dimindex::('B::type => bool) => nat) - (hollight.UNIV::'B::type => bool)))) - ((op =::'A::type => 'A::type => bool) - (($::('A::type, 'B::type) cart => nat => 'A::type) - ((lambda::(nat => 'A::type) => ('A::type, 'B::type) cart) - (g::nat => 'A::type)) - x) - (g x)))" +lemma LAMBDA_BETA: "(op ==>::prop => prop => prop) + ((Trueprop::bool => prop) + ((op &::bool => bool => bool) + ((op <=::nat => nat => bool) (1::nat) (x::nat)) + ((op <=::nat => nat => bool) x + ((dimindex::('B::type => bool) => nat) (UNIV::'B::type => bool))))) + ((Trueprop::bool => prop) + ((op =::'A::type => 'A::type => bool) + (($::('A::type, 'B::type) cart => nat => 'A::type) + ((lambda::(nat => 'A::type) => ('A::type, 'B::type) cart) + (g::nat => 'A::type)) + x) + (g x)))" by (import hollight LAMBDA_BETA) -lemma LAMBDA_UNIQUE: "(All::(('A::type, 'B::type) cart => bool) => bool) - (%x::('A::type, 'B::type) cart. - (All::((nat => 'A::type) => bool) => bool) - (%xa::nat => 'A::type. - (op =::bool => bool => bool) - ((All::(nat => bool) => bool) - (%i::nat. - (op -->::bool => bool => bool) - ((op &::bool => bool => bool) - ((<=::nat => nat => bool) - ((NUMERAL_BIT1::nat => nat) (0::nat)) i) - ((<=::nat => nat => bool) i - ((dimindex::('B::type => bool) => nat) - (hollight.UNIV::'B::type => bool)))) - ((op =::'A::type => 'A::type => bool) - (($::('A::type, 'B::type) cart => nat => 'A::type) x i) - (xa i)))) - ((op =::('A::type, 'B::type) cart - => ('A::type, 'B::type) cart => bool) - ((lambda::(nat => 'A::type) => ('A::type, 'B::type) cart) xa) - x)))" +lemma LAMBDA_UNIQUE: "(op =::bool => bool => bool) + ((All::(nat => bool) => bool) + (%i::nat. + (op -->::bool => bool => bool) + ((op &::bool => bool => bool) + ((op <=::nat => nat => bool) (1::nat) i) + ((op <=::nat => nat => bool) i + ((dimindex::('B::type => bool) => nat) + (UNIV::'B::type => bool)))) + ((op =::'A::type => 'A::type => bool) + (($::('A::type, 'B::type) cart => nat => 'A::type) + (x::('A::type, 'B::type) cart) i) + ((xa::nat => 'A::type) i)))) + ((op =::('A::type, 'B::type) cart => ('A::type, 'B::type) cart => bool) + ((lambda::(nat => 'A::type) => ('A::type, 'B::type) cart) xa) x)" by (import hollight LAMBDA_UNIQUE) -lemma LAMBDA_ETA: "ALL x::('q_46825::type, 'q_46829::type) cart. lambda ($ x) = x" +lemma LAMBDA_ETA: "lambda ($ (x::('q_73734, 'q_73738) cart)) = x" by (import hollight LAMBDA_ETA) -typedef (open) ('A, 'B) finite_sum = "(Collect::('A::type finite_image + 'B::type finite_image => bool) - => ('A::type finite_image + 'B::type finite_image) set) - (%x::'A::type finite_image + 'B::type finite_image. True::bool)" morphisms "dest_finite_sum" "mk_finite_sum" - apply (rule light_ex_imp_nonempty[where t="(Eps::('A::type finite_image + 'B::type finite_image => bool) - => 'A::type finite_image + 'B::type finite_image) - (%x::'A::type finite_image + 'B::type finite_image. True::bool)"]) +lemma FINITE_INDEX_INRANGE: "(Ex::(nat => bool) => bool) + (%xa::nat. + (op &::bool => bool => bool) ((op <=::nat => nat => bool) (1::nat) xa) + ((op &::bool => bool => bool) + ((op <=::nat => nat => bool) xa + ((dimindex::('N::type => bool) => nat) (UNIV::'N::type => bool))) + ((All::(('A::type, 'N::type) cart => bool) => bool) + (%xb::('A::type, 'N::type) cart. + (op =::'A::type => 'A::type => bool) + (($::('A::type, 'N::type) cart => nat => 'A::type) xb + (x::nat)) + (($::('A::type, 'N::type) cart => nat => 'A::type) xb xa)))))" + by (import hollight FINITE_INDEX_INRANGE) + +lemma CART_EQ_FULL: "((x::('A, 'N) cart) = (y::('A, 'N) cart)) = (ALL i::nat. $ x i = $ y i)" + by (import hollight CART_EQ_FULL) + +typedef (open) ('A, 'B) finite_sum = "{x::nat. + x : dotdot (NUMERAL (NUMERAL_BIT1 (0::nat))) + (dimindex UNIV + dimindex UNIV)}" morphisms "dest_finite_sum" "mk_finite_sum" + apply (rule light_ex_imp_nonempty[where t="SOME x::nat. + x : dotdot (NUMERAL (NUMERAL_BIT1 (0::nat))) + (dimindex UNIV + dimindex UNIV)"]) by (import hollight TYDEF_finite_sum) syntax @@ -7238,7 +6863,8 @@ [where a="a :: ('A, 'B) finite_sum" and r=r , OF type_definition_finite_sum] -definition pastecart :: "('A, 'M) cart => ('A, 'N) cart => ('A, ('M, 'N) finite_sum) cart" where +definition + pastecart :: "('A, 'M) cart => ('A, 'N) cart => ('A, ('M, 'N) finite_sum) cart" where "(op ==::(('A::type, 'M::type) cart => ('A::type, 'N::type) cart => ('A::type, ('M::type, 'N::type) finite_sum) cart) @@ -7253,15 +6879,15 @@ (lambda::(nat => 'A::type) => ('A::type, ('M::type, 'N::type) finite_sum) cart) (%i::nat. - (COND::bool => 'A::type => 'A::type => 'A::type) - ((<=::nat => nat => bool) i + (If::bool => 'A::type => 'A::type => 'A::type) + ((op <=::nat => nat => bool) i ((dimindex::('M::type => bool) => nat) - (hollight.UNIV::'M::type => bool))) + (UNIV::'M::type => bool))) (($::('A::type, 'M::type) cart => nat => 'A::type) u i) (($::('A::type, 'N::type) cart => nat => 'A::type) ua ((op -::nat => nat => nat) i ((dimindex::('M::type => bool) => nat) - (hollight.UNIV::'M::type => bool))))))" + (UNIV::'M::type => bool))))))" lemma DEF_pastecart: "(op =::(('A::type, 'M::type) cart => ('A::type, 'N::type) cart @@ -7277,26 +6903,26 @@ (lambda::(nat => 'A::type) => ('A::type, ('M::type, 'N::type) finite_sum) cart) (%i::nat. - (COND::bool => 'A::type => 'A::type => 'A::type) - ((<=::nat => nat => bool) i + (If::bool => 'A::type => 'A::type => 'A::type) + ((op <=::nat => nat => bool) i ((dimindex::('M::type => bool) => nat) - (hollight.UNIV::'M::type => bool))) + (UNIV::'M::type => bool))) (($::('A::type, 'M::type) cart => nat => 'A::type) u i) (($::('A::type, 'N::type) cart => nat => 'A::type) ua ((op -::nat => nat => nat) i ((dimindex::('M::type => bool) => nat) - (hollight.UNIV::'M::type => bool))))))" + (UNIV::'M::type => bool))))))" by (import hollight DEF_pastecart) -definition fstcart :: "('A, ('M, 'N) finite_sum) cart => ('A, 'M) cart" where - "fstcart == -%u::('A::type, ('M::type, 'N::type) finite_sum) cart. lambda ($ u)" - -lemma DEF_fstcart: "fstcart = -(%u::('A::type, ('M::type, 'N::type) finite_sum) cart. lambda ($ u))" +definition + fstcart :: "('A, ('M, 'N) finite_sum) cart => ('A, 'M) cart" where + "fstcart == %u::('A, ('M, 'N) finite_sum) cart. lambda ($ u)" + +lemma DEF_fstcart: "fstcart = (%u::('A, ('M, 'N) finite_sum) cart. lambda ($ u))" by (import hollight DEF_fstcart) -definition sndcart :: "('A, ('M, 'N) finite_sum) cart => ('A, 'N) cart" where +definition + sndcart :: "('A, ('M, 'N) finite_sum) cart => ('A, 'N) cart" where "(op ==::(('A::type, ('M::type, 'N::type) finite_sum) cart => ('A::type, 'N::type) cart) => (('A::type, ('M::type, 'N::type) finite_sum) cart @@ -7309,9 +6935,9 @@ (%i::nat. ($::('A::type, ('M::type, 'N::type) finite_sum) cart => nat => 'A::type) - u ((plus::nat => nat => nat) i + u ((op +::nat => nat => nat) i ((dimindex::('M::type => bool) => nat) - (hollight.UNIV::'M::type => bool)))))" + (UNIV::'M::type => bool)))))" lemma DEF_sndcart: "(op =::(('A::type, ('M::type, 'N::type) finite_sum) cart => ('A::type, 'N::type) cart) @@ -7325,1974 +6951,520 @@ (%i::nat. ($::('A::type, ('M::type, 'N::type) finite_sum) cart => nat => 'A::type) - u ((plus::nat => nat => nat) i + u ((op +::nat => nat => nat) i ((dimindex::('M::type => bool) => nat) - (hollight.UNIV::'M::type => bool)))))" + (UNIV::'M::type => bool)))))" by (import hollight DEF_sndcart) -lemma DIMINDEX_HAS_SIZE_FINITE_SUM: "(HAS_SIZE::(('M::type, 'N::type) finite_sum => bool) => nat => bool) - (hollight.UNIV::('M::type, 'N::type) finite_sum => bool) - ((plus::nat => nat => nat) - ((dimindex::('M::type => bool) => nat) (hollight.UNIV::'M::type => bool)) - ((dimindex::('N::type => bool) => nat) - (hollight.UNIV::'N::type => bool)))" - by (import hollight DIMINDEX_HAS_SIZE_FINITE_SUM) - -lemma DIMINDEX_FINITE_SUM: "(op =::nat => nat => bool) - ((dimindex::(('M::type, 'N::type) finite_sum => bool) => nat) - (hollight.UNIV::('M::type, 'N::type) finite_sum => bool)) - ((plus::nat => nat => nat) - ((dimindex::('M::type => bool) => nat) (hollight.UNIV::'M::type => bool)) - ((dimindex::('N::type => bool) => nat) - (hollight.UNIV::'N::type => bool)))" - by (import hollight DIMINDEX_FINITE_SUM) - -lemma FSTCART_PASTECART: "ALL (x::('A::type, 'M::type) cart) xa::('A::type, 'N::type) cart. - fstcart (pastecart x xa) = x" - by (import hollight FSTCART_PASTECART) - -lemma SNDCART_PASTECART: "ALL (x::('A::type, 'M::type) cart) xa::('A::type, 'N::type) cart. - sndcart (pastecart x xa) = xa" - by (import hollight SNDCART_PASTECART) - -lemma PASTECART_FST_SND: "ALL x::('q_47149::type, ('q_47146::type, 'q_47144::type) finite_sum) cart. - pastecart (fstcart x) (sndcart x) = x" - by (import hollight PASTECART_FST_SND) - -lemma PASTECART_EQ: "ALL (x::('q_47187::type, ('q_47177::type, 'q_47188::type) finite_sum) cart) - y::('q_47187::type, ('q_47177::type, 'q_47188::type) finite_sum) cart. - (x = y) = (fstcart x = fstcart y & sndcart x = sndcart y)" - by (import hollight PASTECART_EQ) - -lemma FORALL_PASTECART: "All (P::('q_47208::type, ('q_47209::type, 'q_47210::type) finite_sum) cart - => bool) = -(ALL (x::('q_47208::type, 'q_47209::type) cart) - y::('q_47208::type, 'q_47210::type) cart. P (pastecart x y))" - by (import hollight FORALL_PASTECART) - -lemma EXISTS_PASTECART: "Ex (P::('q_47230::type, ('q_47231::type, 'q_47232::type) finite_sum) cart - => bool) = -(EX (x::('q_47230::type, 'q_47231::type) cart) - y::('q_47230::type, 'q_47232::type) cart. P (pastecart x y))" - by (import hollight EXISTS_PASTECART) - -lemma SURJECTIVE_IFF_INJECTIVE_GEN: "ALL (s::'A::type => bool) (t::'B::type => bool) f::'A::type => 'B::type. - FINITE s & FINITE t & CARD s = CARD t & SUBSET (IMAGE f s) t --> - (ALL y::'B::type. IN y t --> (EX x::'A::type. IN x s & f x = y)) = - (ALL (x::'A::type) y::'A::type. IN x s & IN y s & f x = f y --> x = y)" - by (import hollight SURJECTIVE_IFF_INJECTIVE_GEN) - -lemma SURJECTIVE_IFF_INJECTIVE: "ALL (x::'A::type => bool) xa::'A::type => 'A::type. - FINITE x & SUBSET (IMAGE xa x) x --> - (ALL y::'A::type. IN y x --> (EX xb::'A::type. IN xb x & xa xb = y)) = - (ALL (xb::'A::type) y::'A::type. - IN xb x & IN y x & xa xb = xa y --> xb = y)" - by (import hollight SURJECTIVE_IFF_INJECTIVE) - -lemma IMAGE_IMP_INJECTIVE_GEN: "ALL (s::'A::type => bool) (t::'B::type => bool) f::'A::type => 'B::type. - FINITE s & CARD s = CARD t & IMAGE f s = t --> - (ALL (x::'A::type) y::'A::type. IN x s & IN y s & f x = f y --> x = y)" - by (import hollight IMAGE_IMP_INJECTIVE_GEN) - -lemma IMAGE_IMP_INJECTIVE: "ALL (s::'q_47557::type => bool) f::'q_47557::type => 'q_47557::type. - FINITE s & IMAGE f s = s --> - (ALL (x::'q_47557::type) y::'q_47557::type. - IN x s & IN y s & f x = f y --> x = y)" - by (import hollight IMAGE_IMP_INJECTIVE) - -lemma CARD_LE_INJ: "ALL (x::'A::type => bool) xa::'B::type => bool. - FINITE x & FINITE xa & <= (CARD x) (CARD xa) --> - (EX f::'A::type => 'B::type. - SUBSET (IMAGE f x) xa & - (ALL (xa::'A::type) y::'A::type. - IN xa x & IN y x & f xa = f y --> xa = y))" - by (import hollight CARD_LE_INJ) - -lemma FORALL_IN_CLAUSES: "(ALL x::'q_47663::type => bool. - (ALL xa::'q_47663::type. IN xa EMPTY --> x xa) = True) & -(ALL (x::'q_47703::type => bool) (xa::'q_47703::type) - xb::'q_47703::type => bool. - (ALL xc::'q_47703::type. IN xc (INSERT xa xb) --> x xc) = - (x xa & (ALL xa::'q_47703::type. IN xa xb --> x xa)))" - by (import hollight FORALL_IN_CLAUSES) - -lemma EXISTS_IN_CLAUSES: "(ALL x::'q_47723::type => bool. - (EX xa::'q_47723::type. IN xa EMPTY & x xa) = False) & -(ALL (x::'q_47763::type => bool) (xa::'q_47763::type) - xb::'q_47763::type => bool. - (EX xc::'q_47763::type. IN xc (INSERT xa xb) & x xc) = - (x xa | (EX xa::'q_47763::type. IN xa xb & x xa)))" - by (import hollight EXISTS_IN_CLAUSES) - -lemma SURJECTIVE_ON_RIGHT_INVERSE: "ALL (x::'q_47819::type => 'q_47820::type) xa::'q_47820::type => bool. - (ALL xb::'q_47820::type. - IN xb xa --> - (EX xa::'q_47819::type. - IN xa (s::'q_47819::type => bool) & x xa = xb)) = - (EX g::'q_47820::type => 'q_47819::type. - ALL y::'q_47820::type. IN y xa --> IN (g y) s & x (g y) = y)" - by (import hollight SURJECTIVE_ON_RIGHT_INVERSE) - -lemma INJECTIVE_ON_LEFT_INVERSE: "ALL (x::'q_47913::type => 'q_47916::type) xa::'q_47913::type => bool. - (ALL (xb::'q_47913::type) y::'q_47913::type. - IN xb xa & IN y xa & x xb = x y --> xb = y) = - (EX xb::'q_47916::type => 'q_47913::type. - ALL xc::'q_47913::type. IN xc xa --> xb (x xc) = xc)" - by (import hollight INJECTIVE_ON_LEFT_INVERSE) - -lemma SURJECTIVE_RIGHT_INVERSE: "(ALL y::'q_47941::type. - EX x::'q_47944::type. (f::'q_47944::type => 'q_47941::type) x = y) = -(EX g::'q_47941::type => 'q_47944::type. ALL y::'q_47941::type. f (g y) = y)" - by (import hollight SURJECTIVE_RIGHT_INVERSE) - -lemma INJECTIVE_LEFT_INVERSE: "(ALL (x::'q_47978::type) xa::'q_47978::type. - (f::'q_47978::type => 'q_47981::type) x = f xa --> x = xa) = -(EX g::'q_47981::type => 'q_47978::type. ALL x::'q_47978::type. g (f x) = x)" - by (import hollight INJECTIVE_LEFT_INVERSE) - -lemma FUNCTION_FACTORS_RIGHT: "ALL (x::'q_48017::type => 'q_48018::type) - xa::'q_48005::type => 'q_48018::type. - (ALL xb::'q_48017::type. EX y::'q_48005::type. xa y = x xb) = - (EX xb::'q_48017::type => 'q_48005::type. x = xa o xb)" - by (import hollight FUNCTION_FACTORS_RIGHT) - -lemma FUNCTION_FACTORS_LEFT: "ALL (x::'q_48090::type => 'q_48091::type) - xa::'q_48090::type => 'q_48070::type. - (ALL (xb::'q_48090::type) y::'q_48090::type. - xa xb = xa y --> x xb = x y) = - (EX xb::'q_48070::type => 'q_48091::type. x = xb o xa)" - by (import hollight FUNCTION_FACTORS_LEFT) - -definition dotdot :: "nat => nat => nat => bool" where - "dotdot == -%(u::nat) ua::nat. - GSPEC (%ub::nat. EX x::nat. SETSPEC ub (<= u x & <= x ua) x)" - -lemma DEF__dot__dot_: "dotdot = -(%(u::nat) ua::nat. - GSPEC (%ub::nat. EX x::nat. SETSPEC ub (<= u x & <= x ua) x))" - by (import hollight DEF__dot__dot_) - -lemma FINITE_NUMSEG: "ALL (m::nat) n::nat. FINITE (dotdot m n)" - by (import hollight FINITE_NUMSEG) - -lemma NUMSEG_COMBINE_R: "ALL (x::'q_48166::type) (p::nat) m::nat. - <= m p & <= p (n::nat) --> - hollight.UNION (dotdot m p) (dotdot (p + NUMERAL_BIT1 0) n) = dotdot m n" - by (import hollight NUMSEG_COMBINE_R) - -lemma NUMSEG_COMBINE_L: "ALL (x::'q_48204::type) (p::nat) m::nat. - <= m p & <= p (n::nat) --> - hollight.UNION (dotdot m (p - NUMERAL_BIT1 0)) (dotdot p n) = dotdot m n" - by (import hollight NUMSEG_COMBINE_L) - -lemma NUMSEG_LREC: "ALL (x::nat) xa::nat. - <= x xa --> INSERT x (dotdot (x + NUMERAL_BIT1 0) xa) = dotdot x xa" - by (import hollight NUMSEG_LREC) - -lemma NUMSEG_RREC: "ALL (x::nat) xa::nat. - <= x xa --> INSERT xa (dotdot x (xa - NUMERAL_BIT1 0)) = dotdot x xa" - by (import hollight NUMSEG_RREC) - -lemma NUMSEG_REC: "ALL (x::nat) xa::nat. - <= x (Suc xa) --> dotdot x (Suc xa) = INSERT (Suc xa) (dotdot x xa)" - by (import hollight NUMSEG_REC) - -lemma IN_NUMSEG: "ALL (x::nat) (xa::nat) xb::nat. IN xb (dotdot x xa) = (<= x xb & <= xb xa)" - by (import hollight IN_NUMSEG) - -lemma NUMSEG_SING: "ALL x::nat. dotdot x x = INSERT x EMPTY" - by (import hollight NUMSEG_SING) - -lemma NUMSEG_EMPTY: "ALL (x::nat) xa::nat. (dotdot x xa = EMPTY) = < xa x" - by (import hollight NUMSEG_EMPTY) - -lemma CARD_NUMSEG_LEMMA: "ALL (m::nat) d::nat. CARD (dotdot m (m + d)) = d + NUMERAL_BIT1 0" - by (import hollight CARD_NUMSEG_LEMMA) - -lemma CARD_NUMSEG: "ALL (m::nat) n::nat. CARD (dotdot m n) = n + NUMERAL_BIT1 0 - m" - by (import hollight CARD_NUMSEG) - -lemma HAS_SIZE_NUMSEG: "ALL (x::nat) xa::nat. HAS_SIZE (dotdot x xa) (xa + NUMERAL_BIT1 0 - x)" - by (import hollight HAS_SIZE_NUMSEG) - -lemma CARD_NUMSEG_1: "ALL x::nat. CARD (dotdot (NUMERAL_BIT1 0) x) = x" - by (import hollight CARD_NUMSEG_1) - -lemma HAS_SIZE_NUMSEG_1: "ALL x::nat. HAS_SIZE (dotdot (NUMERAL_BIT1 0) x) x" - by (import hollight HAS_SIZE_NUMSEG_1) - -lemma NUMSEG_CLAUSES: "(ALL m::nat. dotdot m 0 = COND (m = 0) (INSERT 0 EMPTY) EMPTY) & -(ALL (m::nat) n::nat. - dotdot m (Suc n) = - COND (<= m (Suc n)) (INSERT (Suc n) (dotdot m n)) (dotdot m n))" - by (import hollight NUMSEG_CLAUSES) - -lemma FINITE_INDEX_NUMSEG: "ALL s::'A::type => bool. - FINITE s = - (EX f::nat => 'A::type. - (ALL (i::nat) j::nat. - IN i (dotdot (NUMERAL_BIT1 0) (CARD s)) & - IN j (dotdot (NUMERAL_BIT1 0) (CARD s)) & f i = f j --> - i = j) & - s = IMAGE f (dotdot (NUMERAL_BIT1 0) (CARD s)))" - by (import hollight FINITE_INDEX_NUMSEG) - -lemma FINITE_INDEX_NUMBERS: "ALL s::'A::type => bool. - FINITE s = - (EX (k::nat => bool) f::nat => 'A::type. - (ALL (i::nat) j::nat. IN i k & IN j k & f i = f j --> i = j) & - FINITE k & s = IMAGE f k)" - by (import hollight FINITE_INDEX_NUMBERS) - -lemma DISJOINT_NUMSEG: "ALL (x::nat) (xa::nat) (xb::nat) xc::nat. - DISJOINT (dotdot x xa) (dotdot xb xc) = - (< xa xb | < xc x | < xa x | < xc xb)" - by (import hollight DISJOINT_NUMSEG) - -lemma NUMSEG_ADD_SPLIT: "ALL (x::nat) (xa::nat) xb::nat. - <= x (xa + NUMERAL_BIT1 0) --> - dotdot x (xa + xb) = - hollight.UNION (dotdot x xa) (dotdot (xa + NUMERAL_BIT1 0) (xa + xb))" - by (import hollight NUMSEG_ADD_SPLIT) - -lemma NUMSEG_OFFSET_IMAGE: "ALL (x::nat) (xa::nat) xb::nat. - dotdot (x + xb) (xa + xb) = IMAGE (%i::nat. i + xb) (dotdot x xa)" - by (import hollight NUMSEG_OFFSET_IMAGE) - -lemma SUBSET_NUMSEG: "ALL (m::nat) (n::nat) (p::nat) q::nat. - SUBSET (dotdot m n) (dotdot p q) = (< n m | <= p m & <= n q)" - by (import hollight SUBSET_NUMSEG) - -definition neutral :: "('q_48985 => 'q_48985 => 'q_48985) => 'q_48985" where - "neutral == -%u::'q_48985::type => 'q_48985::type => 'q_48985::type. - SOME x::'q_48985::type. ALL y::'q_48985::type. u x y = y & u y x = y" - -lemma DEF_neutral: "neutral = -(%u::'q_48985::type => 'q_48985::type => 'q_48985::type. - SOME x::'q_48985::type. ALL y::'q_48985::type. u x y = y & u y x = y)" - by (import hollight DEF_neutral) - -definition monoidal :: "('A => 'A => 'A) => bool" where - "monoidal == -%u::'A::type => 'A::type => 'A::type. - (ALL (x::'A::type) y::'A::type. u x y = u y x) & - (ALL (x::'A::type) (y::'A::type) z::'A::type. - u x (u y z) = u (u x y) z) & - (ALL x::'A::type. u (neutral u) x = x)" - -lemma DEF_monoidal: "monoidal = -(%u::'A::type => 'A::type => 'A::type. - (ALL (x::'A::type) y::'A::type. u x y = u y x) & - (ALL (x::'A::type) (y::'A::type) z::'A::type. - u x (u y z) = u (u x y) z) & - (ALL x::'A::type. u (neutral u) x = x))" - by (import hollight DEF_monoidal) - -definition support :: "('B => 'B => 'B) => ('A => 'B) => ('A => bool) => 'A => bool" where - "support == -%(u::'B::type => 'B::type => 'B::type) (ua::'A::type => 'B::type) - ub::'A::type => bool. - GSPEC - (%uc::'A::type. - EX x::'A::type. SETSPEC uc (IN x ub & ua x ~= neutral u) x)" - -lemma DEF_support: "support = -(%(u::'B::type => 'B::type => 'B::type) (ua::'A::type => 'B::type) - ub::'A::type => bool. - GSPEC - (%uc::'A::type. - EX x::'A::type. SETSPEC uc (IN x ub & ua x ~= neutral u) x))" - by (import hollight DEF_support) - -definition iterate :: "('q_49090 => 'q_49090 => 'q_49090) -=> ('A => bool) => ('A => 'q_49090) => 'q_49090" where - "iterate == -%(u::'q_49090::type => 'q_49090::type => 'q_49090::type) - (ua::'A::type => bool) ub::'A::type => 'q_49090::type. - ITSET (%x::'A::type. u (ub x)) (support u ub ua) (neutral u)" - -lemma DEF_iterate: "iterate = -(%(u::'q_49090::type => 'q_49090::type => 'q_49090::type) - (ua::'A::type => bool) ub::'A::type => 'q_49090::type. - ITSET (%x::'A::type. u (ub x)) (support u ub ua) (neutral u))" - by (import hollight DEF_iterate) - -lemma IN_SUPPORT: "ALL (x::'q_49133::type => 'q_49133::type => 'q_49133::type) - (xa::'q_49136::type => 'q_49133::type) (xb::'q_49136::type) - xc::'q_49136::type => bool. - IN xb (support x xa xc) = (IN xb xc & xa xb ~= neutral x)" - by (import hollight IN_SUPPORT) - -lemma SUPPORT_SUPPORT: "ALL (x::'q_49158::type => 'q_49158::type => 'q_49158::type) - (xa::'q_49169::type => 'q_49158::type) xb::'q_49169::type => bool. - support x xa (support x xa xb) = support x xa xb" - by (import hollight SUPPORT_SUPPORT) - -lemma SUPPORT_EMPTY: "ALL (x::'q_49194::type => 'q_49194::type => 'q_49194::type) - (xa::'q_49208::type => 'q_49194::type) xb::'q_49208::type => bool. - (ALL xc::'q_49208::type. IN xc xb --> xa xc = neutral x) = - (support x xa xb = EMPTY)" - by (import hollight SUPPORT_EMPTY) - -lemma SUPPORT_SUBSET: "ALL (x::'q_49228::type => 'q_49228::type => 'q_49228::type) - (xa::'q_49229::type => 'q_49228::type) xb::'q_49229::type => bool. - SUBSET (support x xa xb) xb" - by (import hollight SUPPORT_SUBSET) - -lemma FINITE_SUPPORT: "ALL (u::'q_49252::type => 'q_49252::type => 'q_49252::type) - (f::'q_49246::type => 'q_49252::type) s::'q_49246::type => bool. - FINITE s --> FINITE (support u f s)" - by (import hollight FINITE_SUPPORT) - -lemma SUPPORT_CLAUSES: "(ALL x::'q_49270::type => 'q_49300::type. - support (u_4247::'q_49300::type => 'q_49300::type => 'q_49300::type) x - EMPTY = - EMPTY) & -(ALL (x::'q_49318::type => 'q_49300::type) (xa::'q_49318::type) - xb::'q_49318::type => bool. - support u_4247 x (INSERT xa xb) = - COND (x xa = neutral u_4247) (support u_4247 x xb) - (INSERT xa (support u_4247 x xb))) & -(ALL (x::'q_49351::type => 'q_49300::type) (xa::'q_49351::type) - xb::'q_49351::type => bool. - support u_4247 x (DELETE xb xa) = DELETE (support u_4247 x xb) xa) & -(ALL (x::'q_49389::type => 'q_49300::type) (xa::'q_49389::type => bool) - xb::'q_49389::type => bool. - support u_4247 x (hollight.UNION xa xb) = - hollight.UNION (support u_4247 x xa) (support u_4247 x xb)) & -(ALL (x::'q_49427::type => 'q_49300::type) (xa::'q_49427::type => bool) - xb::'q_49427::type => bool. - support u_4247 x (hollight.INTER xa xb) = - hollight.INTER (support u_4247 x xa) (support u_4247 x xb)) & -(ALL (x::'q_49463::type => 'q_49300::type) (xa::'q_49463::type => bool) - xb::'q_49463::type => bool. - support u_4247 x (DIFF xa xb) = - DIFF (support u_4247 x xa) (support u_4247 x xb))" - by (import hollight SUPPORT_CLAUSES) - -lemma ITERATE_SUPPORT: "ALL (x::'q_49484::type => 'q_49484::type => 'q_49484::type) - (xa::'q_49485::type => 'q_49484::type) xb::'q_49485::type => bool. - FINITE (support x xa xb) --> - iterate x (support x xa xb) xa = iterate x xb xa" - by (import hollight ITERATE_SUPPORT) - -lemma SUPPORT_DELTA: "ALL (x::'q_49529::type => 'q_49529::type => 'q_49529::type) - (xa::'q_49557::type => bool) (xb::'q_49557::type => 'q_49529::type) - xc::'q_49557::type. - support x (%xa::'q_49557::type. COND (xa = xc) (xb xa) (neutral x)) xa = - COND (IN xc xa) (support x xb (INSERT xc EMPTY)) EMPTY" - by (import hollight SUPPORT_DELTA) - -lemma FINITE_SUPPORT_DELTA: "ALL (x::'q_49578::type => 'q_49578::type => 'q_49578::type) - (xa::'q_49587::type => 'q_49578::type) xb::'q_49587::type. - FINITE - (support x (%xc::'q_49587::type. COND (xc = xb) (xa xc) (neutral x)) - (s::'q_49587::type => bool))" - by (import hollight FINITE_SUPPORT_DELTA) - -lemma ITERATE_CLAUSES_GEN: "ALL u_4247::'B::type => 'B::type => 'B::type. - monoidal u_4247 --> - (ALL f::'A::type => 'B::type. iterate u_4247 EMPTY f = neutral u_4247) & - (ALL (f::'A::type => 'B::type) (x::'A::type) s::'A::type => bool. - monoidal u_4247 & FINITE (support u_4247 f s) --> - iterate u_4247 (INSERT x s) f = - COND (IN x s) (iterate u_4247 s f) - (u_4247 (f x) (iterate u_4247 s f)))" - by (import hollight ITERATE_CLAUSES_GEN) - -lemma ITERATE_CLAUSES: "ALL x::'q_49755::type => 'q_49755::type => 'q_49755::type. - monoidal x --> - (ALL f::'q_49713::type => 'q_49755::type. - iterate x EMPTY f = neutral x) & - (ALL (f::'q_49757::type => 'q_49755::type) (xa::'q_49757::type) - s::'q_49757::type => bool. - FINITE s --> - iterate x (INSERT xa s) f = - COND (IN xa s) (iterate x s f) (x (f xa) (iterate x s f)))" - by (import hollight ITERATE_CLAUSES) - -lemma ITERATE_UNION: "ALL u_4247::'q_49843::type => 'q_49843::type => 'q_49843::type. - monoidal u_4247 --> - (ALL (f::'q_49828::type => 'q_49843::type) (s::'q_49828::type => bool) - x::'q_49828::type => bool. - FINITE s & FINITE x & DISJOINT s x --> - iterate u_4247 (hollight.UNION s x) f = - u_4247 (iterate u_4247 s f) (iterate u_4247 x f))" - by (import hollight ITERATE_UNION) - -lemma ITERATE_UNION_GEN: "ALL u_4247::'B::type => 'B::type => 'B::type. - monoidal u_4247 --> - (ALL (f::'A::type => 'B::type) (s::'A::type => bool) t::'A::type => bool. - FINITE (support u_4247 f s) & - FINITE (support u_4247 f t) & - DISJOINT (support u_4247 f s) (support u_4247 f t) --> - iterate u_4247 (hollight.UNION s t) f = - u_4247 (iterate u_4247 s f) (iterate u_4247 t f))" - by (import hollight ITERATE_UNION_GEN) - -lemma ITERATE_DIFF: "ALL u::'q_49986::type => 'q_49986::type => 'q_49986::type. - monoidal u --> - (ALL (f::'q_49982::type => 'q_49986::type) (s::'q_49982::type => bool) - t::'q_49982::type => bool. - FINITE s & SUBSET t s --> - u (iterate u (DIFF s t) f) (iterate u t f) = iterate u s f)" - by (import hollight ITERATE_DIFF) - -lemma ITERATE_DIFF_GEN: "ALL u_4247::'B::type => 'B::type => 'B::type. - monoidal u_4247 --> - (ALL (f::'A::type => 'B::type) (s::'A::type => bool) t::'A::type => bool. - FINITE (support u_4247 f s) & - SUBSET (support u_4247 f t) (support u_4247 f s) --> - u_4247 (iterate u_4247 (DIFF s t) f) (iterate u_4247 t f) = - iterate u_4247 s f)" - by (import hollight ITERATE_DIFF_GEN) - -lemma ITERATE_CLOSED: "ALL u_4247::'B::type => 'B::type => 'B::type. - monoidal u_4247 --> - (ALL P::'B::type => bool. - P (neutral u_4247) & - (ALL (x::'B::type) y::'B::type. P x & P y --> P (u_4247 x y)) --> - (ALL (f::'A::type => 'B::type) x::'A::type => bool. - FINITE x & (ALL xa::'A::type. IN xa x --> P (f xa)) --> - P (iterate u_4247 x f)))" - by (import hollight ITERATE_CLOSED) - -lemma ITERATE_CLOSED_GEN: "ALL u_4247::'B::type => 'B::type => 'B::type. - monoidal u_4247 --> - (ALL P::'B::type => bool. - P (neutral u_4247) & - (ALL (x::'B::type) y::'B::type. P x & P y --> P (u_4247 x y)) --> - (ALL (f::'A::type => 'B::type) s::'A::type => bool. - FINITE (support u_4247 f s) & - (ALL x::'A::type. IN x s & f x ~= neutral u_4247 --> P (f x)) --> - P (iterate u_4247 s f)))" - by (import hollight ITERATE_CLOSED_GEN) - -lemma ITERATE_RELATED: "ALL u_4247::'B::type => 'B::type => 'B::type. - monoidal u_4247 --> - (ALL R::'B::type => 'B::type => bool. - R (neutral u_4247) (neutral u_4247) & - (ALL (x1::'B::type) (y1::'B::type) (x2::'B::type) y2::'B::type. - R x1 x2 & R y1 y2 --> R (u_4247 x1 y1) (u_4247 x2 y2)) --> - (ALL (f::'A::type => 'B::type) (g::'A::type => 'B::type) - x::'A::type => bool. - FINITE x & (ALL xa::'A::type. IN xa x --> R (f xa) (g xa)) --> - R (iterate u_4247 x f) (iterate u_4247 x g)))" - by (import hollight ITERATE_RELATED) - -lemma ITERATE_EQ_NEUTRAL: "ALL u_4247::'B::type => 'B::type => 'B::type. - monoidal u_4247 --> - (ALL (f::'A::type => 'B::type) s::'A::type => bool. - (ALL x::'A::type. IN x s --> f x = neutral u_4247) --> - iterate u_4247 s f = neutral u_4247)" - by (import hollight ITERATE_EQ_NEUTRAL) - -lemma ITERATE_SING: "ALL x::'B::type => 'B::type => 'B::type. - monoidal x --> - (ALL (f::'A::type => 'B::type) xa::'A::type. - iterate x (INSERT xa EMPTY) f = f xa)" - by (import hollight ITERATE_SING) - -lemma ITERATE_DELTA: "ALL u_4247::'q_50438::type => 'q_50438::type => 'q_50438::type. - monoidal u_4247 --> - (ALL (x::'q_50457::type => 'q_50438::type) (xa::'q_50457::type) - xb::'q_50457::type => bool. - iterate u_4247 xb - (%xb::'q_50457::type. COND (xb = xa) (x xb) (neutral u_4247)) = - COND (IN xa xb) (x xa) (neutral u_4247))" - by (import hollight ITERATE_DELTA) - -lemma ITERATE_IMAGE: "ALL u_4247::'q_50536::type => 'q_50536::type => 'q_50536::type. - monoidal u_4247 --> - (ALL (f::'q_50535::type => 'q_50507::type) - (g::'q_50507::type => 'q_50536::type) x::'q_50535::type => bool. - FINITE x & - (ALL (xa::'q_50535::type) y::'q_50535::type. - IN xa x & IN y x & f xa = f y --> xa = y) --> - iterate u_4247 (IMAGE f x) g = iterate u_4247 x (g o f))" - by (import hollight ITERATE_IMAGE) - -lemma ITERATE_BIJECTION: "ALL u_4247::'B::type => 'B::type => 'B::type. - monoidal u_4247 --> - (ALL (f::'A::type => 'B::type) (p::'A::type => 'A::type) - s::'A::type => bool. - FINITE s & - (ALL x::'A::type. IN x s --> IN (p x) s) & - (ALL y::'A::type. IN y s --> (EX! x::'A::type. IN x s & p x = y)) --> - iterate u_4247 s f = iterate u_4247 s (f o p))" - by (import hollight ITERATE_BIJECTION) - -lemma ITERATE_ITERATE_PRODUCT: "ALL u_4247::'C::type => 'C::type => 'C::type. - monoidal u_4247 --> - (ALL (x::'A::type => bool) (xa::'A::type => 'B::type => bool) - xb::'A::type => 'B::type => 'C::type. - FINITE x & (ALL i::'A::type. IN i x --> FINITE (xa i)) --> - iterate u_4247 x (%i::'A::type. iterate u_4247 (xa i) (xb i)) = - iterate u_4247 - (GSPEC - (%u::'A::type * 'B::type. - EX (i::'A::type) j::'B::type. - SETSPEC u (IN i x & IN j (xa i)) (i, j))) - (GABS - (%f::'A::type * 'B::type => 'C::type. - ALL (i::'A::type) j::'B::type. GEQ (f (i, j)) (xb i j))))" - by (import hollight ITERATE_ITERATE_PRODUCT) - -lemma ITERATE_EQ: "ALL u_4247::'q_50867::type => 'q_50867::type => 'q_50867::type. - monoidal u_4247 --> - (ALL (f::'q_50871::type => 'q_50867::type) - (g::'q_50871::type => 'q_50867::type) x::'q_50871::type => bool. - FINITE x & (ALL xa::'q_50871::type. IN xa x --> f xa = g xa) --> - iterate u_4247 x f = iterate u_4247 x g)" - by (import hollight ITERATE_EQ) - -lemma ITERATE_EQ_GENERAL: "ALL u_4247::'C::type => 'C::type => 'C::type. - monoidal u_4247 --> - (ALL (s::'A::type => bool) (t::'B::type => bool) - (f::'A::type => 'C::type) (g::'B::type => 'C::type) - h::'A::type => 'B::type. - FINITE s & - (ALL y::'B::type. IN y t --> (EX! x::'A::type. IN x s & h x = y)) & - (ALL x::'A::type. IN x s --> IN (h x) t & g (h x) = f x) --> - iterate u_4247 s f = iterate u_4247 t g)" - by (import hollight ITERATE_EQ_GENERAL) - -definition nsum :: "('q_51017 => bool) => ('q_51017 => nat) => nat" where - "(op ==::(('q_51017::type => bool) => ('q_51017::type => nat) => nat) - => (('q_51017::type => bool) => ('q_51017::type => nat) => nat) - => prop) - (nsum::('q_51017::type => bool) => ('q_51017::type => nat) => nat) - ((iterate::(nat => nat => nat) - => ('q_51017::type => bool) => ('q_51017::type => nat) => nat) - (plus::nat => nat => nat))" - -lemma DEF_nsum: "(op =::(('q_51017::type => bool) => ('q_51017::type => nat) => nat) - => (('q_51017::type => bool) => ('q_51017::type => nat) => nat) - => bool) - (nsum::('q_51017::type => bool) => ('q_51017::type => nat) => nat) - ((iterate::(nat => nat => nat) - => ('q_51017::type => bool) => ('q_51017::type => nat) => nat) - (plus::nat => nat => nat))" - by (import hollight DEF_nsum) - -lemma NEUTRAL_ADD: "(op =::nat => nat => bool) - ((neutral::(nat => nat => nat) => nat) (plus::nat => nat => nat)) (0::nat)" - by (import hollight NEUTRAL_ADD) - -lemma NEUTRAL_MUL: "neutral op * = NUMERAL_BIT1 0" - by (import hollight NEUTRAL_MUL) - -lemma MONOIDAL_ADD: "(monoidal::(nat => nat => nat) => bool) (plus::nat => nat => nat)" - by (import hollight MONOIDAL_ADD) - -lemma MONOIDAL_MUL: "(monoidal::(nat => nat => nat) => bool) (op *::nat => nat => nat)" - by (import hollight MONOIDAL_MUL) - -lemma NSUM_CLAUSES: "(ALL x::'q_51055::type => nat. nsum EMPTY x = 0) & -(ALL (x::'q_51094::type) (xa::'q_51094::type => nat) - xb::'q_51094::type => bool. - FINITE xb --> - nsum (INSERT x xb) xa = COND (IN x xb) (nsum xb xa) (xa x + nsum xb xa))" - by (import hollight NSUM_CLAUSES) - -lemma NSUM_UNION: "ALL (x::'q_51120::type => nat) (xa::'q_51120::type => bool) - xb::'q_51120::type => bool. - FINITE xa & FINITE xb & DISJOINT xa xb --> - nsum (hollight.UNION xa xb) x = nsum xa x + nsum xb x" - by (import hollight NSUM_UNION) - -lemma NSUM_DIFF: "ALL (f::'q_51175::type => nat) (s::'q_51175::type => bool) - t::'q_51175::type => bool. - FINITE s & SUBSET t s --> nsum (DIFF s t) f = nsum s f - nsum t f" - by (import hollight NSUM_DIFF) - -lemma NSUM_SUPPORT: "ALL (x::'q_51214::type => nat) xa::'q_51214::type => bool. - FINITE (support plus x xa) --> nsum (support plus x xa) x = nsum xa x" - by (import hollight NSUM_SUPPORT) - -lemma NSUM_ADD: "ALL (f::'q_51260::type => nat) (g::'q_51260::type => nat) - s::'q_51260::type => bool. - FINITE s --> nsum s (%x::'q_51260::type. f x + g x) = nsum s f + nsum s g" - by (import hollight NSUM_ADD) - -lemma NSUM_CMUL: "ALL (f::'q_51298::type => nat) (c::nat) s::'q_51298::type => bool. - FINITE s --> nsum s (%x::'q_51298::type. c * f x) = c * nsum s f" - by (import hollight NSUM_CMUL) - -lemma NSUM_LE: "ALL (x::'q_51337::type => nat) (xa::'q_51337::type => nat) - xb::'q_51337::type => bool. - FINITE xb & (ALL xc::'q_51337::type. IN xc xb --> <= (x xc) (xa xc)) --> - <= (nsum xb x) (nsum xb xa)" - by (import hollight NSUM_LE) - -lemma NSUM_LT: "ALL (f::'A::type => nat) (g::'A::type => nat) s::'A::type => bool. - FINITE s & - (ALL x::'A::type. IN x s --> <= (f x) (g x)) & - (EX x::'A::type. IN x s & < (f x) (g x)) --> - < (nsum s f) (nsum s g)" - by (import hollight NSUM_LT) - -lemma NSUM_LT_ALL: "ALL (f::'q_51459::type => nat) (g::'q_51459::type => nat) - s::'q_51459::type => bool. - FINITE s & - s ~= EMPTY & (ALL x::'q_51459::type. IN x s --> < (f x) (g x)) --> - < (nsum s f) (nsum s g)" - by (import hollight NSUM_LT_ALL) - -lemma NSUM_EQ: "ALL (x::'q_51501::type => nat) (xa::'q_51501::type => nat) - xb::'q_51501::type => bool. - FINITE xb & (ALL xc::'q_51501::type. IN xc xb --> x xc = xa xc) --> - nsum xb x = nsum xb xa" - by (import hollight NSUM_EQ) - -lemma NSUM_CONST: "ALL (c::nat) s::'q_51531::type => bool. - FINITE s --> nsum s (%n::'q_51531::type. c) = CARD s * c" - by (import hollight NSUM_CONST) - -lemma NSUM_EQ_0: "ALL (x::'A::type => nat) xa::'A::type => bool. - (ALL xb::'A::type. IN xb xa --> x xb = 0) --> nsum xa x = 0" - by (import hollight NSUM_EQ_0) - -lemma NSUM_0: "ALL x::'A::type => bool. nsum x (%n::'A::type. 0) = 0" - by (import hollight NSUM_0) - -lemma NSUM_POS_LE: "ALL (x::'q_51610::type => nat) xa::'q_51610::type => bool. - FINITE xa & (ALL xb::'q_51610::type. IN xb xa --> <= 0 (x xb)) --> - <= 0 (nsum xa x)" - by (import hollight NSUM_POS_LE) - -lemma NSUM_POS_BOUND: "ALL (f::'A::type => nat) (b::nat) x::'A::type => bool. - FINITE x & - (ALL xa::'A::type. IN xa x --> <= 0 (f xa)) & <= (nsum x f) b --> - (ALL xa::'A::type. IN xa x --> <= (f xa) b)" - by (import hollight NSUM_POS_BOUND) - -lemma NSUM_POS_EQ_0: "ALL (x::'q_51745::type => nat) xa::'q_51745::type => bool. - FINITE xa & - (ALL xb::'q_51745::type. IN xb xa --> <= 0 (x xb)) & nsum xa x = 0 --> - (ALL xb::'q_51745::type. IN xb xa --> x xb = 0)" - by (import hollight NSUM_POS_EQ_0) - -lemma NSUM_SING: "ALL (x::'q_51765::type => nat) xa::'q_51765::type. - nsum (INSERT xa EMPTY) x = x xa" - by (import hollight NSUM_SING) - -lemma NSUM_DELTA: "ALL (x::'A::type => bool) xa::'A::type. - nsum x (%x::'A::type. COND (x = xa) (b::nat) 0) = COND (IN xa x) b 0" - by (import hollight NSUM_DELTA) - -lemma NSUM_SWAP: "ALL (f::'A::type => 'B::type => nat) (x::'A::type => bool) - xa::'B::type => bool. - FINITE x & FINITE xa --> - nsum x (%i::'A::type. nsum xa (f i)) = - nsum xa (%j::'B::type. nsum x (%i::'A::type. f i j))" - by (import hollight NSUM_SWAP) - -lemma NSUM_IMAGE: "ALL (x::'q_51905::type => 'q_51881::type) (xa::'q_51881::type => nat) - xb::'q_51905::type => bool. - FINITE xb & - (ALL (xa::'q_51905::type) y::'q_51905::type. - IN xa xb & IN y xb & x xa = x y --> xa = y) --> - nsum (IMAGE x xb) xa = nsum xb (xa o x)" - by (import hollight NSUM_IMAGE) - -lemma NSUM_SUPERSET: "ALL (f::'A::type => nat) (u::'A::type => bool) v::'A::type => bool. - FINITE u & - SUBSET u v & (ALL x::'A::type. IN x v & ~ IN x u --> f x = 0) --> - nsum v f = nsum u f" - by (import hollight NSUM_SUPERSET) - -lemma NSUM_UNION_RZERO: "ALL (f::'A::type => nat) (u::'A::type => bool) v::'A::type => bool. - FINITE u & (ALL x::'A::type. IN x v & ~ IN x u --> f x = 0) --> - nsum (hollight.UNION u v) f = nsum u f" - by (import hollight NSUM_UNION_RZERO) - -lemma NSUM_UNION_LZERO: "ALL (f::'A::type => nat) (u::'A::type => bool) v::'A::type => bool. - FINITE v & (ALL x::'A::type. IN x u & ~ IN x v --> f x = 0) --> - nsum (hollight.UNION u v) f = nsum v f" - by (import hollight NSUM_UNION_LZERO) - -lemma NSUM_RESTRICT: "ALL (f::'q_52126::type => nat) s::'q_52126::type => bool. - FINITE s --> - nsum s (%x::'q_52126::type. COND (IN x s) (f x) 0) = nsum s f" - by (import hollight NSUM_RESTRICT) - -lemma NSUM_BOUND: "ALL (x::'A::type => bool) (xa::'A::type => nat) xb::nat. - FINITE x & (ALL xc::'A::type. IN xc x --> <= (xa xc) xb) --> - <= (nsum x xa) (CARD x * xb)" - by (import hollight NSUM_BOUND) - -lemma NSUM_BOUND_GEN: "ALL (x::'A::type => bool) (xa::'q_52201::type) b::nat. - FINITE x & - x ~= EMPTY & - (ALL xa::'A::type. - IN xa x --> <= ((f::'A::type => nat) xa) (DIV b (CARD x))) --> - <= (nsum x f) b" - by (import hollight NSUM_BOUND_GEN) - -lemma NSUM_BOUND_LT: "ALL (s::'A::type => bool) (f::'A::type => nat) b::nat. - FINITE s & - (ALL x::'A::type. IN x s --> <= (f x) b) & - (EX x::'A::type. IN x s & < (f x) b) --> - < (nsum s f) (CARD s * b)" - by (import hollight NSUM_BOUND_LT) - -lemma NSUM_BOUND_LT_ALL: "ALL (s::'q_52344::type => bool) (f::'q_52344::type => nat) b::nat. - FINITE s & s ~= EMPTY & (ALL x::'q_52344::type. IN x s --> < (f x) b) --> - < (nsum s f) (CARD s * b)" - by (import hollight NSUM_BOUND_LT_ALL) - -lemma NSUM_BOUND_LT_GEN: "ALL (s::'A::type => bool) (t::'q_52386::type) b::nat. - FINITE s & - s ~= EMPTY & - (ALL x::'A::type. - IN x s --> < ((f::'A::type => nat) x) (DIV b (CARD s))) --> - < (nsum s f) b" - by (import hollight NSUM_BOUND_LT_GEN) - -lemma NSUM_UNION_EQ: "ALL (s::'q_52445::type => bool) (t::'q_52445::type => bool) - u::'q_52445::type => bool. - FINITE u & hollight.INTER s t = EMPTY & hollight.UNION s t = u --> - nsum s (f::'q_52445::type => nat) + nsum t f = nsum u f" - by (import hollight NSUM_UNION_EQ) - -lemma NSUM_EQ_SUPERSET: "ALL (f::'A::type => nat) (s::'A::type => bool) t::'A::type => bool. - FINITE t & - SUBSET t s & - (ALL x::'A::type. IN x t --> f x = (g::'A::type => nat) x) & - (ALL x::'A::type. IN x s & ~ IN x t --> f x = 0) --> - nsum s f = nsum t g" - by (import hollight NSUM_EQ_SUPERSET) - -lemma NSUM_RESTRICT_SET: "ALL (s::'A::type => bool) (f::'A::type => nat) r::'q_52556::type. - FINITE s --> - nsum - (GSPEC - (%u::'A::type. - EX x::'A::type. SETSPEC u (IN x s & (P::'A::type => bool) x) x)) - f = - nsum s (%x::'A::type. COND (P x) (f x) 0)" - by (import hollight NSUM_RESTRICT_SET) - -lemma NSUM_NSUM_RESTRICT: "ALL (R::'q_52685::type => 'q_52684::type => bool) - (f::'q_52685::type => 'q_52684::type => nat) (s::'q_52685::type => bool) - t::'q_52684::type => bool. - FINITE s & FINITE t --> - nsum s - (%x::'q_52685::type. - nsum - (GSPEC - (%u::'q_52684::type. - EX y::'q_52684::type. SETSPEC u (IN y t & R x y) y)) - (f x)) = - nsum t - (%y::'q_52684::type. - nsum - (GSPEC - (%u::'q_52685::type. - EX x::'q_52685::type. SETSPEC u (IN x s & R x y) x)) - (%x::'q_52685::type. f x y))" - by (import hollight NSUM_NSUM_RESTRICT) - -lemma CARD_EQ_NSUM: "ALL x::'q_52704::type => bool. - FINITE x --> CARD x = nsum x (%x::'q_52704::type. NUMERAL_BIT1 0)" - by (import hollight CARD_EQ_NSUM) - -lemma NSUM_MULTICOUNT_GEN: "ALL (R::'A::type => 'B::type => bool) (s::'A::type => bool) - (t::'B::type => bool) k::'B::type => nat. - FINITE s & - FINITE t & - (ALL j::'B::type. - IN j t --> - CARD - (GSPEC - (%u::'A::type. EX i::'A::type. SETSPEC u (IN i s & R i j) i)) = - k j) --> - nsum s - (%i::'A::type. - CARD - (GSPEC - (%u::'B::type. EX j::'B::type. SETSPEC u (IN j t & R i j) j))) = - nsum t k" - by (import hollight NSUM_MULTICOUNT_GEN) - -lemma NSUM_MULTICOUNT: "ALL (R::'A::type => 'B::type => bool) (s::'A::type => bool) - (t::'B::type => bool) k::nat. - FINITE s & - FINITE t & - (ALL j::'B::type. - IN j t --> - CARD - (GSPEC - (%u::'A::type. EX i::'A::type. SETSPEC u (IN i s & R i j) i)) = - k) --> - nsum s - (%i::'A::type. - CARD - (GSPEC - (%u::'B::type. EX j::'B::type. SETSPEC u (IN j t & R i j) j))) = - k * CARD t" - by (import hollight NSUM_MULTICOUNT) - -lemma NSUM_IMAGE_GEN: "ALL (f::'A::type => 'B::type) (g::'A::type => nat) s::'A::type => bool. - FINITE s --> - nsum s g = - nsum (IMAGE f s) - (%y::'B::type. - nsum - (GSPEC - (%u::'A::type. EX x::'A::type. SETSPEC u (IN x s & f x = y) x)) - g)" - by (import hollight NSUM_IMAGE_GEN) - -lemma NSUM_SUBSET: "ALL (u::'A::type => bool) (v::'A::type => bool) f::'A::type => nat. - FINITE u & FINITE v & (ALL x::'A::type. IN x (DIFF u v) --> f x = 0) --> - <= (nsum u f) (nsum v f)" - by (import hollight NSUM_SUBSET) - -lemma NSUM_SUBSET_SIMPLE: "ALL (u::'q_53164::type => bool) (v::'q_53164::type => bool) - f::'q_53164::type => nat. - FINITE v & SUBSET u v --> <= (nsum u f) (nsum v f)" - by (import hollight NSUM_SUBSET_SIMPLE) - -lemma NSUM_ADD_NUMSEG: "ALL (x::nat => nat) (xa::nat => nat) (xb::nat) xc::nat. - nsum (dotdot xb xc) (%i::nat. x i + xa i) = - nsum (dotdot xb xc) x + nsum (dotdot xb xc) xa" - by (import hollight NSUM_ADD_NUMSEG) - -lemma NSUM_CMUL_NUMSEG: "ALL (x::nat => nat) (xa::nat) (xb::nat) xc::nat. - nsum (dotdot xb xc) (%i::nat. xa * x i) = xa * nsum (dotdot xb xc) x" - by (import hollight NSUM_CMUL_NUMSEG) - -lemma NSUM_LE_NUMSEG: "ALL (x::nat => nat) (xa::nat => nat) (xb::nat) xc::nat. - (ALL i::nat. <= xb i & <= i xc --> <= (x i) (xa i)) --> - <= (nsum (dotdot xb xc) x) (nsum (dotdot xb xc) xa)" - by (import hollight NSUM_LE_NUMSEG) - -lemma NSUM_EQ_NUMSEG: "ALL (f::nat => nat) (g::nat => nat) (m::nat) n::nat. - (ALL i::nat. <= m i & <= i n --> f i = g i) --> - nsum (dotdot m n) f = nsum (dotdot m n) g" - by (import hollight NSUM_EQ_NUMSEG) - -lemma NSUM_CONST_NUMSEG: "ALL (x::nat) (xa::nat) xb::nat. - nsum (dotdot xa xb) (%n::nat. x) = (xb + NUMERAL_BIT1 0 - xa) * x" - by (import hollight NSUM_CONST_NUMSEG) - -lemma NSUM_EQ_0_NUMSEG: "ALL (x::nat => nat) xa::'q_53403::type. - (ALL i::nat. <= (m::nat) i & <= i (n::nat) --> x i = 0) --> - nsum (dotdot m n) x = 0" - by (import hollight NSUM_EQ_0_NUMSEG) - -lemma NSUM_TRIV_NUMSEG: "ALL (f::nat => nat) (m::nat) n::nat. < n m --> nsum (dotdot m n) f = 0" - by (import hollight NSUM_TRIV_NUMSEG) - -lemma NSUM_POS_LE_NUMSEG: "ALL (x::nat) (xa::nat) xb::nat => nat. - (ALL p::nat. <= x p & <= p xa --> <= 0 (xb p)) --> - <= 0 (nsum (dotdot x xa) xb)" - by (import hollight NSUM_POS_LE_NUMSEG) - -lemma NSUM_POS_EQ_0_NUMSEG: "ALL (f::nat => nat) (m::nat) n::nat. - (ALL p::nat. <= m p & <= p n --> <= 0 (f p)) & - nsum (dotdot m n) f = 0 --> - (ALL p::nat. <= m p & <= p n --> f p = 0)" - by (import hollight NSUM_POS_EQ_0_NUMSEG) - -lemma NSUM_SING_NUMSEG: "ALL (x::nat => nat) xa::nat. nsum (dotdot xa xa) x = x xa" - by (import hollight NSUM_SING_NUMSEG) - -lemma NSUM_CLAUSES_NUMSEG: "(ALL x::nat. nsum (dotdot x 0) (f::nat => nat) = COND (x = 0) (f 0) 0) & -(ALL (x::nat) xa::nat. - nsum (dotdot x (Suc xa)) f = - COND (<= x (Suc xa)) (nsum (dotdot x xa) f + f (Suc xa)) - (nsum (dotdot x xa) f))" - by (import hollight NSUM_CLAUSES_NUMSEG) - -lemma NSUM_SWAP_NUMSEG: "ALL (a::nat) (b::nat) (c::nat) (d::nat) f::nat => nat => nat. - nsum (dotdot a b) (%i::nat. nsum (dotdot c d) (f i)) = - nsum (dotdot c d) (%j::nat. nsum (dotdot a b) (%i::nat. f i j))" - by (import hollight NSUM_SWAP_NUMSEG) - -lemma NSUM_ADD_SPLIT: "ALL (x::nat => nat) (xa::nat) (xb::nat) xc::nat. - <= xa (xb + NUMERAL_BIT1 0) --> - nsum (dotdot xa (xb + xc)) x = - nsum (dotdot xa xb) x + nsum (dotdot (xb + NUMERAL_BIT1 0) (xb + xc)) x" - by (import hollight NSUM_ADD_SPLIT) - -lemma NSUM_OFFSET: "ALL (x::nat => nat) (xa::nat) xb::nat. - nsum (dotdot (xa + xb) ((n::nat) + xb)) x = - nsum (dotdot xa n) (%i::nat. x (i + xb))" - by (import hollight NSUM_OFFSET) - -lemma NSUM_OFFSET_0: "ALL (x::nat => nat) (xa::nat) xb::nat. - <= xa xb --> - nsum (dotdot xa xb) x = nsum (dotdot 0 (xb - xa)) (%i::nat. x (i + xa))" - by (import hollight NSUM_OFFSET_0) - -lemma NSUM_CLAUSES_LEFT: "ALL (x::nat => nat) (xa::nat) xb::nat. - <= xa xb --> - nsum (dotdot xa xb) x = x xa + nsum (dotdot (xa + NUMERAL_BIT1 0) xb) x" - by (import hollight NSUM_CLAUSES_LEFT) - -lemma NSUM_CLAUSES_RIGHT: "ALL (f::nat => nat) (m::nat) n::nat. - < 0 n & <= m n --> - nsum (dotdot m n) f = nsum (dotdot m (n - NUMERAL_BIT1 0)) f + f n" - by (import hollight NSUM_CLAUSES_RIGHT) - -lemma NSUM_BIJECTION: "ALL (x::'A::type => nat) (xa::'A::type => 'A::type) xb::'A::type => bool. - FINITE xb & - (ALL x::'A::type. IN x xb --> IN (xa x) xb) & - (ALL y::'A::type. IN y xb --> (EX! x::'A::type. IN x xb & xa x = y)) --> - nsum xb x = nsum xb (x o xa)" - by (import hollight NSUM_BIJECTION) - -lemma NSUM_NSUM_PRODUCT: "ALL (x::'A::type => bool) (xa::'A::type => 'B::type => bool) - xb::'A::type => 'B::type => nat. - FINITE x & (ALL i::'A::type. IN i x --> FINITE (xa i)) --> - nsum x (%x::'A::type. nsum (xa x) (xb x)) = - nsum - (GSPEC - (%u::'A::type * 'B::type. - EX (i::'A::type) j::'B::type. - SETSPEC u (IN i x & IN j (xa i)) (i, j))) - (GABS - (%f::'A::type * 'B::type => nat. - ALL (i::'A::type) j::'B::type. GEQ (f (i, j)) (xb i j)))" - by (import hollight NSUM_NSUM_PRODUCT) - -lemma NSUM_EQ_GENERAL: "ALL (x::'A::type => bool) (xa::'B::type => bool) (xb::'A::type => nat) - (xc::'B::type => nat) xd::'A::type => 'B::type. - FINITE x & - (ALL y::'B::type. IN y xa --> (EX! xa::'A::type. IN xa x & xd xa = y)) & - (ALL xe::'A::type. IN xe x --> IN (xd xe) xa & xc (xd xe) = xb xe) --> - nsum x xb = nsum xa xc" - by (import hollight NSUM_EQ_GENERAL) - -consts - sum :: "('q_54215 => bool) => ('q_54215 => hollight.real) => hollight.real" - -defs - sum_def: "(op ==::(('q_54215::type => bool) - => ('q_54215::type => hollight.real) => hollight.real) - => (('q_54215::type => bool) - => ('q_54215::type => hollight.real) => hollight.real) - => prop) - (hollight.sum::('q_54215::type => bool) - => ('q_54215::type => hollight.real) => hollight.real) - ((iterate::(hollight.real => hollight.real => hollight.real) - => ('q_54215::type => bool) - => ('q_54215::type => hollight.real) => hollight.real) - (real_add::hollight.real => hollight.real => hollight.real))" - -lemma DEF_sum: "(op =::(('q_54215::type => bool) - => ('q_54215::type => hollight.real) => hollight.real) - => (('q_54215::type => bool) - => ('q_54215::type => hollight.real) => hollight.real) - => bool) - (hollight.sum::('q_54215::type => bool) - => ('q_54215::type => hollight.real) => hollight.real) - ((iterate::(hollight.real => hollight.real => hollight.real) - => ('q_54215::type => bool) - => ('q_54215::type => hollight.real) => hollight.real) - (real_add::hollight.real => hollight.real => hollight.real))" - by (import hollight DEF_sum) - -lemma NEUTRAL_REAL_ADD: "neutral real_add = real_of_num 0" - by (import hollight NEUTRAL_REAL_ADD) - -lemma NEUTRAL_REAL_MUL: "neutral real_mul = real_of_num (NUMERAL_BIT1 0)" - by (import hollight NEUTRAL_REAL_MUL) - -lemma MONOIDAL_REAL_ADD: "monoidal real_add" - by (import hollight MONOIDAL_REAL_ADD) - -lemma MONOIDAL_REAL_MUL: "monoidal real_mul" - by (import hollight MONOIDAL_REAL_MUL) - -lemma SUM_CLAUSES: "(ALL x::'q_54257::type => hollight.real. - hollight.sum EMPTY x = real_of_num 0) & -(ALL (x::'q_54298::type) (xa::'q_54298::type => hollight.real) - xb::'q_54298::type => bool. - FINITE xb --> - hollight.sum (INSERT x xb) xa = - COND (IN x xb) (hollight.sum xb xa) - (real_add (xa x) (hollight.sum xb xa)))" - by (import hollight SUM_CLAUSES) - -lemma SUM_UNION: "ALL (x::'q_54324::type => hollight.real) (xa::'q_54324::type => bool) - xb::'q_54324::type => bool. - FINITE xa & FINITE xb & DISJOINT xa xb --> - hollight.sum (hollight.UNION xa xb) x = - real_add (hollight.sum xa x) (hollight.sum xb x)" - by (import hollight SUM_UNION) - -lemma SUM_DIFF: "ALL (x::'q_54364::type => hollight.real) (xa::'q_54364::type => bool) - xb::'q_54364::type => bool. - FINITE xa & SUBSET xb xa --> - hollight.sum (DIFF xa xb) x = - real_sub (hollight.sum xa x) (hollight.sum xb x)" - by (import hollight SUM_DIFF) - -lemma SUM_SUPPORT: "ALL (x::'q_54403::type => hollight.real) xa::'q_54403::type => bool. - FINITE (support real_add x xa) --> - hollight.sum (support real_add x xa) x = hollight.sum xa x" - by (import hollight SUM_SUPPORT) - -lemma SUM_ADD: "ALL (f::'q_54449::type => hollight.real) - (g::'q_54449::type => hollight.real) s::'q_54449::type => bool. - FINITE s --> - hollight.sum s (%x::'q_54449::type. real_add (f x) (g x)) = - real_add (hollight.sum s f) (hollight.sum s g)" - by (import hollight SUM_ADD) - -lemma SUM_CMUL: "ALL (f::'q_54487::type => hollight.real) (c::hollight.real) - s::'q_54487::type => bool. - FINITE s --> - hollight.sum s (%x::'q_54487::type. real_mul c (f x)) = - real_mul c (hollight.sum s f)" - by (import hollight SUM_CMUL) - -lemma SUM_NEG: "ALL (x::'q_54530::type => hollight.real) xa::'q_54530::type => bool. - FINITE xa --> - hollight.sum xa (%xa::'q_54530::type. real_neg (x xa)) = - real_neg (hollight.sum xa x)" - by (import hollight SUM_NEG) - -lemma SUM_SUB: "ALL (x::'q_54565::type => hollight.real) - (xa::'q_54565::type => hollight.real) xb::'q_54565::type => bool. - FINITE xb --> - hollight.sum xb (%xb::'q_54565::type. real_sub (x xb) (xa xb)) = - real_sub (hollight.sum xb x) (hollight.sum xb xa)" - by (import hollight SUM_SUB) - -lemma SUM_LE: "ALL (x::'q_54607::type => hollight.real) - (xa::'q_54607::type => hollight.real) xb::'q_54607::type => bool. - FINITE xb & - (ALL xc::'q_54607::type. IN xc xb --> real_le (x xc) (xa xc)) --> - real_le (hollight.sum xb x) (hollight.sum xb xa)" - by (import hollight SUM_LE) - -lemma SUM_LT: "ALL (f::'A::type => hollight.real) (g::'A::type => hollight.real) - s::'A::type => bool. - FINITE s & - (ALL x::'A::type. IN x s --> real_le (f x) (g x)) & - (EX x::'A::type. IN x s & real_lt (f x) (g x)) --> - real_lt (hollight.sum s f) (hollight.sum s g)" - by (import hollight SUM_LT) - -lemma SUM_LT_ALL: "ALL (f::'q_54729::type => hollight.real) - (g::'q_54729::type => hollight.real) s::'q_54729::type => bool. - FINITE s & - s ~= EMPTY & (ALL x::'q_54729::type. IN x s --> real_lt (f x) (g x)) --> - real_lt (hollight.sum s f) (hollight.sum s g)" - by (import hollight SUM_LT_ALL) - -lemma SUM_EQ: "ALL (x::'q_54771::type => hollight.real) - (xa::'q_54771::type => hollight.real) xb::'q_54771::type => bool. - FINITE xb & (ALL xc::'q_54771::type. IN xc xb --> x xc = xa xc) --> - hollight.sum xb x = hollight.sum xb xa" - by (import hollight SUM_EQ) - -lemma SUM_ABS: "ALL (f::'q_54830::type => hollight.real) s::'q_54830::type => bool. - FINITE s --> - real_le (real_abs (hollight.sum s f)) - (hollight.sum s (%x::'q_54830::type. real_abs (f x)))" - by (import hollight SUM_ABS) - -lemma SUM_CONST: "ALL (c::hollight.real) s::'q_54851::type => bool. - FINITE s --> - hollight.sum s (%n::'q_54851::type. c) = - real_mul (real_of_num (CARD s)) c" - by (import hollight SUM_CONST) - -lemma SUM_EQ_0: "ALL (x::'A::type => hollight.real) xa::'A::type => bool. - (ALL xb::'A::type. IN xb xa --> x xb = real_of_num 0) --> - hollight.sum xa x = real_of_num 0" - by (import hollight SUM_EQ_0) - -lemma SUM_0: "ALL x::'A::type => bool. - hollight.sum x (%n::'A::type. real_of_num 0) = real_of_num 0" - by (import hollight SUM_0) - -lemma SUM_POS_LE: "ALL (x::'q_54944::type => hollight.real) xa::'q_54944::type => bool. - FINITE xa & - (ALL xb::'q_54944::type. IN xb xa --> real_le (real_of_num 0) (x xb)) --> - real_le (real_of_num 0) (hollight.sum xa x)" - by (import hollight SUM_POS_LE) - -lemma SUM_POS_BOUND: "ALL (f::'A::type => hollight.real) (b::hollight.real) x::'A::type => bool. - FINITE x & - (ALL xa::'A::type. IN xa x --> real_le (real_of_num 0) (f xa)) & - real_le (hollight.sum x f) b --> - (ALL xa::'A::type. IN xa x --> real_le (f xa) b)" - by (import hollight SUM_POS_BOUND) - -lemma SUM_POS_EQ_0: "ALL (x::'q_55091::type => hollight.real) xa::'q_55091::type => bool. - FINITE xa & - (ALL xb::'q_55091::type. IN xb xa --> real_le (real_of_num 0) (x xb)) & - hollight.sum xa x = real_of_num 0 --> - (ALL xb::'q_55091::type. IN xb xa --> x xb = real_of_num 0)" - by (import hollight SUM_POS_EQ_0) - -lemma SUM_SING: "ALL (x::'q_55113::type => hollight.real) xa::'q_55113::type. - hollight.sum (INSERT xa EMPTY) x = x xa" - by (import hollight SUM_SING) - -lemma SUM_DELTA: "ALL (x::'A::type => bool) xa::'A::type. - hollight.sum x - (%x::'A::type. COND (x = xa) (b::hollight.real) (real_of_num 0)) = - COND (IN xa x) b (real_of_num 0)" - by (import hollight SUM_DELTA) - -lemma SUM_SWAP: "ALL (f::'A::type => 'B::type => hollight.real) (x::'A::type => bool) - xa::'B::type => bool. - FINITE x & FINITE xa --> - hollight.sum x (%i::'A::type. hollight.sum xa (f i)) = - hollight.sum xa (%j::'B::type. hollight.sum x (%i::'A::type. f i j))" - by (import hollight SUM_SWAP) - -lemma SUM_IMAGE: "ALL (x::'q_55257::type => 'q_55233::type) - (xa::'q_55233::type => hollight.real) xb::'q_55257::type => bool. - FINITE xb & - (ALL (xa::'q_55257::type) y::'q_55257::type. - IN xa xb & IN y xb & x xa = x y --> xa = y) --> - hollight.sum (IMAGE x xb) xa = hollight.sum xb (xa o x)" - by (import hollight SUM_IMAGE) - -lemma SUM_SUPERSET: "ALL (f::'A::type => hollight.real) (u::'A::type => bool) - v::'A::type => bool. - FINITE u & - SUBSET u v & - (ALL x::'A::type. IN x v & ~ IN x u --> f x = real_of_num 0) --> - hollight.sum v f = hollight.sum u f" - by (import hollight SUM_SUPERSET) - -lemma SUM_UNION_RZERO: "ALL (f::'A::type => hollight.real) (u::'A::type => bool) - v::'A::type => bool. - FINITE u & - (ALL x::'A::type. IN x v & ~ IN x u --> f x = real_of_num 0) --> - hollight.sum (hollight.UNION u v) f = hollight.sum u f" - by (import hollight SUM_UNION_RZERO) - -lemma SUM_UNION_LZERO: "ALL (f::'A::type => hollight.real) (u::'A::type => bool) - v::'A::type => bool. - FINITE v & - (ALL x::'A::type. IN x u & ~ IN x v --> f x = real_of_num 0) --> - hollight.sum (hollight.UNION u v) f = hollight.sum v f" - by (import hollight SUM_UNION_LZERO) - -lemma SUM_RESTRICT: "ALL (f::'q_55484::type => hollight.real) s::'q_55484::type => bool. - FINITE s --> - hollight.sum s - (%x::'q_55484::type. COND (IN x s) (f x) (real_of_num 0)) = - hollight.sum s f" - by (import hollight SUM_RESTRICT) - -lemma SUM_BOUND: "ALL (x::'A::type => bool) (xa::'A::type => hollight.real) xb::hollight.real. - FINITE x & (ALL xc::'A::type. IN xc x --> real_le (xa xc) xb) --> - real_le (hollight.sum x xa) (real_mul (real_of_num (CARD x)) xb)" - by (import hollight SUM_BOUND) - -lemma SUM_BOUND_GEN: "ALL (s::'A::type => bool) (t::'q_55543::type) b::hollight.real. - FINITE s & - s ~= EMPTY & - (ALL x::'A::type. - IN x s --> - real_le ((f::'A::type => hollight.real) x) - (real_div b (real_of_num (CARD s)))) --> - real_le (hollight.sum s f) b" - by (import hollight SUM_BOUND_GEN) - -lemma SUM_ABS_BOUND: "ALL (s::'A::type => bool) (f::'A::type => hollight.real) b::hollight.real. - FINITE s & (ALL x::'A::type. IN x s --> real_le (real_abs (f x)) b) --> - real_le (real_abs (hollight.sum s f)) (real_mul (real_of_num (CARD s)) b)" - by (import hollight SUM_ABS_BOUND) - -lemma SUM_BOUND_LT: "ALL (s::'A::type => bool) (f::'A::type => hollight.real) b::hollight.real. - FINITE s & - (ALL x::'A::type. IN x s --> real_le (f x) b) & - (EX x::'A::type. IN x s & real_lt (f x) b) --> - real_lt (hollight.sum s f) (real_mul (real_of_num (CARD s)) b)" - by (import hollight SUM_BOUND_LT) - -lemma SUM_BOUND_LT_ALL: "ALL (s::'q_55748::type => bool) (f::'q_55748::type => hollight.real) - b::hollight.real. - FINITE s & - s ~= EMPTY & (ALL x::'q_55748::type. IN x s --> real_lt (f x) b) --> - real_lt (hollight.sum s f) (real_mul (real_of_num (CARD s)) b)" - by (import hollight SUM_BOUND_LT_ALL) - -lemma SUM_BOUND_LT_GEN: "ALL (s::'A::type => bool) (t::'q_55770::type) b::hollight.real. - FINITE s & - s ~= EMPTY & - (ALL x::'A::type. - IN x s --> - real_lt ((f::'A::type => hollight.real) x) - (real_div b (real_of_num (CARD s)))) --> - real_lt (hollight.sum s f) b" - by (import hollight SUM_BOUND_LT_GEN) - -lemma SUM_UNION_EQ: "ALL (s::'q_55831::type => bool) (t::'q_55831::type => bool) - u::'q_55831::type => bool. - FINITE u & hollight.INTER s t = EMPTY & hollight.UNION s t = u --> - real_add (hollight.sum s (f::'q_55831::type => hollight.real)) - (hollight.sum t f) = - hollight.sum u f" - by (import hollight SUM_UNION_EQ) - -lemma SUM_EQ_SUPERSET: "ALL (f::'A::type => hollight.real) (s::'A::type => bool) - t::'A::type => bool. - FINITE t & - SUBSET t s & - (ALL x::'A::type. IN x t --> f x = (g::'A::type => hollight.real) x) & - (ALL x::'A::type. IN x s & ~ IN x t --> f x = real_of_num 0) --> - hollight.sum s f = hollight.sum t g" - by (import hollight SUM_EQ_SUPERSET) - -lemma SUM_RESTRICT_SET: "ALL (s::'A::type => bool) (f::'A::type => hollight.real) r::'q_55944::type. - FINITE s --> - hollight.sum - (GSPEC - (%u::'A::type. - EX x::'A::type. SETSPEC u (IN x s & (P::'A::type => bool) x) x)) - f = - hollight.sum s (%x::'A::type. COND (P x) (f x) (real_of_num 0))" - by (import hollight SUM_RESTRICT_SET) - -lemma SUM_SUM_RESTRICT: "ALL (R::'q_56075::type => 'q_56074::type => bool) - (f::'q_56075::type => 'q_56074::type => hollight.real) - (s::'q_56075::type => bool) t::'q_56074::type => bool. - FINITE s & FINITE t --> - hollight.sum s - (%x::'q_56075::type. - hollight.sum - (GSPEC - (%u::'q_56074::type. - EX y::'q_56074::type. SETSPEC u (IN y t & R x y) y)) - (f x)) = - hollight.sum t - (%y::'q_56074::type. - hollight.sum - (GSPEC - (%u::'q_56075::type. - EX x::'q_56075::type. SETSPEC u (IN x s & R x y) x)) - (%x::'q_56075::type. f x y))" - by (import hollight SUM_SUM_RESTRICT) - -lemma CARD_EQ_SUM: "ALL x::'q_56096::type => bool. - FINITE x --> - real_of_num (CARD x) = - hollight.sum x (%x::'q_56096::type. real_of_num (NUMERAL_BIT1 0))" - by (import hollight CARD_EQ_SUM) - -lemma SUM_MULTICOUNT_GEN: "ALL (R::'A::type => 'B::type => bool) (s::'A::type => bool) - (t::'B::type => bool) k::'B::type => nat. - FINITE s & - FINITE t & - (ALL j::'B::type. - IN j t --> - CARD - (GSPEC - (%u::'A::type. EX i::'A::type. SETSPEC u (IN i s & R i j) i)) = - k j) --> - hollight.sum s - (%i::'A::type. - real_of_num - (CARD - (GSPEC - (%u::'B::type. - EX j::'B::type. SETSPEC u (IN j t & R i j) j)))) = - hollight.sum t (%i::'B::type. real_of_num (k i))" - by (import hollight SUM_MULTICOUNT_GEN) - -lemma SUM_MULTICOUNT: "ALL (R::'A::type => 'B::type => bool) (s::'A::type => bool) - (t::'B::type => bool) k::nat. - FINITE s & - FINITE t & - (ALL j::'B::type. - IN j t --> - CARD - (GSPEC - (%u::'A::type. EX i::'A::type. SETSPEC u (IN i s & R i j) i)) = - k) --> - hollight.sum s - (%i::'A::type. - real_of_num - (CARD - (GSPEC - (%u::'B::type. - EX j::'B::type. SETSPEC u (IN j t & R i j) j)))) = - real_of_num (k * CARD t)" - by (import hollight SUM_MULTICOUNT) - -lemma SUM_IMAGE_GEN: "ALL (f::'A::type => 'B::type) (g::'A::type => hollight.real) - s::'A::type => bool. - FINITE s --> - hollight.sum s g = - hollight.sum (IMAGE f s) - (%y::'B::type. - hollight.sum - (GSPEC - (%u::'A::type. EX x::'A::type. SETSPEC u (IN x s & f x = y) x)) - g)" - by (import hollight SUM_IMAGE_GEN) - -lemma REAL_OF_NUM_SUM: "ALL (f::'q_56493::type => nat) s::'q_56493::type => bool. - FINITE s --> - real_of_num (nsum s f) = - hollight.sum s (%x::'q_56493::type. real_of_num (f x))" - by (import hollight REAL_OF_NUM_SUM) - -lemma SUM_SUBSET: "ALL (u::'A::type => bool) (v::'A::type => bool) - f::'A::type => hollight.real. - FINITE u & - FINITE v & - (ALL x::'A::type. IN x (DIFF u v) --> real_le (f x) (real_of_num 0)) & - (ALL x::'A::type. IN x (DIFF v u) --> real_le (real_of_num 0) (f x)) --> - real_le (hollight.sum u f) (hollight.sum v f)" - by (import hollight SUM_SUBSET) - -lemma SUM_SUBSET_SIMPLE: "ALL (u::'A::type => bool) (v::'A::type => bool) - f::'A::type => hollight.real. - FINITE v & - SUBSET u v & - (ALL x::'A::type. IN x (DIFF v u) --> real_le (real_of_num 0) (f x)) --> - real_le (hollight.sum u f) (hollight.sum v f)" - by (import hollight SUM_SUBSET_SIMPLE) - -lemma SUM_ADD_NUMSEG: "ALL (x::nat => hollight.real) (xa::nat => hollight.real) (xb::nat) xc::nat. - hollight.sum (dotdot xb xc) (%i::nat. real_add (x i) (xa i)) = - real_add (hollight.sum (dotdot xb xc) x) (hollight.sum (dotdot xb xc) xa)" - by (import hollight SUM_ADD_NUMSEG) - -lemma SUM_CMUL_NUMSEG: "ALL (x::nat => hollight.real) (xa::hollight.real) (xb::nat) xc::nat. - hollight.sum (dotdot xb xc) (%i::nat. real_mul xa (x i)) = - real_mul xa (hollight.sum (dotdot xb xc) x)" - by (import hollight SUM_CMUL_NUMSEG) - -lemma SUM_NEG_NUMSEG: "ALL (x::nat => hollight.real) (xa::nat) xb::nat. - hollight.sum (dotdot xa xb) (%i::nat. real_neg (x i)) = - real_neg (hollight.sum (dotdot xa xb) x)" - by (import hollight SUM_NEG_NUMSEG) - -lemma SUM_SUB_NUMSEG: "ALL (x::nat => hollight.real) (xa::nat => hollight.real) (xb::nat) xc::nat. - hollight.sum (dotdot xb xc) (%i::nat. real_sub (x i) (xa i)) = - real_sub (hollight.sum (dotdot xb xc) x) (hollight.sum (dotdot xb xc) xa)" - by (import hollight SUM_SUB_NUMSEG) - -lemma SUM_LE_NUMSEG: "ALL (x::nat => hollight.real) (xa::nat => hollight.real) (xb::nat) xc::nat. - (ALL i::nat. <= xb i & <= i xc --> real_le (x i) (xa i)) --> - real_le (hollight.sum (dotdot xb xc) x) (hollight.sum (dotdot xb xc) xa)" - by (import hollight SUM_LE_NUMSEG) - -lemma SUM_EQ_NUMSEG: "ALL (f::nat => hollight.real) (g::nat => hollight.real) (m::nat) n::nat. - (ALL i::nat. <= m i & <= i n --> f i = g i) --> - hollight.sum (dotdot m n) f = hollight.sum (dotdot m n) g" - by (import hollight SUM_EQ_NUMSEG) - -lemma SUM_ABS_NUMSEG: "ALL (x::nat => hollight.real) (xa::nat) xb::nat. - real_le (real_abs (hollight.sum (dotdot xa xb) x)) - (hollight.sum (dotdot xa xb) (%i::nat. real_abs (x i)))" - by (import hollight SUM_ABS_NUMSEG) - -lemma SUM_CONST_NUMSEG: "ALL (x::hollight.real) (xa::nat) xb::nat. - hollight.sum (dotdot xa xb) (%n::nat. x) = - real_mul (real_of_num (xb + NUMERAL_BIT1 0 - xa)) x" - by (import hollight SUM_CONST_NUMSEG) - -lemma SUM_EQ_0_NUMSEG: "ALL (x::nat => hollight.real) xa::'q_57019::type. - (ALL i::nat. <= (m::nat) i & <= i (n::nat) --> x i = real_of_num 0) --> - hollight.sum (dotdot m n) x = real_of_num 0" - by (import hollight SUM_EQ_0_NUMSEG) - -lemma SUM_TRIV_NUMSEG: "ALL (f::nat => hollight.real) (m::nat) n::nat. - < n m --> hollight.sum (dotdot m n) f = real_of_num 0" - by (import hollight SUM_TRIV_NUMSEG) - -lemma SUM_POS_LE_NUMSEG: "ALL (x::nat) (xa::nat) xb::nat => hollight.real. - (ALL p::nat. <= x p & <= p xa --> real_le (real_of_num 0) (xb p)) --> - real_le (real_of_num 0) (hollight.sum (dotdot x xa) xb)" - by (import hollight SUM_POS_LE_NUMSEG) - -lemma SUM_POS_EQ_0_NUMSEG: "ALL (f::nat => hollight.real) (m::nat) n::nat. - (ALL p::nat. <= m p & <= p n --> real_le (real_of_num 0) (f p)) & - hollight.sum (dotdot m n) f = real_of_num 0 --> - (ALL p::nat. <= m p & <= p n --> f p = real_of_num 0)" - by (import hollight SUM_POS_EQ_0_NUMSEG) - -lemma SUM_SING_NUMSEG: "ALL (x::nat => hollight.real) xa::nat. hollight.sum (dotdot xa xa) x = x xa" - by (import hollight SUM_SING_NUMSEG) - -lemma SUM_CLAUSES_NUMSEG: "(ALL x::nat. - hollight.sum (dotdot x 0) (f::nat => hollight.real) = - COND (x = 0) (f 0) (real_of_num 0)) & -(ALL (x::nat) xa::nat. - hollight.sum (dotdot x (Suc xa)) f = - COND (<= x (Suc xa)) - (real_add (hollight.sum (dotdot x xa) f) (f (Suc xa))) - (hollight.sum (dotdot x xa) f))" - by (import hollight SUM_CLAUSES_NUMSEG) - -lemma SUM_SWAP_NUMSEG: "ALL (a::nat) (b::nat) (c::nat) (d::nat) f::nat => nat => hollight.real. - hollight.sum (dotdot a b) (%i::nat. hollight.sum (dotdot c d) (f i)) = - hollight.sum (dotdot c d) - (%j::nat. hollight.sum (dotdot a b) (%i::nat. f i j))" - by (import hollight SUM_SWAP_NUMSEG) - -lemma SUM_ADD_SPLIT: "ALL (x::nat => hollight.real) (xa::nat) (xb::nat) xc::nat. - <= xa (xb + NUMERAL_BIT1 0) --> - hollight.sum (dotdot xa (xb + xc)) x = - real_add (hollight.sum (dotdot xa xb) x) - (hollight.sum (dotdot (xb + NUMERAL_BIT1 0) (xb + xc)) x)" - by (import hollight SUM_ADD_SPLIT) - -lemma SUM_OFFSET: "ALL (x::nat => hollight.real) (xa::nat) xb::nat. - hollight.sum (dotdot (xa + xb) ((n::nat) + xb)) x = - hollight.sum (dotdot xa n) (%i::nat. x (i + xb))" - by (import hollight SUM_OFFSET) - -lemma SUM_OFFSET_0: "ALL (x::nat => hollight.real) (xa::nat) xb::nat. - <= xa xb --> - hollight.sum (dotdot xa xb) x = - hollight.sum (dotdot 0 (xb - xa)) (%i::nat. x (i + xa))" - by (import hollight SUM_OFFSET_0) - -lemma SUM_CLAUSES_LEFT: "ALL (x::nat => hollight.real) (xa::nat) xb::nat. - <= xa xb --> - hollight.sum (dotdot xa xb) x = - real_add (x xa) (hollight.sum (dotdot (xa + NUMERAL_BIT1 0) xb) x)" - by (import hollight SUM_CLAUSES_LEFT) - -lemma SUM_CLAUSES_RIGHT: "ALL (f::nat => hollight.real) (m::nat) n::nat. - < 0 n & <= m n --> - hollight.sum (dotdot m n) f = - real_add (hollight.sum (dotdot m (n - NUMERAL_BIT1 0)) f) (f n)" - by (import hollight SUM_CLAUSES_RIGHT) - -lemma REAL_OF_NUM_SUM_NUMSEG: "ALL (x::nat => nat) (xa::nat) xb::nat. - real_of_num (nsum (dotdot xa xb) x) = - hollight.sum (dotdot xa xb) (%i::nat. real_of_num (x i))" - by (import hollight REAL_OF_NUM_SUM_NUMSEG) - -lemma SUM_BIJECTION: "ALL (x::'A::type => hollight.real) (xa::'A::type => 'A::type) - xb::'A::type => bool. - FINITE xb & - (ALL x::'A::type. IN x xb --> IN (xa x) xb) & - (ALL y::'A::type. IN y xb --> (EX! x::'A::type. IN x xb & xa x = y)) --> - hollight.sum xb x = hollight.sum xb (x o xa)" - by (import hollight SUM_BIJECTION) - -lemma SUM_SUM_PRODUCT: "ALL (x::'A::type => bool) (xa::'A::type => 'B::type => bool) - xb::'A::type => 'B::type => hollight.real. - FINITE x & (ALL i::'A::type. IN i x --> FINITE (xa i)) --> - hollight.sum x (%x::'A::type. hollight.sum (xa x) (xb x)) = - hollight.sum - (GSPEC - (%u::'A::type * 'B::type. - EX (i::'A::type) j::'B::type. - SETSPEC u (IN i x & IN j (xa i)) (i, j))) - (GABS - (%f::'A::type * 'B::type => hollight.real. - ALL (i::'A::type) j::'B::type. GEQ (f (i, j)) (xb i j)))" - by (import hollight SUM_SUM_PRODUCT) - -lemma SUM_EQ_GENERAL: "ALL (x::'A::type => bool) (xa::'B::type => bool) - (xb::'A::type => hollight.real) (xc::'B::type => hollight.real) - xd::'A::type => 'B::type. - FINITE x & - (ALL y::'B::type. IN y xa --> (EX! xa::'A::type. IN xa x & xd xa = y)) & - (ALL xe::'A::type. IN xe x --> IN (xd xe) xa & xc (xd xe) = xb xe) --> - hollight.sum x xb = hollight.sum xa xc" - by (import hollight SUM_EQ_GENERAL) - -definition CASEWISE :: "(('q_57926 => 'q_57930) * ('q_57931 => 'q_57926 => 'q_57890)) hollight.list -=> 'q_57931 => 'q_57930 => 'q_57890" where +lemma FINITE_SUM_IMAGE: "(op =::(('A::type, 'B::type) finite_sum => bool) + => (('A::type, 'B::type) finite_sum => bool) => bool) + (UNIV::('A::type, 'B::type) finite_sum => bool) + ((op `::(nat => ('A::type, 'B::type) finite_sum) + => (nat => bool) => ('A::type, 'B::type) finite_sum => bool) + (mk_finite_sum::nat => ('A::type, 'B::type) finite_sum) + ((atLeastAtMost::nat => nat => nat => bool) (1::nat) + ((op +::nat => nat => nat) + ((dimindex::('A::type => bool) => nat) (UNIV::'A::type => bool)) + ((dimindex::('B::type => bool) => nat) (UNIV::'B::type => bool)))))" + by (import hollight FINITE_SUM_IMAGE) + +lemma HAS_SIZE_1: "(HAS_SIZE::(unit => bool) => nat => bool) (UNIV::unit => bool) (1::nat)" + by (import hollight HAS_SIZE_1) + +typedef (open) N_2 = "{x. x : dotdot (NUMERAL (NUMERAL_BIT1 0)) + (NUMERAL (NUMERAL_BIT0 (NUMERAL_BIT1 0)))}" morphisms "dest_auto_define_finite_type_2" "mk_auto_define_finite_type_2" + apply (rule light_ex_imp_nonempty[where t="SOME x. + x : dotdot (NUMERAL (NUMERAL_BIT1 0)) + (NUMERAL (NUMERAL_BIT0 (NUMERAL_BIT1 0)))"]) + by (import hollight TYDEF_2) + +syntax + dest_auto_define_finite_type_2 :: _ + +syntax + mk_auto_define_finite_type_2 :: _ + +lemmas "TYDEF_2_@intern" = typedef_hol2hollight + [where a="a :: N_2" and r=r , + OF type_definition_N_2] + +typedef (open) N_3 = "{x. x : dotdot (NUMERAL (NUMERAL_BIT1 0)) + (NUMERAL (NUMERAL_BIT1 (NUMERAL_BIT1 0)))}" morphisms "dest_auto_define_finite_type_3" "mk_auto_define_finite_type_3" + apply (rule light_ex_imp_nonempty[where t="SOME x. + x : dotdot (NUMERAL (NUMERAL_BIT1 0)) + (NUMERAL (NUMERAL_BIT1 (NUMERAL_BIT1 0)))"]) + by (import hollight TYDEF_3) + +syntax + dest_auto_define_finite_type_3 :: _ + +syntax + mk_auto_define_finite_type_3 :: _ + +lemmas "TYDEF_3_@intern" = typedef_hol2hollight + [where a="a :: N_3" and r=r , + OF type_definition_N_3] + +lemma FINITE_CART: "(op ==>::prop => prop => prop) + ((all::(nat => prop) => prop) + (%i::nat. + (op ==>::prop => prop => prop) + ((Trueprop::bool => prop) + ((op &::bool => bool => bool) + ((op <=::nat => nat => bool) (1::nat) i) + ((op <=::nat => nat => bool) i + ((dimindex::('N::type => bool) => nat) + (UNIV::'N::type => bool))))) + ((Trueprop::bool => prop) + ((finite::('A::type => bool) => bool) + ((Collect::('A::type => bool) => 'A::type => bool) + (%u::'A::type. + (Ex::('A::type => bool) => bool) + (%x::'A::type. + (op &::bool => bool => bool) + ((P::nat => 'A::type => bool) i x) + ((op =::'A::type => 'A::type => bool) u x)))))))) + ((Trueprop::bool => prop) + ((finite::(('A::type, 'N::type) cart => bool) => bool) + ((Collect::(('A::type, 'N::type) cart => bool) + => ('A::type, 'N::type) cart => bool) + (%u::('A::type, 'N::type) cart. + (Ex::(('A::type, 'N::type) cart => bool) => bool) + (%v::('A::type, 'N::type) cart. + (op &::bool => bool => bool) + ((All::(nat => bool) => bool) + (%i::nat. + (op -->::bool => bool => bool) + ((op &::bool => bool => bool) + ((op <=::nat => nat => bool) (1::nat) i) + ((op <=::nat => nat => bool) i + ((dimindex::('N::type => bool) => nat) + (UNIV::'N::type => bool)))) + (P i (($::('A::type, 'N::type) cart + => nat => 'A::type) + v i)))) + ((op =::('A::type, 'N::type) cart + => ('A::type, 'N::type) cart => bool) + u v))))))" + by (import hollight FINITE_CART) + +definition + vector :: "'A list => ('A, 'N) cart" where + "(op ==::('A::type list => ('A::type, 'N::type) cart) + => ('A::type list => ('A::type, 'N::type) cart) => prop) + (vector::'A::type list => ('A::type, 'N::type) cart) + (%u::'A::type list. + (lambda::(nat => 'A::type) => ('A::type, 'N::type) cart) + (%i::nat. + (op !::'A::type list => nat => 'A::type) u + ((op -::nat => nat => nat) i (1::nat))))" + +lemma DEF_vector: "(op =::('A::type list => ('A::type, 'N::type) cart) + => ('A::type list => ('A::type, 'N::type) cart) => bool) + (vector::'A::type list => ('A::type, 'N::type) cart) + (%u::'A::type list. + (lambda::(nat => 'A::type) => ('A::type, 'N::type) cart) + (%i::nat. + (op !::'A::type list => nat => 'A::type) u + ((op -::nat => nat => nat) i (1::nat))))" + by (import hollight DEF_vector) + +definition + CASEWISE :: "(('q_74835 => 'q_74839) * ('q_74840 => 'q_74835 => 'q_74799)) list +=> 'q_74840 => 'q_74839 => 'q_74799" where "CASEWISE == -SOME CASEWISE::(('q_57926::type => 'q_57930::type) * - ('q_57931::type - => 'q_57926::type => 'q_57890::type)) hollight.list - => 'q_57931::type => 'q_57930::type => 'q_57890::type. - (ALL (f::'q_57931::type) x::'q_57930::type. - CASEWISE NIL f x = (SOME y::'q_57890::type. True)) & - (ALL (h::('q_57926::type => 'q_57930::type) * - ('q_57931::type => 'q_57926::type => 'q_57890::type)) - (t::(('q_57926::type => 'q_57930::type) * - ('q_57931::type - => 'q_57926::type => 'q_57890::type)) hollight.list) - (f::'q_57931::type) x::'q_57930::type. - CASEWISE (CONS h t) f x = - COND (EX y::'q_57926::type. fst h y = x) - (snd h f (SOME y::'q_57926::type. fst h y = x)) (CASEWISE t f x))" +SOME CASEWISE::(('q_74835 => 'q_74839) * + ('q_74840 => 'q_74835 => 'q_74799)) list + => 'q_74840 => 'q_74839 => 'q_74799. + (ALL (f::'q_74840) x::'q_74839. + CASEWISE [] f x = (SOME y::'q_74799. True)) & + (ALL (h::('q_74835 => 'q_74839) * ('q_74840 => 'q_74835 => 'q_74799)) + (t::(('q_74835 => 'q_74839) * + ('q_74840 => 'q_74835 => 'q_74799)) list) + (f::'q_74840) x::'q_74839. + CASEWISE (h # t) f x = + (if EX y::'q_74835. fst h y = x + then snd h f (SOME y::'q_74835. fst h y = x) else CASEWISE t f x))" lemma DEF_CASEWISE: "CASEWISE = -(SOME CASEWISE::(('q_57926::type => 'q_57930::type) * - ('q_57931::type - => 'q_57926::type => 'q_57890::type)) hollight.list - => 'q_57931::type => 'q_57930::type => 'q_57890::type. - (ALL (f::'q_57931::type) x::'q_57930::type. - CASEWISE NIL f x = (SOME y::'q_57890::type. True)) & - (ALL (h::('q_57926::type => 'q_57930::type) * - ('q_57931::type => 'q_57926::type => 'q_57890::type)) - (t::(('q_57926::type => 'q_57930::type) * - ('q_57931::type - => 'q_57926::type => 'q_57890::type)) hollight.list) - (f::'q_57931::type) x::'q_57930::type. - CASEWISE (CONS h t) f x = - COND (EX y::'q_57926::type. fst h y = x) - (snd h f (SOME y::'q_57926::type. fst h y = x)) (CASEWISE t f x)))" +(SOME CASEWISE::(('q_74835 => 'q_74839) * + ('q_74840 => 'q_74835 => 'q_74799)) list + => 'q_74840 => 'q_74839 => 'q_74799. + (ALL (f::'q_74840) x::'q_74839. + CASEWISE [] f x = (SOME y::'q_74799. True)) & + (ALL (h::('q_74835 => 'q_74839) * ('q_74840 => 'q_74835 => 'q_74799)) + (t::(('q_74835 => 'q_74839) * + ('q_74840 => 'q_74835 => 'q_74799)) list) + (f::'q_74840) x::'q_74839. + CASEWISE (h # t) f x = + (if EX y::'q_74835. fst h y = x + then snd h f (SOME y::'q_74835. fst h y = x) else CASEWISE t f x)))" by (import hollight DEF_CASEWISE) -lemma CASEWISE: "(op &::bool => bool => bool) - ((op =::'q_57950::type => 'q_57950::type => bool) - ((CASEWISE::(('q_57942::type => 'q_57990::type) * - ('q_57991::type - => 'q_57942::type => 'q_57950::type)) hollight.list - => 'q_57991::type => 'q_57990::type => 'q_57950::type) - (NIL::(('q_57942::type => 'q_57990::type) * - ('q_57991::type - => 'q_57942::type => 'q_57950::type)) hollight.list) - (f::'q_57991::type) (x::'q_57990::type)) - ((Eps::('q_57950::type => bool) => 'q_57950::type) - (%y::'q_57950::type. True::bool))) - ((op =::'q_57951::type => 'q_57951::type => bool) - ((CASEWISE::(('q_57993::type => 'q_57990::type) * - ('q_57991::type - => 'q_57993::type => 'q_57951::type)) hollight.list - => 'q_57991::type => 'q_57990::type => 'q_57951::type) - ((CONS::('q_57993::type => 'q_57990::type) * - ('q_57991::type => 'q_57993::type => 'q_57951::type) - => (('q_57993::type => 'q_57990::type) * - ('q_57991::type - => 'q_57993::type => 'q_57951::type)) hollight.list - => (('q_57993::type => 'q_57990::type) * - ('q_57991::type - => 'q_57993::type => 'q_57951::type)) hollight.list) - ((Pair::('q_57993::type => 'q_57990::type) - => ('q_57991::type => 'q_57993::type => 'q_57951::type) - => ('q_57993::type => 'q_57990::type) * - ('q_57991::type => 'q_57993::type => 'q_57951::type)) - (s::'q_57993::type => 'q_57990::type) - (t::'q_57991::type => 'q_57993::type => 'q_57951::type)) - (clauses::(('q_57993::type => 'q_57990::type) * - ('q_57991::type - => 'q_57993::type => 'q_57951::type)) hollight.list)) - f x) - ((COND::bool => 'q_57951::type => 'q_57951::type => 'q_57951::type) - ((Ex::('q_57993::type => bool) => bool) - (%y::'q_57993::type. - (op =::'q_57990::type => 'q_57990::type => bool) (s y) x)) - (t f ((Eps::('q_57993::type => bool) => 'q_57993::type) - (%y::'q_57993::type. - (op =::'q_57990::type => 'q_57990::type => bool) (s y) x))) - ((CASEWISE::(('q_57993::type => 'q_57990::type) * - ('q_57991::type - => 'q_57993::type => 'q_57951::type)) hollight.list - => 'q_57991::type => 'q_57990::type => 'q_57951::type) - clauses f x)))" - by (import hollight CASEWISE) - -lemma CASEWISE_CASES: "ALL (clauses::(('q_58085::type => 'q_58082::type) * - ('q_58083::type - => 'q_58085::type => 'q_58092::type)) hollight.list) - (c::'q_58083::type) x::'q_58082::type. - (EX (s::'q_58085::type => 'q_58082::type) - (t::'q_58083::type => 'q_58085::type => 'q_58092::type) - a::'q_58085::type. - MEM (s, t) clauses & s a = x & CASEWISE clauses c x = t c a) | - ~ (EX (s::'q_58085::type => 'q_58082::type) - (t::'q_58083::type => 'q_58085::type => 'q_58092::type) - a::'q_58085::type. MEM (s, t) clauses & s a = x) & - CASEWISE clauses c x = (SOME y::'q_58092::type. True)" - by (import hollight CASEWISE_CASES) - -lemma CASEWISE_WORKS: "ALL (x::(('P::type => 'A::type) * - ('C::type => 'P::type => 'B::type)) hollight.list) - xa::'C::type. - (ALL (s::'P::type => 'A::type) (t::'C::type => 'P::type => 'B::type) - (s'::'P::type => 'A::type) (t'::'C::type => 'P::type => 'B::type) - (xb::'P::type) y::'P::type. - MEM (s, t) x & MEM (s', t') x & s xb = s' y --> - t xa xb = t' xa y) --> - ALL_list - (GABS - (%f::('P::type => 'A::type) * ('C::type => 'P::type => 'B::type) - => bool. - ALL (s::'P::type => 'A::type) t::'C::type => 'P::type => 'B::type. - GEQ (f (s, t)) - (ALL xb::'P::type. CASEWISE x xa (s xb) = t xa xb))) - x" - by (import hollight CASEWISE_WORKS) - -definition admissible :: "('q_58228 => 'q_58221 => bool) -=> (('q_58228 => 'q_58224) => 'q_58234 => bool) - => ('q_58234 => 'q_58221) - => (('q_58228 => 'q_58224) => 'q_58234 => 'q_58229) => bool" where +definition + admissible :: "('q_75137 => 'q_75130 => bool) +=> (('q_75137 => 'q_75133) => 'q_75143 => bool) + => ('q_75143 => 'q_75130) + => (('q_75137 => 'q_75133) => 'q_75143 => 'q_75138) => bool" where "admissible == -%(u::'q_58228::type => 'q_58221::type => bool) - (ua::('q_58228::type => 'q_58224::type) => 'q_58234::type => bool) - (ub::'q_58234::type => 'q_58221::type) - uc::('q_58228::type => 'q_58224::type) - => 'q_58234::type => 'q_58229::type. - ALL (f::'q_58228::type => 'q_58224::type) - (g::'q_58228::type => 'q_58224::type) a::'q_58234::type. - ua f a & - ua g a & (ALL z::'q_58228::type. u z (ub a) --> f z = g z) --> +%(u::'q_75137 => 'q_75130 => bool) + (ua::('q_75137 => 'q_75133) => 'q_75143 => bool) + (ub::'q_75143 => 'q_75130) + uc::('q_75137 => 'q_75133) => 'q_75143 => 'q_75138. + ALL (f::'q_75137 => 'q_75133) (g::'q_75137 => 'q_75133) a::'q_75143. + ua f a & ua g a & (ALL z::'q_75137. u z (ub a) --> f z = g z) --> uc f a = uc g a" lemma DEF_admissible: "admissible = -(%(u::'q_58228::type => 'q_58221::type => bool) - (ua::('q_58228::type => 'q_58224::type) => 'q_58234::type => bool) - (ub::'q_58234::type => 'q_58221::type) - uc::('q_58228::type => 'q_58224::type) - => 'q_58234::type => 'q_58229::type. - ALL (f::'q_58228::type => 'q_58224::type) - (g::'q_58228::type => 'q_58224::type) a::'q_58234::type. - ua f a & - ua g a & (ALL z::'q_58228::type. u z (ub a) --> f z = g z) --> +(%(u::'q_75137 => 'q_75130 => bool) + (ua::('q_75137 => 'q_75133) => 'q_75143 => bool) + (ub::'q_75143 => 'q_75130) + uc::('q_75137 => 'q_75133) => 'q_75143 => 'q_75138. + ALL (f::'q_75137 => 'q_75133) (g::'q_75137 => 'q_75133) a::'q_75143. + ua f a & ua g a & (ALL z::'q_75137. u z (ub a) --> f z = g z) --> uc f a = uc g a)" by (import hollight DEF_admissible) -definition tailadmissible :: "('A => 'A => bool) +definition + tailadmissible :: "('A => 'A => bool) => (('A => 'B) => 'P => bool) - => ('P => 'A) => (('A => 'B) => 'P => 'B) => bool" where + => ('P => 'A) => (('A => 'B) => 'P => 'B) => bool" where "tailadmissible == -%(u::'A::type => 'A::type => bool) - (ua::('A::type => 'B::type) => 'P::type => bool) - (ub::'P::type => 'A::type) - uc::('A::type => 'B::type) => 'P::type => 'B::type. - EX (P::('A::type => 'B::type) => 'P::type => bool) - (G::('A::type => 'B::type) => 'P::type => 'A::type) - H::('A::type => 'B::type) => 'P::type => 'B::type. - (ALL (f::'A::type => 'B::type) (a::'P::type) y::'A::type. +%(u::'A => 'A => bool) (ua::('A => 'B) => 'P => bool) (ub::'P => 'A) + uc::('A => 'B) => 'P => 'B. + EX (P::('A => 'B) => 'P => bool) (G::('A => 'B) => 'P => 'A) + H::('A => 'B) => 'P => 'B. + (ALL (f::'A => 'B) (a::'P) y::'A. P f a & u y (G f a) --> u y (ub a)) & - (ALL (f::'A::type => 'B::type) (g::'A::type => 'B::type) a::'P::type. - (ALL z::'A::type. u z (ub a) --> f z = g z) --> + (ALL (f::'A => 'B) (g::'A => 'B) a::'P. + (ALL z::'A. u z (ub a) --> f z = g z) --> P f a = P g a & G f a = G g a & H f a = H g a) & - (ALL (f::'A::type => 'B::type) a::'P::type. - ua f a --> uc f a = COND (P f a) (f (G f a)) (H f a))" + (ALL (f::'A => 'B) a::'P. + ua f a --> uc f a = (if P f a then f (G f a) else H f a))" lemma DEF_tailadmissible: "tailadmissible = -(%(u::'A::type => 'A::type => bool) - (ua::('A::type => 'B::type) => 'P::type => bool) - (ub::'P::type => 'A::type) - uc::('A::type => 'B::type) => 'P::type => 'B::type. - EX (P::('A::type => 'B::type) => 'P::type => bool) - (G::('A::type => 'B::type) => 'P::type => 'A::type) - H::('A::type => 'B::type) => 'P::type => 'B::type. - (ALL (f::'A::type => 'B::type) (a::'P::type) y::'A::type. +(%(u::'A => 'A => bool) (ua::('A => 'B) => 'P => bool) (ub::'P => 'A) + uc::('A => 'B) => 'P => 'B. + EX (P::('A => 'B) => 'P => bool) (G::('A => 'B) => 'P => 'A) + H::('A => 'B) => 'P => 'B. + (ALL (f::'A => 'B) (a::'P) y::'A. P f a & u y (G f a) --> u y (ub a)) & - (ALL (f::'A::type => 'B::type) (g::'A::type => 'B::type) a::'P::type. - (ALL z::'A::type. u z (ub a) --> f z = g z) --> + (ALL (f::'A => 'B) (g::'A => 'B) a::'P. + (ALL z::'A. u z (ub a) --> f z = g z) --> P f a = P g a & G f a = G g a & H f a = H g a) & - (ALL (f::'A::type => 'B::type) a::'P::type. - ua f a --> uc f a = COND (P f a) (f (G f a)) (H f a)))" + (ALL (f::'A => 'B) a::'P. + ua f a --> uc f a = (if P f a then f (G f a) else H f a)))" by (import hollight DEF_tailadmissible) -definition superadmissible :: "('q_58378 => 'q_58378 => bool) -=> (('q_58378 => 'q_58380) => 'q_58386 => bool) - => ('q_58386 => 'q_58378) - => (('q_58378 => 'q_58380) => 'q_58386 => 'q_58380) => bool" where +definition + superadmissible :: "('q_75287 => 'q_75287 => bool) +=> (('q_75287 => 'q_75289) => 'q_75295 => bool) + => ('q_75295 => 'q_75287) + => (('q_75287 => 'q_75289) => 'q_75295 => 'q_75289) => bool" where "superadmissible == -%(u::'q_58378::type => 'q_58378::type => bool) - (ua::('q_58378::type => 'q_58380::type) => 'q_58386::type => bool) - (ub::'q_58386::type => 'q_58378::type) - uc::('q_58378::type => 'q_58380::type) - => 'q_58386::type => 'q_58380::type. - admissible u - (%(f::'q_58378::type => 'q_58380::type) a::'q_58386::type. True) ub - ua --> +%(u::'q_75287 => 'q_75287 => bool) + (ua::('q_75287 => 'q_75289) => 'q_75295 => bool) + (ub::'q_75295 => 'q_75287) + uc::('q_75287 => 'q_75289) => 'q_75295 => 'q_75289. + admissible u (%(f::'q_75287 => 'q_75289) a::'q_75295. True) ub ua --> tailadmissible u ua ub uc" lemma DEF_superadmissible: "superadmissible = -(%(u::'q_58378::type => 'q_58378::type => bool) - (ua::('q_58378::type => 'q_58380::type) => 'q_58386::type => bool) - (ub::'q_58386::type => 'q_58378::type) - uc::('q_58378::type => 'q_58380::type) - => 'q_58386::type => 'q_58380::type. - admissible u - (%(f::'q_58378::type => 'q_58380::type) a::'q_58386::type. True) ub - ua --> +(%(u::'q_75287 => 'q_75287 => bool) + (ua::('q_75287 => 'q_75289) => 'q_75295 => bool) + (ub::'q_75295 => 'q_75287) + uc::('q_75287 => 'q_75289) => 'q_75295 => 'q_75289. + admissible u (%(f::'q_75287 => 'q_75289) a::'q_75295. True) ub ua --> tailadmissible u ua ub uc)" by (import hollight DEF_superadmissible) -lemma SUPERADMISSIBLE_COND: "ALL (x::'A::type => 'A::type => bool) - (xa::('A::type => 'B::type) => 'P::type => bool) - (xb::('A::type => 'B::type) => 'P::type => bool) - (xc::'P::type => 'A::type) - (xd::('A::type => 'B::type) => 'P::type => 'B::type) - xe::('A::type => 'B::type) => 'P::type => 'B::type. - admissible x xa xc xb & - superadmissible x - (%(f::'A::type => 'B::type) x::'P::type. xa f x & xb f x) xc xd & - superadmissible x - (%(f::'A::type => 'B::type) x::'P::type. xa f x & ~ xb f x) xc xe --> - superadmissible x xa xc - (%(f::'A::type => 'B::type) x::'P::type. - COND (xb f x) (xd f x) (xe f x))" - by (import hollight SUPERADMISSIBLE_COND) - -lemma ADMISSIBLE_IMP_SUPERADMISSIBLE: "ALL (x::'A::type => 'A::type => bool) - (xa::('A::type => 'B::type) => 'P::type => bool) - (xb::'P::type => 'A::type) - xc::('A::type => 'B::type) => 'P::type => 'B::type. - admissible x xa xb xc --> superadmissible x xa xb xc" - by (import hollight ADMISSIBLE_IMP_SUPERADMISSIBLE) - -lemma TAIL_IMP_SUPERADMISSIBLE: "ALL (x::'A::type => 'A::type => bool) - (xa::('A::type => 'B::type) => 'P::type => bool) - (xb::'P::type => 'A::type) - xc::('A::type => 'B::type) => 'P::type => 'A::type. - admissible x xa xb xc & - (ALL (f::'A::type => 'B::type) a::'P::type. - xa f a --> (ALL y::'A::type. x y (xc f a) --> x y (xb a))) --> - superadmissible x xa xb - (%(f::'A::type => 'B::type) x::'P::type. f (xc f x))" - by (import hollight TAIL_IMP_SUPERADMISSIBLE) - -lemma ADMISSIBLE_COND: "ALL (u_353::'A::type => 'q_58766::type => bool) - (p::('A::type => 'B::type) => 'P::type => bool) - (P::('A::type => 'B::type) => 'P::type => bool) - (s::'P::type => 'q_58766::type) - (h::('A::type => 'B::type) => 'P::type => 'B::type) - k::('A::type => 'B::type) => 'P::type => 'B::type. - admissible u_353 p s P & - admissible u_353 (%(f::'A::type => 'B::type) x::'P::type. p f x & P f x) - s h & - admissible u_353 - (%(f::'A::type => 'B::type) x::'P::type. p f x & ~ P f x) s k --> - admissible u_353 p s - (%(f::'A::type => 'B::type) x::'P::type. COND (P f x) (h f x) (k f x))" - by (import hollight ADMISSIBLE_COND) - -lemma ADMISSIBLE_CONST: "admissible (u_353::'q_58841::type => 'q_58840::type => bool) - (p::('q_58841::type => 'q_58842::type) => 'q_58843::type => bool) - (s::'q_58843::type => 'q_58840::type) - (%f::'q_58841::type => 'q_58842::type. c::'q_58843::type => 'q_58844::type)" +lemma MATCH_SEQPATTERN: "_MATCH (x::'q_75330) + (_SEQPATTERN (r::'q_75330 => 'q_75323 => bool) + (s::'q_75330 => 'q_75323 => bool)) = +(if Ex (r x) then _MATCH x r else _MATCH x s)" + by (import hollight MATCH_SEQPATTERN) + +lemma ADMISSIBLE_CONST: "admissible (u_556::'q_75351 => 'q_75350 => bool) + (x::('q_75351 => 'q_75352) => 'q_75353 => bool) (xa::'q_75353 => 'q_75350) + (%f::'q_75351 => 'q_75352. xb::'q_75353 => 'q_75354)" by (import hollight ADMISSIBLE_CONST) -lemma ADMISSIBLE_COMB: "ALL (x::'A::type => 'A::type => bool) - (xa::('A::type => 'B::type) => 'P::type => bool) - (xb::'P::type => 'A::type) - (xc::('A::type => 'B::type) => 'P::type => 'C::type => 'D::type) - xd::('A::type => 'B::type) => 'P::type => 'C::type. - admissible x xa xb xc & admissible x xa xb xd --> - admissible x xa xb - (%(f::'A::type => 'B::type) x::'P::type. xc f x (xd f x))" +lemma ADMISSIBLE_BASE: "(!!(f::'A => 'B) a::'P. + (xa::('A => 'B) => 'P => bool) f a + ==> (x::'A => 'A => bool) ((xc::'P => 'A) a) ((xb::'P => 'A) a)) +==> admissible x xa xb (%(f::'A => 'B) x::'P. f (xc x))" + by (import hollight ADMISSIBLE_BASE) + +lemma ADMISSIBLE_COMB: "admissible (x::'A => 'A => bool) (xa::('A => 'B) => 'P => bool) + (xb::'P => 'A) (xc::('A => 'B) => 'P => 'C => 'D) & +admissible x xa xb (xd::('A => 'B) => 'P => 'C) +==> admissible x xa xb (%(f::'A => 'B) x::'P. xc f x (xd f x))" by (import hollight ADMISSIBLE_COMB) -lemma ADMISSIBLE_BASE: "ALL (x::'A::type => 'A::type => bool) - (xa::('A::type => 'B::type) => 'P::type => bool) - (xb::'P::type => 'A::type) xc::'P::type => 'A::type. - (ALL (f::'A::type => 'B::type) a::'P::type. - xa f a --> x (xc a) (xb a)) --> - admissible x xa xb (%(f::'A::type => 'B::type) x::'P::type. f (xc x))" - by (import hollight ADMISSIBLE_BASE) - -lemma ADMISSIBLE_NEST: "ALL (x::'A::type => 'A::type => bool) - (xa::('A::type => 'B::type) => 'P::type => bool) - (xb::'P::type => 'A::type) - xc::('A::type => 'B::type) => 'P::type => 'A::type. - admissible x xa xb xc & - (ALL (f::'A::type => 'B::type) a::'P::type. - xa f a --> x (xc f a) (xb a)) --> - admissible x xa xb (%(f::'A::type => 'B::type) x::'P::type. f (xc f x))" +lemma ADMISSIBLE_RAND: "admissible (x::'A => 'A => bool) (xa::('A => 'B) => 'P => bool) + (xb::'P => 'A) (xd::('A => 'B) => 'P => 'C) +==> admissible x xa xb + (%(f::'A => 'B) x::'P. (xc::'P => 'C => 'D) x (xd f x))" + by (import hollight ADMISSIBLE_RAND) + +lemma ADMISSIBLE_LAMBDA: "admissible (x::'A => 'A => bool) + (%f::'A => 'B. + SOME fa::'C * 'P => bool. + ALL (u::'C) x::'P. fa (u, x) = (xa::('A => 'B) => 'P => bool) f x) + (SOME f::'C * 'P => 'A. ALL (u::'C) x::'P. f (u, x) = (xb::'P => 'A) x) + (%f::'A => 'B. + SOME fa::'C * 'P => bool. + ALL (u::'C) x::'P. + fa (u, x) = (xc::('A => 'B) => 'C => 'P => bool) f u x) +==> admissible x xa xb (%(f::'A => 'B) (x::'P) u::'C. xc f u x)" + by (import hollight ADMISSIBLE_LAMBDA) + +lemma ADMISSIBLE_NEST: "admissible (x::'A => 'A => bool) (xa::('A => 'B) => 'P => bool) + (xb::'P => 'A) (xc::('A => 'B) => 'P => 'A) & +(ALL (f::'A => 'B) a::'P. xa f a --> x (xc f a) (xb a)) +==> admissible x xa xb (%(f::'A => 'B) x::'P. f (xc f x))" by (import hollight ADMISSIBLE_NEST) -lemma ADMISSIBLE_LAMBDA: "ALL (x::'A::type => 'A::type => bool) - (xa::('A::type => 'B::type) => 'P::type => bool) - (xb::'P::type => 'A::type) - xc::'C::type => ('A::type => 'B::type) => 'P::type => 'D::type. - (ALL xd::'C::type. admissible x xa xb (xc xd)) --> - admissible x xa xb - (%(f::'A::type => 'B::type) (x::'P::type) u::'C::type. xc u f x)" - by (import hollight ADMISSIBLE_LAMBDA) - -lemma ADMISSIBLE_NSUM: "ALL (x::'B::type => 'A::type => bool) - (xa::('B::type => 'C::type) => 'P::type => bool) - (xb::'P::type => 'A::type) - (xc::('B::type => 'C::type) => 'P::type => nat => nat) - (xd::'P::type => nat) xe::'P::type => nat. - (ALL xf::nat. - admissible x - (%(f::'B::type => 'C::type) x::'P::type. - <= (xd x) xf & <= xf (xe x) & xa f x) - xb (%(f::'B::type => 'C::type) x::'P::type. xc f x xf)) --> - admissible x xa xb - (%(f::'B::type => 'C::type) x::'P::type. - nsum (dotdot (xd x) (xe x)) (xc f x))" +lemma ADMISSIBLE_COND: "admissible (u_556::'q_75688 => 'q_75687 => bool) + (p::('q_75688 => 'q_75719) => 'P => bool) (s::'P => 'q_75687) + (P::('q_75688 => 'q_75719) => 'P => bool) & +admissible u_556 (%(f::'q_75688 => 'q_75719) x::'P. p f x & P f x) s + (h::('q_75688 => 'q_75719) => 'P => 'q_75744) & +admissible u_556 (%(f::'q_75688 => 'q_75719) x::'P. p f x & ~ P f x) s + (k::('q_75688 => 'q_75719) => 'P => 'q_75744) +==> admissible u_556 p s + (%(f::'q_75688 => 'q_75719) x::'P. if P f x then h f x else k f x)" + by (import hollight ADMISSIBLE_COND) + +lemma ADMISSIBLE_MATCH: "admissible (x::'q_75790 => 'q_75789 => bool) + (xa::('q_75790 => 'q_75791) => 'P => bool) (xb::'P => 'q_75789) + (xc::('q_75790 => 'q_75791) => 'P => 'q_75826) & +admissible x xa xb + (%(f::'q_75790 => 'q_75791) x::'P. + (c::('q_75790 => 'q_75791) => 'P => 'q_75826 => 'q_75823 => bool) f x + (xc f x)) +==> admissible x xa xb + (%(f::'q_75790 => 'q_75791) x::'P. _MATCH (xc f x) (c f x))" + by (import hollight ADMISSIBLE_MATCH) + +lemma ADMISSIBLE_SEQPATTERN: "admissible (x::'q_75867 => 'q_75866 => bool) + (xa::('q_75867 => 'q_75929) => 'P => bool) (xb::'P => 'q_75866) + (%(f::'q_75867 => 'q_75929) x::'P. + Ex ((xc::('q_75867 => 'q_75929) => 'P => 'q_75955 => 'q_75945 => bool) + f x ((xe::('q_75867 => 'q_75929) => 'P => 'q_75955) f x))) & +admissible x + (%(f::'q_75867 => 'q_75929) x::'P. xa f x & Ex (xc f x (xe f x))) xb + (%(f::'q_75867 => 'q_75929) x::'P. xc f x (xe f x)) & +admissible x + (%(f::'q_75867 => 'q_75929) x::'P. xa f x & ~ Ex (xc f x (xe f x))) xb + (%(f::'q_75867 => 'q_75929) x::'P. + (xd::('q_75867 => 'q_75929) => 'P => 'q_75955 => 'q_75945 => bool) f x + (xe f x)) +==> admissible x xa xb + (%(f::'q_75867 => 'q_75929) x::'P. + _SEQPATTERN (xc f x) (xd f x) (xe f x))" + by (import hollight ADMISSIBLE_SEQPATTERN) + +lemma ADMISSIBLE_UNGUARDED_PATTERN: "admissible (u_556::'q_76041 => 'q_76040 => bool) + (p::('q_76041 => 'q_76088) => 'P => bool) (s::'P => 'q_76040) + (pat::('q_76041 => 'q_76088) => 'P => 'q_76121) & +admissible u_556 p s (e::('q_76041 => 'q_76088) => 'P => 'q_76121) & +admissible u_556 (%(f::'q_76041 => 'q_76088) x::'P. p f x & pat f x = e f x) + s (t::('q_76041 => 'q_76088) => 'P => 'q_76128) & +admissible u_556 (%(f::'q_76041 => 'q_76088) x::'P. p f x & pat f x = e f x) + s (y::('q_76041 => 'q_76088) => 'P => 'q_76128) +==> admissible u_556 p s + (%(f::'q_76041 => 'q_76088) x::'P. + _UNGUARDED_PATTERN (pat f x = e f x) (t f x = y f x))" + by (import hollight ADMISSIBLE_UNGUARDED_PATTERN) + +lemma ADMISSIBLE_GUARDED_PATTERN: "admissible (u_556::'q_76215 => 'q_76214 => bool) + (p::('q_76215 => 'q_76292) => 'P => bool) (s::'P => 'q_76214) + (pat::('q_76215 => 'q_76292) => 'P => 'q_76330) & +admissible u_556 p s (e::('q_76215 => 'q_76292) => 'P => 'q_76330) & +admissible u_556 + (%(f::'q_76215 => 'q_76292) x::'P. + p f x & + pat f x = e f x & (q::('q_76215 => 'q_76292) => 'P => bool) f x) + s (t::('q_76215 => 'q_76292) => 'P => 'q_76339) & +admissible u_556 (%(f::'q_76215 => 'q_76292) x::'P. p f x & pat f x = e f x) + s q & +admissible u_556 + (%(f::'q_76215 => 'q_76292) x::'P. p f x & pat f x = e f x & q f x) s + (y::('q_76215 => 'q_76292) => 'P => 'q_76339) +==> admissible u_556 p s + (%(f::'q_76215 => 'q_76292) x::'P. + _GUARDED_PATTERN (pat f x = e f x) (q f x) (t f x = y f x))" + by (import hollight ADMISSIBLE_GUARDED_PATTERN) + +lemma ADMISSIBLE_NSUM: "admissible (x::'B => 'A => bool) + (%f::'B => 'C. + SOME fa::nat * 'P => bool. + ALL (k::nat) x::'P. + fa (k, x) = + ((xd::'P => nat) x <= k & + k <= (xe::'P => nat) x & (xa::('B => 'C) => 'P => bool) f x)) + (SOME f::nat * 'P => 'A. ALL (k::nat) x::'P. f (k, x) = (xb::'P => 'A) x) + (%f::'B => 'C. + SOME fa::nat * 'P => nat. + ALL (k::nat) x::'P. + fa (k, x) = (xc::('B => 'C) => 'P => nat => nat) f x k) +==> admissible x xa xb (%(f::'B => 'C) x::'P. nsum {xd x..xe x} (xc f x))" by (import hollight ADMISSIBLE_NSUM) -lemma ADMISSIBLE_SUM: "ALL (x::'B::type => 'A::type => bool) - (xa::('B::type => 'C::type) => 'P::type => bool) - (xb::'P::type => 'A::type) - (xc::('B::type => 'C::type) => 'P::type => nat => hollight.real) - (xd::'P::type => nat) xe::'P::type => nat. - (ALL xf::nat. - admissible x - (%(f::'B::type => 'C::type) x::'P::type. - <= (xd x) xf & <= xf (xe x) & xa f x) - xb (%(f::'B::type => 'C::type) x::'P::type. xc f x xf)) --> - admissible x xa xb - (%(f::'B::type => 'C::type) x::'P::type. - hollight.sum (dotdot (xd x) (xe x)) (xc f x))" +lemma ADMISSIBLE_SUM: "admissible (x::'B => 'A => bool) + (%f::'B => 'C. + SOME fa::nat * 'P => bool. + ALL (k::nat) x::'P. + fa (k, x) = + ((xd::'P => nat) x <= k & + k <= (xe::'P => nat) x & (xa::('B => 'C) => 'P => bool) f x)) + (SOME f::nat * 'P => 'A. ALL (k::nat) x::'P. f (k, x) = (xb::'P => 'A) x) + (%f::'B => 'C. + SOME fa::nat * 'P => hollight.real. + ALL (k::nat) x::'P. + fa (k, x) = (xc::('B => 'C) => 'P => nat => hollight.real) f x k) +==> admissible x xa xb + (%(f::'B => 'C) x::'P. hollight.sum {xd x..xe x} (xc f x))" by (import hollight ADMISSIBLE_SUM) -lemma WF_REC_CASES: "ALL (u_353::'A::type => 'A::type => bool) - clauses::(('P::type => 'A::type) * - (('A::type => 'B::type) - => 'P::type => 'B::type)) hollight.list. - WF u_353 & - ALL_list - (GABS - (%f::('P::type => 'A::type) * - (('A::type => 'B::type) => 'P::type => 'B::type) - => bool. - ALL (s::'P::type => 'A::type) - t::('A::type => 'B::type) => 'P::type => 'B::type. - GEQ (f (s, t)) - (EX (P::('A::type => 'B::type) => 'P::type => bool) - (G::('A::type => 'B::type) => 'P::type => 'A::type) - H::('A::type => 'B::type) => 'P::type => 'B::type. - (ALL (f::'A::type => 'B::type) (a::'P::type) y::'A::type. - P f a & u_353 y (G f a) --> u_353 y (s a)) & - (ALL (f::'A::type => 'B::type) (g::'A::type => 'B::type) - a::'P::type. - (ALL z::'A::type. u_353 z (s a) --> f z = g z) --> - P f a = P g a & G f a = G g a & H f a = H g a) & - (ALL (f::'A::type => 'B::type) a::'P::type. - t f a = COND (P f a) (f (G f a)) (H f a))))) - clauses --> - (EX f::'A::type => 'B::type. ALL x::'A::type. f x = CASEWISE clauses f x)" - by (import hollight WF_REC_CASES) - -lemma RECURSION_CASEWISE: "ALL clauses::(('P::type => 'A::type) * - (('A::type => 'B::type) - => 'P::type => 'B::type)) hollight.list. - (EX u::'A::type => 'A::type => bool. - WF u & - ALL_list - (GABS - (%f::('P::type => 'A::type) * - (('A::type => 'B::type) => 'P::type => 'B::type) - => bool. - ALL (s::'P::type => 'A::type) - t::('A::type => 'B::type) => 'P::type => 'B::type. - GEQ (f (s, t)) - (tailadmissible u - (%(f::'A::type => 'B::type) a::'P::type. True) s t))) - clauses) & - (ALL (x::'P::type => 'A::type) - (xa::('A::type => 'B::type) => 'P::type => 'B::type) - (xb::'P::type => 'A::type) - (xc::('A::type => 'B::type) => 'P::type => 'B::type) - (xd::'A::type => 'B::type) (xe::'P::type) xf::'P::type. - MEM (x, xa) clauses & MEM (xb, xc) clauses --> - x xe = xb xf --> xa xd xe = xc xd xf) --> - (EX f::'A::type => 'B::type. - ALL_list - (GABS - (%fa::('P::type => 'A::type) * - (('A::type => 'B::type) => 'P::type => 'B::type) - => bool. - ALL (s::'P::type => 'A::type) - t::('A::type => 'B::type) => 'P::type => 'B::type. - GEQ (fa (s, t)) (ALL x::'P::type. f (s x) = t f x))) - clauses)" - by (import hollight RECURSION_CASEWISE) - -lemma cth: "ALL (p1::'A::type => 'q_59947::type) - (p2::'q_59958::type => 'A::type => 'q_59952::type) - (p1'::'A::type => 'q_59947::type) - p2'::'q_59958::type => 'A::type => 'q_59952::type. - (ALL (c::'q_59958::type) (x::'A::type) y::'A::type. - p1 x = p1' y --> p2 c x = p2' c y) --> - (ALL (c::'q_59958::type) (x::'A::type) y::'A::type. - p1' x = p1 y --> p2' c x = p2 c y)" +lemma ADMISSIBLE_MAP: "admissible (x::'A => 'q_76632 => bool) (xa::('A => 'B) => 'P => bool) + (xb::'P => 'q_76632) (xd::('A => 'B) => 'P => 'q_76647 list) & +admissible x + (%f::'A => 'B. + SOME fa::'q_76647 * 'P => bool. + ALL (y::'q_76647) x::'P. fa (y, x) = (xa f x & y : set (xd f x))) + (SOME f::'q_76647 * 'P => 'q_76632. + ALL (y::'q_76647) x::'P. f (y, x) = xb x) + (%f::'A => 'B. + SOME fa::'q_76647 * 'P => 'q_76641. + ALL (y::'q_76647) x::'P. + fa (y, x) = (xc::('A => 'B) => 'P => 'q_76647 => 'q_76641) f x y) +==> admissible x xa xb (%(f::'A => 'B) x::'P. map (xc f x) (xd f x))" + by (import hollight ADMISSIBLE_MAP) + +lemma ADMISSIBLE_MATCH_SEQPATTERN: "admissible (x::'q_76705 => 'q_76704 => bool) + (xa::('q_76705 => 'q_76770) => 'P => bool) (xb::'P => 'q_76704) + (%(f::'q_76705 => 'q_76770) x::'P. + Ex ((xc::('q_76705 => 'q_76770) => 'P => 'q_76825 => 'q_76794 => bool) + f x ((xe::('q_76705 => 'q_76770) => 'P => 'q_76825) f x))) & +admissible x + (%(f::'q_76705 => 'q_76770) x::'P. xa f x & Ex (xc f x (xe f x))) xb + (%(f::'q_76705 => 'q_76770) x::'P. _MATCH (xe f x) (xc f x)) & +admissible x + (%(f::'q_76705 => 'q_76770) x::'P. xa f x & ~ Ex (xc f x (xe f x))) xb + (%(f::'q_76705 => 'q_76770) x::'P. + _MATCH (xe f x) + ((xd::('q_76705 => 'q_76770) => 'P => 'q_76825 => 'q_76794 => bool) f + x)) +==> admissible x xa xb + (%(x::'q_76705 => 'q_76770) xa::'P. + _MATCH (xe x xa) (_SEQPATTERN (xc x xa) (xd x xa)))" + by (import hollight ADMISSIBLE_MATCH_SEQPATTERN) + +lemma ADMISSIBLE_IMP_SUPERADMISSIBLE: "admissible (x::'A => 'A => bool) (xa::('A => 'B) => 'P => bool) + (xb::'P => 'A) (xc::('A => 'B) => 'P => 'B) +==> superadmissible x xa xb xc" + by (import hollight ADMISSIBLE_IMP_SUPERADMISSIBLE) + +lemma SUPERADMISSIBLE_CONST: "superadmissible (u_556::'q_76904 => 'q_76904 => bool) + (p::('q_76904 => 'q_76906) => 'q_76905 => bool) (s::'q_76905 => 'q_76904) + (%f::'q_76904 => 'q_76906. c::'q_76905 => 'q_76906)" + by (import hollight SUPERADMISSIBLE_CONST) + +lemma SUPERADMISSIBLE_TAIL: "admissible (x::'A => 'A => bool) (xa::('A => 'B) => 'P => bool) + (xb::'P => 'A) (xc::('A => 'B) => 'P => 'A) & +(ALL (f::'A => 'B) a::'P. + xa f a --> (ALL y::'A. x y (xc f a) --> x y (xb a))) +==> superadmissible x xa xb (%(f::'A => 'B) x::'P. f (xc f x))" + by (import hollight SUPERADMISSIBLE_TAIL) + +lemma SUPERADMISSIBLE_COND: "admissible (x::'A => 'A => bool) (xa::('A => 'B) => 'P => bool) + (xc::'P => 'A) (xb::('A => 'B) => 'P => bool) & +superadmissible x (%(f::'A => 'B) x::'P. xa f x & xb f x) xc + (xd::('A => 'B) => 'P => 'B) & +superadmissible x (%(f::'A => 'B) x::'P. xa f x & ~ xb f x) xc + (xe::('A => 'B) => 'P => 'B) +==> superadmissible x xa xc + (%(f::'A => 'B) x::'P. if xb f x then xd f x else xe f x)" + by (import hollight SUPERADMISSIBLE_COND) + +lemma SUPERADMISSIBLE_MATCH_SEQPATTERN: "admissible (x::'q_77225 => 'q_77225 => bool) + (xa::('q_77225 => 'q_77341) => 'P => bool) (xb::'P => 'q_77225) + (%(f::'q_77225 => 'q_77341) x::'P. + Ex ((xc::('q_77225 => 'q_77341) => 'P => 'q_77340 => 'q_77341 => bool) + f x ((xe::('q_77225 => 'q_77341) => 'P => 'q_77340) f x))) & +superadmissible x + (%(f::'q_77225 => 'q_77341) x::'P. xa f x & Ex (xc f x (xe f x))) xb + (%(f::'q_77225 => 'q_77341) x::'P. _MATCH (xe f x) (xc f x)) & +superadmissible x + (%(f::'q_77225 => 'q_77341) x::'P. xa f x & ~ Ex (xc f x (xe f x))) xb + (%(f::'q_77225 => 'q_77341) x::'P. + _MATCH (xe f x) + ((xd::('q_77225 => 'q_77341) => 'P => 'q_77340 => 'q_77341 => bool) f + x)) +==> superadmissible x xa xb + (%(x::'q_77225 => 'q_77341) xa::'P. + _MATCH (xe x xa) (_SEQPATTERN (xc x xa) (xd x xa)))" + by (import hollight SUPERADMISSIBLE_MATCH_SEQPATTERN) + +lemma SUPERADMISSIBLE_MATCH_UNGUARDED_PATTERN: "(ALL (f::'A => 'B) (a::'P) (t::'Q) u::'Q. + (p::('A => 'B) => 'P => bool) f a & + (pat::'Q => 'D) t = (e::'P => 'D) a & pat u = e a --> + (arg::'P => 'Q => 'A) a t = arg a u) & +(ALL (f::'A => 'B) (a::'P) t::'Q. + p f a & pat t = e a --> + (ALL y::'A. + (u_556::'A => 'A => bool) y (arg a t) --> + u_556 y ((s::'P => 'A) a))) +==> superadmissible u_556 p s + (%(f::'A => 'B) x::'P. + _MATCH (e x) + (%(u::'D) v::'B. + EX t::'Q. _UNGUARDED_PATTERN (pat t = u) (f (arg x t) = v)))" + by (import hollight SUPERADMISSIBLE_MATCH_UNGUARDED_PATTERN) + +lemma SUPERADMISSIBLE_MATCH_GUARDED_PATTERN: "(ALL (f::'A => 'B) (a::'P) (t::'Q) u::'Q. + (p::('A => 'B) => 'P => bool) f a & + (pat::'Q => 'D) t = (e::'P => 'D) a & + (q::'P => 'Q => bool) a t & pat u = e a & q a u --> + (arg::'P => 'Q => 'A) a t = arg a u) & +(ALL (f::'A => 'B) (a::'P) t::'Q. + p f a & q a t & pat t = e a --> + (ALL y::'A. + (u_556::'A => 'A => bool) y (arg a t) --> + u_556 y ((s::'P => 'A) a))) +==> superadmissible u_556 p s + (%(f::'A => 'B) x::'P. + _MATCH (e x) + (%(u::'D) v::'B. + EX t::'Q. + _GUARDED_PATTERN (pat t = u) (q x t) (f (arg x t) = v)))" + by (import hollight SUPERADMISSIBLE_MATCH_GUARDED_PATTERN) + +lemma cth: "[| !!(c::'q_78547) (x::'A) y::'A. + (p1::'A => 'q_78536) x = (p1'::'A => 'q_78536) y + ==> (p2::'q_78547 => 'A => 'q_78541) c x = + (p2'::'q_78547 => 'A => 'q_78541) c y; + p1' (x::'A) = p1 (y::'A) |] +==> p2' (c::'q_78547) x = p2 c y" by (import hollight cth) -lemma RECURSION_CASEWISE_PAIRWISE: "ALL x::(('q_59995::type => 'q_59975::type) * - (('q_59975::type => 'q_59991::type) - => 'q_59995::type => 'q_59991::type)) hollight.list. - (EX u::'q_59975::type => 'q_59975::type => bool. - WF u & - ALL_list - (GABS - (%f::('q_59995::type => 'q_59975::type) * - (('q_59975::type => 'q_59991::type) - => 'q_59995::type => 'q_59991::type) - => bool. - ALL (s::'q_59995::type => 'q_59975::type) - t::('q_59975::type => 'q_59991::type) - => 'q_59995::type => 'q_59991::type. - GEQ (f (s, t)) - (tailadmissible u - (%(f::'q_59975::type => 'q_59991::type) - a::'q_59995::type. True) - s t))) - x) & - ALL_list - (GABS - (%f::('q_59995::type => 'q_59975::type) * - (('q_59975::type => 'q_59991::type) - => 'q_59995::type => 'q_59991::type) - => bool. - ALL (a::'q_59995::type => 'q_59975::type) - b::('q_59975::type => 'q_59991::type) - => 'q_59995::type => 'q_59991::type. - GEQ (f (a, b)) - (ALL (c::'q_59975::type => 'q_59991::type) (x::'q_59995::type) - y::'q_59995::type. a x = a y --> b c x = b c y))) - x & - PAIRWISE - (GABS - (%f::('q_59995::type => 'q_59975::type) * - (('q_59975::type => 'q_59991::type) - => 'q_59995::type => 'q_59991::type) - => ('q_59995::type => 'q_59975::type) * - (('q_59975::type => 'q_59991::type) - => 'q_59995::type => 'q_59991::type) - => bool. - ALL (s::'q_59995::type => 'q_59975::type) - t::('q_59975::type => 'q_59991::type) - => 'q_59995::type => 'q_59991::type. - GEQ (f (s, t)) - (GABS - (%f::('q_59995::type => 'q_59975::type) * - (('q_59975::type => 'q_59991::type) - => 'q_59995::type => 'q_59991::type) - => bool. - ALL (s'::'q_59995::type => 'q_59975::type) - t'::('q_59975::type => 'q_59991::type) - => 'q_59995::type => 'q_59991::type. - GEQ (f (s', t')) - (ALL (c::'q_59975::type => 'q_59991::type) - (x::'q_59995::type) y::'q_59995::type. - s x = s' y --> t c x = t' c y))))) - x --> - (EX f::'q_59975::type => 'q_59991::type. - ALL_list - (GABS - (%fa::('q_59995::type => 'q_59975::type) * - (('q_59975::type => 'q_59991::type) - => 'q_59995::type => 'q_59991::type) - => bool. - ALL (s::'q_59995::type => 'q_59975::type) - t::('q_59975::type => 'q_59991::type) - => 'q_59995::type => 'q_59991::type. - GEQ (fa (s, t)) (ALL x::'q_59995::type. f (s x) = t f x))) - x)" - by (import hollight RECURSION_CASEWISE_PAIRWISE) - -lemma SUPERADMISSIBLE_T: "superadmissible (u_353::'q_60105::type => 'q_60105::type => bool) - (%(f::'q_60105::type => 'q_60107::type) x::'q_60111::type. True) - (s::'q_60111::type => 'q_60105::type) - (t::('q_60105::type => 'q_60107::type) - => 'q_60111::type => 'q_60107::type) = -tailadmissible u_353 - (%(f::'q_60105::type => 'q_60107::type) x::'q_60111::type. True) s t" +lemma SUPERADMISSIBLE_T: "superadmissible (u_556::'q_78694 => 'q_78694 => bool) + (%(f::'q_78694 => 'q_78696) x::'q_78700. True) (s::'q_78700 => 'q_78694) + (t::('q_78694 => 'q_78696) => 'q_78700 => 'q_78696) = +tailadmissible u_556 (%(f::'q_78694 => 'q_78696) x::'q_78700. True) s t" by (import hollight SUPERADMISSIBLE_T) +lemma elemma0: "(ALL z::'q_78985 * 'q_78984. + (f::'q_78985 * 'q_78984 => 'q_78975) z = + (g::'q_78985 * 'q_78984 => 'q_78975) z) = +(ALL (x::'q_78985) y::'q_78984. f (x, y) = g (x, y)) & +(P::'q_79002 * 'q_79001 => 'q_78994) = +(SOME f::'q_79002 * 'q_79001 => 'q_78994. + ALL (x::'q_79002) y::'q_79001. f (x, y) = P (x, y))" + by (import hollight elemma0) + ;end_setup end diff -r 2bd54d4b5f3d -r fea3ed6951e3 src/HOL/Import/HOLLight/hollight.imp --- a/src/HOL/Import/HOLLight/hollight.imp Wed Jul 13 00:23:24 2011 +0900 +++ b/src/HOL/Import/HOLLight/hollight.imp Wed Jul 13 00:29:33 2011 +0900 @@ -2,21 +2,158 @@ import_segment "hollight" +def_maps + "vector" > "vector_def" + "treal_of_num" > "treal_of_num_def" + "treal_neg" > "treal_neg_def" + "treal_mul" > "treal_mul_def" + "treal_le" > "treal_le_def" + "treal_inv" > "treal_inv_def" + "treal_eq" > "treal_eq_def" + "treal_add" > "treal_add_def" + "tailadmissible" > "tailadmissible_def" + "support" > "support_def" + "superadmissible" > "superadmissible_def" + "sum" > "sum_def" + "sndcart" > "sndcart_def" + "rem" > "rem_def" + "real_sub" > "real_sub_def" + "real_sgn" > "real_sgn_def" + "real_pow" > "real_pow_def" + "real_of_num" > "real_of_num_def" + "real_neg" > "real_neg_def" + "real_mul" > "real_mul_def" + "real_mod" > "real_mod_def" + "real_min" > "real_min_def" + "real_max" > "real_max_def" + "real_lt" > "real_lt_def" + "real_le" > "real_le_def" + "real_inv" > "real_inv_def" + "real_gt" > "real_gt_def" + "real_ge" > "real_ge_def" + "real_div" > "real_div_def" + "real_add" > "real_add_def" + "real_abs" > "real_abs_def" + "pastecart" > "pastecart_def" + "pairwise" > "pairwise_def" + "num_of_int" > "num_of_int_def" + "num_mod" > "num_mod_def" + "num_gcd" > "num_gcd_def" + "num_divides" > "num_divides_def" + "num_coprime" > "num_coprime_def" + "nsum" > "nsum_def" + "neutral" > "neutral_def" + "nadd_rinv" > "nadd_rinv_def" + "nadd_of_num" > "nadd_of_num_def" + "nadd_mul" > "nadd_mul_def" + "nadd_le" > "nadd_le_def" + "nadd_inv" > "nadd_inv_def" + "nadd_eq" > "nadd_eq_def" + "nadd_add" > "nadd_add_def" + "monoidal" > "monoidal_def" + "minimal" > "minimal_def" + "lambda" > "lambda_def" + "iterate" > "iterate_def" + "is_nadd" > "is_nadd_def" + "integer" > "integer_def" + "int_sub" > "int_sub_def" + "int_sgn" > "int_sgn_def" + "int_pow" > "int_pow_def" + "int_of_num" > "int_of_num_def" + "int_neg" > "int_neg_def" + "int_mul" > "int_mul_def" + "int_mod" > "int_mod_def" + "int_min" > "int_min_def" + "int_max" > "int_max_def" + "int_lt" > "int_lt_def" + "int_le" > "int_le_def" + "int_gt" > "int_gt_def" + "int_ge" > "int_ge_def" + "int_gcd" > "int_gcd_def" + "int_divides" > "int_divides_def" + "int_coprime" > "int_coprime_def" + "int_add" > "int_add_def" + "int_abs" > "int_abs_def" + "hreal_of_num" > "hreal_of_num_def" + "hreal_mul" > "hreal_mul_def" + "hreal_le" > "hreal_le_def" + "hreal_inv" > "hreal_inv_def" + "hreal_add" > "hreal_add_def" + "fstcart" > "fstcart_def" + "eqeq" > "eqeq_def" + "div" > "div_def" + "dist" > "dist_def" + "dimindex" > "dimindex_def" + "admissible" > "admissible_def" + "_UNGUARDED_PATTERN" > "_UNGUARDED_PATTERN_def" + "_SEQPATTERN" > "_SEQPATTERN_def" + "_MATCH" > "_MATCH_def" + "_GUARDED_PATTERN" > "_GUARDED_PATTERN_def" + "_FUNCTION" > "_FUNCTION_def" + "_FALSITY_" > "_FALSITY__def" + "_11937" > "_11937_def" + "ZRECSPACE" > "ZRECSPACE_def" + "ZCONSTR" > "ZCONSTR_def" + "ZBOT" > "ZBOT_def" + "UNCURRY" > "UNCURRY_def" + "SURJ" > "SURJ_def" + "SING" > "SING_def" + "REST" > "REST_def" + "PASSOC" > "PASSOC_def" + "PAIRWISE" > "PAIRWISE_def" + "NUM_REP" > "NUM_REP_def" + "NUMSUM" > "NUMSUM_def" + "NUMSND" > "NUMSND_def" + "NUMRIGHT" > "NUMRIGHT_def" + "NUMPAIR" > "NUMPAIR_def" + "NUMLEFT" > "NUMLEFT_def" + "NUMFST" > "NUMFST_def" + "LET_END" > "LET_END_def" + "ITSET" > "ITSET_def" + "ISO" > "ISO_def" + "INJP" > "INJP_def" + "INJN" > "INJN_def" + "INJF" > "INJF_def" + "INJA" > "INJA_def" + "INJ" > "INJ_def" + "IND_SUC" > "IND_SUC_def" + "IND_0" > "IND_0_def" + "HAS_SIZE" > "HAS_SIZE_def" + "FNIL" > "FNIL_def" + "FINREC" > "FINREC_def" + "FCONS" > "FCONS_def" + "DECIMAL" > "DECIMAL_def" + "CROSS" > "CROSS_def" + "COUNTABLE" > "COUNTABLE_def" + "CONSTR" > "CONSTR_def" + "CASEWISE" > "CASEWISE_def" + "CARD" > "CARD_def" + "BOTTOM" > "BOTTOM_def" + "BIJ" > "BIJ_def" + "ASCII" > "ASCII_def" + ">_c" > ">_c_def" + ">=_c" > ">=_c_def" + "=_c" > "=_c_def" + "<_c" > "<_c_def" + "<=_c" > "<=_c_def" + "$" > "$_def" + type_maps - "sum" > "+" + "sum" > "Sum_Type.sum" "recspace" > "HOLLight.hollight.recspace" "real" > "HOLLight.hollight.real" "prod" > "Product_Type.prod" - "option" > "HOLLight.hollight.option" + "option" > "Datatype.option" "num" > "Nat.nat" "nadd" > "HOLLight.hollight.nadd" - "list" > "HOLLight.hollight.list" + "list" > "List.list" "int" > "HOLLight.hollight.int" "ind" > "Nat.ind" "hreal" > "HOLLight.hollight.hreal" "fun" > "fun" "finite_sum" > "HOLLight.hollight.finite_sum" "finite_image" > "HOLLight.hollight.finite_image" + "char" > "HOLLight.hollight.char" "cart" > "HOLLight.hollight.cart" "bool" > "HOL.bool" "N_3" > "HOLLight.hollight.N_3" @@ -25,8 +162,7 @@ const_maps "~" > "HOL.Not" - "two_2" > "HOLLight.hollight.two_2" - "two_1" > "HOLLight.hollight.two_1" + "vector" > "HOLLight.hollight.vector" "treal_of_num" > "HOLLight.hollight.treal_of_num" "treal_neg" > "HOLLight.hollight.treal_neg" "treal_mul" > "HOLLight.hollight.treal_mul" @@ -34,20 +170,20 @@ "treal_inv" > "HOLLight.hollight.treal_inv" "treal_eq" > "HOLLight.hollight.treal_eq" "treal_add" > "HOLLight.hollight.treal_add" - "three_3" > "HOLLight.hollight.three_3" - "three_2" > "HOLLight.hollight.three_2" - "three_1" > "HOLLight.hollight.three_1" "tailadmissible" > "HOLLight.hollight.tailadmissible" "support" > "HOLLight.hollight.support" "superadmissible" > "HOLLight.hollight.superadmissible" "sum" > "HOLLight.hollight.sum" "sndcart" > "HOLLight.hollight.sndcart" - "set_of_list" > "HOLLight.hollight.set_of_list" + "set_of_list" > "List.set" + "rem" > "HOLLight.hollight.rem" "real_sub" > "HOLLight.hollight.real_sub" + "real_sgn" > "HOLLight.hollight.real_sgn" "real_pow" > "HOLLight.hollight.real_pow" "real_of_num" > "HOLLight.hollight.real_of_num" "real_neg" > "HOLLight.hollight.real_neg" "real_mul" > "HOLLight.hollight.real_mul" + "real_mod" > "HOLLight.hollight.real_mod" "real_min" > "HOLLight.hollight.real_min" "real_max" > "HOLLight.hollight.real_max" "real_lt" > "HOLLight.hollight.real_lt" @@ -62,6 +198,11 @@ "pairwise" > "HOLLight.hollight.pairwise" "one" > "Product_Type.Unity" "o" > "Fun.comp" + "num_of_int" > "HOLLight.hollight.num_of_int" + "num_mod" > "HOLLight.hollight.num_mod" + "num_gcd" > "HOLLight.hollight.num_gcd" + "num_divides" > "HOLLight.hollight.num_divides" + "num_coprime" > "HOLLight.hollight.num_coprime" "nsum" > "HOLLight.hollight.nsum" "neutral" > "HOLLight.hollight.neutral" "nadd_rinv" > "HOLLight.hollight.nadd_rinv" @@ -72,28 +213,31 @@ "nadd_eq" > "HOLLight.hollight.nadd_eq" "nadd_add" > "HOLLight.hollight.nadd_add" "monoidal" > "HOLLight.hollight.monoidal" - "mod_real" > "HOLLight.hollight.mod_real" - "mod_nat" > "HOLLight.hollight.mod_nat" - "mod_int" > "HOLLight.hollight.mod_int" "mk_pair" > "Product_Type.Pair_Rep" + "mk_num" > "Fun.id" "minimal" > "HOLLight.hollight.minimal" - "measure" > "HOLLight.hollight.measure" - "list_of_set" > "HOLLight.hollight.list_of_set" + "list_EX" > "List.list_ex" + "list_ALL" > "List.list_all" "lambda" > "HOLLight.hollight.lambda" "iterate" > "HOLLight.hollight.iterate" "is_nadd" > "HOLLight.hollight.is_nadd" - "is_int" > "HOLLight.hollight.is_int" + "integer" > "HOLLight.hollight.integer" "int_sub" > "HOLLight.hollight.int_sub" + "int_sgn" > "HOLLight.hollight.int_sgn" "int_pow" > "HOLLight.hollight.int_pow" "int_of_num" > "HOLLight.hollight.int_of_num" "int_neg" > "HOLLight.hollight.int_neg" "int_mul" > "HOLLight.hollight.int_mul" + "int_mod" > "HOLLight.hollight.int_mod" "int_min" > "HOLLight.hollight.int_min" "int_max" > "HOLLight.hollight.int_max" "int_lt" > "HOLLight.hollight.int_lt" "int_le" > "HOLLight.hollight.int_le" "int_gt" > "HOLLight.hollight.int_gt" "int_ge" > "HOLLight.hollight.int_ge" + "int_gcd" > "HOLLight.hollight.int_gcd" + "int_divides" > "HOLLight.hollight.int_divides" + "int_coprime" > "HOLLight.hollight.int_coprime" "int_add" > "HOLLight.hollight.int_add" "int_abs" > "HOLLight.hollight.int_abs" "hreal_of_num" > "HOLLight.hollight.hreal_of_num" @@ -102,76 +246,79 @@ "hreal_inv" > "HOLLight.hollight.hreal_inv" "hreal_add" > "HOLLight.hollight.hreal_add" "fstcart" > "HOLLight.hollight.fstcart" - "finite_index" > "HOLLight.hollight.finite_index" "eqeq" > "HOLLight.hollight.eqeq" - "dotdot" > "HOLLight.hollight.dotdot" + "div" > "HOLLight.hollight.div" "dist" > "HOLLight.hollight.dist" "dimindex" > "HOLLight.hollight.dimindex" "admissible" > "HOLLight.hollight.admissible" + "_UNGUARDED_PATTERN" > "HOLLight.hollight._UNGUARDED_PATTERN" + "_SEQPATTERN" > "HOLLight.hollight._SEQPATTERN" + "_MATCH" > "HOLLight.hollight._MATCH" + "_GUARDED_PATTERN" > "HOLLight.hollight._GUARDED_PATTERN" + "_FUNCTION" > "HOLLight.hollight._FUNCTION" "_FALSITY_" > "HOLLight.hollight._FALSITY_" - "_10328" > "HOLLight.hollight._10328" - "_10327" > "HOLLight.hollight._10327" - "_10326" > "HOLLight.hollight._10326" - "_10303" > "HOLLight.hollight._10303" - "_10302" > "HOLLight.hollight._10302" - "_0" > "0" :: "nat" - "\\/" > HOL.disj + "_11937" > "HOLLight.hollight._11937" + "_0" > "Groups.zero_class.zero" :: "nat" + "\\/" > "HOL.disj" "ZRECSPACE" > "HOLLight.hollight.ZRECSPACE" - "ZIP" > "HOLLight.hollight.ZIP" + "ZIP" > "List.zip" "ZCONSTR" > "HOLLight.hollight.ZCONSTR" "ZBOT" > "HOLLight.hollight.ZBOT" - "WF" > "HOLLight.hollight.WF" - "UNIV" > "HOLLight.hollight.UNIV" - "UNIONS" > "HOLLight.hollight.UNIONS" - "UNION" > "HOLLight.hollight.UNION" + "WF" > "Wellfounded.wfP" + "UNIV" > "Orderings.top_class.top" :: "'a => bool" + "UNIONS" > "Complete_Lattice.Sup_class.Sup" :: "(('a => bool) => bool) => 'a => bool" + "UNION" > "Lattices.semilattice_sup_class.sup" :: "('a => bool) => ('a => bool) => 'a => bool" "UNCURRY" > "HOLLight.hollight.UNCURRY" - "TL" > "HOLLight.hollight.TL" - "T" > "True" + "TL" > "List.tl" + "T" > "HOL.True" "SURJ" > "HOLLight.hollight.SURJ" "SUC" > "Nat.Suc" - "SUBSET" > "HOLLight.hollight.SUBSET" - "SOME" > "HOLLight.hollight.SOME" + "SUBSET" > "Orderings.ord_class.less_eq" :: "('a => bool) => ('a => bool) => bool" + "SOME" > "Datatype.Some" "SND" > "Product_Type.snd" "SING" > "HOLLight.hollight.SING" - "SETSPEC" > "HOLLight.hollight.SETSPEC" - "REVERSE" > "HOLLight.hollight.REVERSE" + "SETSPEC" > "HOLLightCompat.SETSPEC" + "REVERSE" > "List.rev" "REST" > "HOLLight.hollight.REST" - "REP_prod" > "Rep_Prod" - "REPLICATE" > "HOLLight.hollight.REPLICATE" - "PSUBSET" > "HOLLight.hollight.PSUBSET" + "REPLICATE" > "List.replicate" + "PSUBSET" > "Orderings.ord_class.less" :: "('a => bool) => ('a => bool) => bool" "PRE" > "HOLLightCompat.Pred" "PASSOC" > "HOLLight.hollight.PASSOC" "PAIRWISE" > "HOLLight.hollight.PAIRWISE" - "OUTR" > "HOLLight.hollight.OUTR" - "OUTL" > "HOLLight.hollight.OUTL" + "OUTR" > "HOLLightCompat.OUTR" + "OUTL" > "HOLLightCompat.OUTL" "ONTO" > "Fun.surj" - "ONE_ONE" > "HOL4Setup.ONE_ONE" - "ODD" > "HOLLight.hollight.ODD" + "ONE_ONE" > "Fun.inj" + "ODD" > "HOLLightCompat.ODD" + "NUM_REP" > "HOLLight.hollight.NUM_REP" "NUMSUM" > "HOLLight.hollight.NUMSUM" "NUMSND" > "HOLLight.hollight.NUMSND" "NUMRIGHT" > "HOLLight.hollight.NUMRIGHT" "NUMPAIR" > "HOLLight.hollight.NUMPAIR" "NUMLEFT" > "HOLLight.hollight.NUMLEFT" "NUMFST" > "HOLLight.hollight.NUMFST" - "NUMERAL" > "HOL4Compat.NUMERAL" - "NULL" > "HOLLight.hollight.NULL" - "NONE" > "HOLLight.hollight.NONE" - "NIL" > "HOLLight.hollight.NIL" - "MOD" > "HOLLight.hollight.MOD" - "MEM" > "HOLLight.hollight.MEM" - "MAP2" > "HOLLight.hollight.MAP2" - "MAP" > "HOLLight.hollight.MAP" + "NUMERAL" > "HOLLightCompat.NUMERAL" + "NULL" > "List.null" + "NONE" > "Datatype.None" + "NIL" > "List.list.Nil" + "MOD" > "Divides.div_class.mod" :: "nat => nat => nat" + "MIN" > "Orderings.ord_class.min" :: "nat => nat => nat" + "MEM" > "HOLLightList.list_mem" + "MEASURE" > "HOLLightCompat.MEASURE" + "MAX" > "Orderings.ord_class.max" :: "nat => nat => nat" + "MAP2" > "HOLLightList.map2" + "MAP" > "List.map" "LET_END" > "HOLLight.hollight.LET_END" - "LET" > "HOL4Compat.LET" - "LENGTH" > "HOLLight.hollight.LENGTH" - "LAST" > "HOLLight.hollight.LAST" + "LET" > "HOLLightCompat.LET" + "LENGTH" > "List.length" + "LAST" > "List.last" "ITSET" > "HOLLight.hollight.ITSET" - "ITLIST2" > "HOLLight.hollight.ITLIST2" - "ITLIST" > "HOLLight.hollight.ITLIST" + "ITLIST2" > "HOLLightList.fold2" + "ITLIST" > "List.foldr" "ISO" > "HOLLight.hollight.ISO" - "INTERS" > "HOLLight.hollight.INTERS" - "INTER" > "HOLLight.hollight.INTER" - "INSERT" > "HOLLight.hollight.INSERT" + "INTERS" > "Complete_Lattice.Inf_class.Inf" :: "(('a => bool) => bool) => 'a => bool" + "INTER" > "Lattices.semilattice_inf_class.inf" :: "('a => bool) => ('a => bool) => 'a => bool" + "INSERT" > "Set.insert" "INR" > "Sum_Type.Inr" "INL" > "Sum_Type.Inl" "INJP" > "HOLLight.hollight.INJP" @@ -179,80 +326,81 @@ "INJF" > "HOLLight.hollight.INJF" "INJA" > "HOLLight.hollight.INJA" "INJ" > "HOLLight.hollight.INJ" - "INFINITE" > "HOLLight.hollight.INFINITE" - "IN" > "HOLLight.hollight.IN" - "IMAGE" > "HOLLight.hollight.IMAGE" + "INFINITE" > "HOLLightCompat.INFINITE" + "IND_SUC" > "HOLLight.hollight.IND_SUC" + "IND_0" > "HOLLight.hollight.IND_0" + "IN" > "Set.member" + "IMAGE" > "Set.image" "I" > "Fun.id" - "HD" > "HOLLight.hollight.HD" + "HD" > "List.hd" "HAS_SIZE" > "HOLLight.hollight.HAS_SIZE" - "GSPEC" > "HOLLight.hollight.GSPEC" - "GEQ" > "HOLLight.hollight.GEQ" - "GABS" > "HOLLight.hollight.GABS" + "GSPEC" > "Set.Collect" + "GEQ" > "HOL.eq" + "GABS" > "Hilbert_Choice.Eps" "FST" > "Product_Type.fst" "FNIL" > "HOLLight.hollight.FNIL" "FINREC" > "HOLLight.hollight.FINREC" - "FINITE" > "HOLLight.hollight.FINITE" - "FILTER" > "HOLLight.hollight.FILTER" + "FINITE" > "Finite_Set.finite" + "FILTER" > "List.filter" "FCONS" > "HOLLight.hollight.FCONS" - "FACT" > "HOLLight.hollight.FACT" - "F" > "False" - "EXP" > "HOLLight.hollight.EXP" - "EX" > "HOLLight.hollight.EX" - "EVEN" > "HOLLight.hollight.EVEN" - "EMPTY" > "HOLLight.hollight.EMPTY" - "EL" > "HOLLight.hollight.EL" - "DIV" > "HOLLight.hollight.DIV" - "DISJOINT" > "HOLLight.hollight.DISJOINT" - "DIFF" > "HOLLight.hollight.DIFF" - "DELETE" > "HOLLight.hollight.DELETE" + "FACT" > "Fact.fact_class.fact" :: "nat => nat" + "F" > "HOL.False" + "EXP" > "Power.power_class.power" :: "nat => nat => nat" + "EVEN" > "Parity.even_odd_class.even" :: "nat => bool" + "EMPTY" > "Orderings.bot_class.bot" :: "'a => bool" + "EL" > "HOLLightList.list_el" + "DIV" > "Divides.div_class.div" :: "nat => nat => nat" + "DISJOINT" > "HOLLightCompat.DISJOINT" + "DIFF" > "Groups.minus_class.minus" :: "('a => bool) => ('a => bool) => 'a => bool" + "DELETE" > "HOLLightCompat.DELETE" "DECIMAL" > "HOLLight.hollight.DECIMAL" - "CURRY" > "HOLLight.hollight.CURRY" + "CURRY" > "Product_Type.curry" + "CROSS" > "HOLLight.hollight.CROSS" "COUNTABLE" > "HOLLight.hollight.COUNTABLE" "CONSTR" > "HOLLight.hollight.CONSTR" - "CONS" > "HOLLight.hollight.CONS" - "COND" > "HOLLight.hollight.COND" - "CHOICE" > "HOLLight.hollight.CHOICE" + "CONS" > "List.list.Cons" + "COND" > "HOL.If" + "CHOICE" > "Hilbert_Choice.Eps" "CASEWISE" > "HOLLight.hollight.CASEWISE" - "CARD_LT" > "HOLLight.hollight.CARD_LT" - "CARD_LE" > "HOLLight.hollight.CARD_LE" - "CARD_GT" > "HOLLight.hollight.CARD_GT" - "CARD_GE" > "HOLLight.hollight.CARD_GE" - "CARD_EQ" > "HOLLight.hollight.CARD_EQ" "CARD" > "HOLLight.hollight.CARD" + "BUTLAST" > "List.butlast" "BOTTOM" > "HOLLight.hollight.BOTTOM" - "BIT1" > "HOL4Compat.NUMERAL_BIT1" + "BIT1" > "HOLLightCompat.NUMERAL_BIT1" "BIT0" > "HOLLightCompat.NUMERAL_BIT0" "BIJ" > "HOLLight.hollight.BIJ" - "ASSOC" > "HOLLight.hollight.ASSOC" - "APPEND" > "HOLLight.hollight.APPEND" - "ALL_list" > "HOLLight.hollight.ALL_list" - "ALL2" > "HOLLight.hollight.ALL2" - "ABS_prod" > "Abs_Prod" + "ASCII" > "HOLLight.hollight.ASCII" + "APPEND" > "List.append" + "ALL2" > "List.list_all2" "@" > "Hilbert_Choice.Eps" "?!" > "HOL.Ex1" "?" > "HOL.Ex" - ">=" > "HOLLight.hollight.>=" - ">" > "HOLLight.hollight.>" - "==>" > HOL.implies - "=" > HOL.eq - "<=" > "HOLLight.hollight.<=" - "<" > "HOLLight.hollight.<" - "/\\" > HOL.conj - "-" > "Groups.minus" :: "nat => nat => nat" + ">_c" > "HOLLight.hollight.>_c" + ">=_c" > "HOLLight.hollight.>=_c" + ">=" > "Orderings.ord_class.greater_eq" :: "nat => nat => bool" + ">" > "Orderings.ord_class.greater" :: "nat => nat => bool" + "=_c" > "HOLLight.hollight.=_c" + "==>" > "HOL.implies" + "=" > "HOL.eq" + "<_c" > "HOLLight.hollight.<_c" + "<=_c" > "HOLLight.hollight.<=_c" + "<=" > "Orderings.ord_class.less_eq" :: "nat => nat => bool" + "<" > "Orderings.ord_class.less" :: "nat => nat => bool" + "/\\" > "HOL.conj" + ".." > "HOLLightCompat.dotdot" + "-" > "Groups.minus_class.minus" :: "nat => nat => nat" "," > "Product_Type.Pair" - "+" > "Groups.plus" :: "nat => nat => nat" - "*" > "Groups.times" :: "nat => nat => nat" + "+" > "Groups.plus_class.plus" :: "nat => nat => nat" + "*" > "Groups.times_class.times" :: "nat => nat => nat" "$" > "HOLLight.hollight.$" "!" > "HOL.All" const_renames - "ALL" > "ALL_list" + "EX" > "list_EX" + "ALL" > "list_ALL" "==" > "eqeq" - ".." > "dotdot" thm_maps - "two_2_def" > "HOLLight.hollight.two_2_def" - "two_1_def" > "HOLLight.hollight.two_1_def" + "vector_def" > "HOLLight.hollight.vector_def" "treal_of_num_def" > "HOLLight.hollight.treal_of_num_def" "treal_neg_def" > "HOLLight.hollight.treal_neg_def" "treal_mul_def" > "HOLLight.hollight.treal_mul_def" @@ -260,26 +408,24 @@ "treal_inv_def" > "HOLLight.hollight.treal_inv_def" "treal_eq_def" > "HOLLight.hollight.treal_eq_def" "treal_add_def" > "HOLLight.hollight.treal_add_def" - "three_3_def" > "HOLLight.hollight.three_3_def" - "three_2_def" > "HOLLight.hollight.three_2_def" - "three_1_def" > "HOLLight.hollight.three_1_def" "th_cond" > "HOLLight.hollight.th_cond" - "th" > "Fun.id_apply" + "th" > "HOL.eta_contract_eq" "tailadmissible_def" > "HOLLight.hollight.tailadmissible_def" "support_def" > "HOLLight.hollight.support_def" "superadmissible_def" > "HOLLight.hollight.superadmissible_def" "sum_def" > "HOLLight.hollight.sum_def" - "sum_RECURSION" > "HOLLightCompat.sum_Recursion" - "sum_INDUCT" > "HOLLightCompat.sumlift.induct" + "string_INFINITE" > "List.infinite_UNIV_listI" "sth" > "HOLLight.hollight.sth" "sndcart_def" > "HOLLight.hollight.sndcart_def" - "set_of_list_def" > "HOLLight.hollight.set_of_list_def" "right_th" > "HOLLight.hollight.right_th" + "rem_def" > "HOLLight.hollight.rem_def" "real_sub_def" > "HOLLight.hollight.real_sub_def" + "real_sgn_def" > "HOLLight.hollight.real_sgn_def" "real_pow_def" > "HOLLight.hollight.real_pow_def" "real_of_num_def" > "HOLLight.hollight.real_of_num_def" "real_neg_def" > "HOLLight.hollight.real_neg_def" "real_mul_def" > "HOLLight.hollight.real_mul_def" + "real_mod_def" > "HOLLight.hollight.real_mod_def" "real_min_def" > "HOLLight.hollight.real_min_def" "real_max_def" > "HOLLight.hollight.real_max_def" "real_lt_def" > "HOLLight.hollight.real_lt_def" @@ -290,27 +436,34 @@ "real_div_def" > "HOLLight.hollight.real_div_def" "real_add_def" > "HOLLight.hollight.real_add_def" "real_abs_def" > "HOLLight.hollight.real_abs_def" + "real_INFINITE" > "HOLLight.hollight.real_INFINITE" "pastecart_def" > "HOLLight.hollight.pastecart_def" "pairwise_def" > "HOLLight.hollight.pairwise_def" "pair_RECURSION" > "HOLLight.hollight.pair_RECURSION" - "pair_INDUCT" > "Datatype.prod.inducts" + "pair_INDUCT" > "Product_Type.prod.induct" "one_axiom" > "HOLLight.hollight.one_axiom" "one_RECURSION" > "HOLLight.hollight.one_RECURSION" - "one_INDUCT" > "Datatype.unit.inducts" + "one_INDUCT" > "Product_Type.unit.induct" "one_Axiom" > "HOLLight.hollight.one_Axiom" - "one" > "HOL4Compat.one" - "o_THM" > "Fun.o_apply" + "one" > "HOLLightCompat.one" + "o_THM" > "Fun.comp_def" "o_ASSOC" > "HOLLight.hollight.o_ASSOC" + "num_of_int_def" > "HOLLight.hollight.num_of_int_def" + "num_mod_def" > "HOLLight.hollight.num_mod_def" + "num_gcd_def" > "HOLLight.hollight.num_gcd_def" + "num_divides_def" > "HOLLight.hollight.num_divides_def" + "num_coprime_def" > "HOLLight.hollight.num_coprime_def" + "num_congruent" > "HOLLight.hollight.num_congruent" "num_WOP" > "HOLLight.hollight.num_WOP" - "num_WF" > "HOLLight.hollight.num_WF" + "num_WF" > "Nat.nat_less_induct" "num_RECURSION_STD" > "HOLLight.hollight.num_RECURSION_STD" "num_MAX" > "HOLLight.hollight.num_MAX" - "num_INFINITE" > "HOLLight.hollight.num_INFINITE" - "num_INDUCTION" > "List.lexn.induct" + "num_INFINITE" > "Finite_Set.infinite_UNIV_char_0" + "num_INDUCTION" > "Fact.fact_nat.induct" "num_FINITE_AVOID" > "HOLLight.hollight.num_FINITE_AVOID" "num_FINITE" > "HOLLight.hollight.num_FINITE" "num_CASES" > "Nat.nat.nchotomy" - "num_Axiom" > "HOLLight.hollight.num_Axiom" + "num_Axiom" > "HOLLightCompat.num_Axiom" "nsum_def" > "HOLLight.hollight.nsum_def" "neutral_def" > "HOLLight.hollight.neutral_def" "nadd_rinv_def" > "HOLLight.hollight.nadd_rinv_def" @@ -321,21 +474,19 @@ "nadd_eq_def" > "HOLLight.hollight.nadd_eq_def" "nadd_add_def" > "HOLLight.hollight.nadd_add_def" "monoidal_def" > "HOLLight.hollight.monoidal_def" - "mod_real_def" > "HOLLight.hollight.mod_real_def" - "mod_nat_def" > "HOLLight.hollight.mod_nat_def" - "mod_int_def" > "HOLLight.hollight.mod_int_def" "minimal_def" > "HOLLight.hollight.minimal_def" - "measure_def" > "HOLLight.hollight.measure_def" - "list_of_set_def" > "HOLLight.hollight.list_of_set_def" - "list_INDUCT" > "HOLLight.hollight.list_INDUCT" - "list_CASES" > "HOLLight.hollight.list_CASES" + "list_INDUCT" > "List.list.induct" + "list_CASES" > "List.list.nchotomy" "lambda_def" > "HOLLight.hollight.lambda_def" "iterate_def" > "HOLLight.hollight.iterate_def" "is_nadd_def" > "HOLLight.hollight.is_nadd_def" "is_nadd_0" > "HOLLight.hollight.is_nadd_0" - "is_int_def" > "HOLLight.hollight.is_int_def" + "is_int" > "HOLLight.hollight.is_int" + "integer_def" > "HOLLight.hollight.integer_def" "int_sub_th" > "HOLLight.hollight.int_sub_th" "int_sub_def" > "HOLLight.hollight.int_sub_def" + "int_sgn_th" > "HOLLight.hollight.int_sgn_th" + "int_sgn_def" > "HOLLight.hollight.int_sgn_def" "int_pow_th" > "HOLLight.hollight.int_pow_th" "int_pow_def" > "HOLLight.hollight.int_pow_def" "int_of_num_th" > "HOLLight.hollight.int_of_num_th" @@ -344,6 +495,7 @@ "int_neg_def" > "HOLLight.hollight.int_neg_def" "int_mul_th" > "HOLLight.hollight.int_mul_th" "int_mul_def" > "HOLLight.hollight.int_mul_def" + "int_mod_def" > "HOLLight.hollight.int_mod_def" "int_min_th" > "HOLLight.hollight.int_min_th" "int_min_def" > "HOLLight.hollight.int_min_def" "int_max_th" > "HOLLight.hollight.int_max_th" @@ -352,7 +504,11 @@ "int_le_def" > "HOLLight.hollight.int_le_def" "int_gt_def" > "HOLLight.hollight.int_gt_def" "int_ge_def" > "HOLLight.hollight.int_ge_def" - "int_eq" > "HOLLight.hollight.int.dest_int_inject" + "int_gcd_def" > "HOLLight.hollight.int_gcd_def" + "int_eq" > "HOLLight.hollight.int.real_of_int_inject" + "int_divides_def" > "HOLLight.hollight.int_divides_def" + "int_coprime_def" > "HOLLight.hollight.int_coprime_def" + "int_congruent" > "HOLLight.hollight.int_congruent" "int_add_th" > "HOLLight.hollight.int_add_th" "int_add_def" > "HOLLight.hollight.int_add_def" "int_abs_th" > "HOLLight.hollight.int_abs_th" @@ -363,34 +519,34 @@ "hreal_inv_def" > "HOLLight.hollight.hreal_inv_def" "hreal_add_def" > "HOLLight.hollight.hreal_add_def" "fstcart_def" > "HOLLight.hollight.fstcart_def" - "finite_index_def" > "HOLLight.hollight.finite_index_def" "eqeq_def" > "HOLLight.hollight.eqeq_def" - "dotdot_def" > "HOLLight.hollight.dotdot_def" + "elemma0" > "HOLLight.hollight.elemma0" + "div_def" > "HOLLight.hollight.div_def" "dist_def" > "HOLLight.hollight.dist_def" "dimindex_def" > "HOLLight.hollight.dimindex_def" "dest_int_rep" > "HOLLight.hollight.dest_int_rep" "cth" > "HOLLight.hollight.cth" "bool_RECURSION" > "HOLLight.hollight.bool_RECURSION" + "bool_INDUCT" > "Product_Type.bool.induct" "ax__3" > "HOL4Setup.INFINITY_AX" - "ax__2" > "Hilbert_Choice.tfl_some" - "ax__1" > "Presburger.fm_modd_pinf" + "ax__2" > "Hilbert_Choice.someI" + "ax__1" > "HOL.eta_contract_eq" "admissible_def" > "HOLLight.hollight.admissible_def" + "_UNGUARDED_PATTERN_def" > "HOLLight.hollight._UNGUARDED_PATTERN_def" + "_SEQPATTERN_def" > "HOLLight.hollight._SEQPATTERN_def" + "_MATCH_def" > "HOLLight.hollight._MATCH_def" + "_GUARDED_PATTERN_def" > "HOLLight.hollight._GUARDED_PATTERN_def" + "_FUNCTION_def" > "HOLLight.hollight._FUNCTION_def" "_FALSITY__def" > "HOLLight.hollight._FALSITY__def" - "_10328_def" > "HOLLight.hollight._10328_def" - "_10327_def" > "HOLLight.hollight._10327_def" - "_10326_def" > "HOLLight.hollight._10326_def" - "_10303_def" > "HOLLight.hollight._10303_def" - "_10302_def" > "HOLLight.hollight._10302_def" + "_11937_def" > "HOLLight.hollight._11937_def" "ZRECSPACE_def" > "HOLLight.hollight.ZRECSPACE_def" - "ZIP_def" > "HOLLight.hollight.ZIP_def" - "ZIP" > "HOLLight.hollight.ZIP" + "ZIP" > "HOLLightList.ZIP" "ZCONSTR_def" > "HOLLight.hollight.ZCONSTR_def" "ZCONSTR_ZBOT" > "HOLLight.hollight.ZCONSTR_ZBOT" "ZBOT_def" > "HOLLight.hollight.ZBOT_def" "WLOG_LT" > "HOLLight.hollight.WLOG_LT" "WLOG_LE" > "HOLLight.hollight.WLOG_LE" "WF_num" > "HOLLight.hollight.WF_num" - "WF_def" > "HOLLight.hollight.WF_def" "WF_UREC_WF" > "HOLLight.hollight.WF_UREC_WF" "WF_UREC" > "HOLLight.hollight.WF_UREC" "WF_SUBSET" > "HOLLight.hollight.WF_SUBSET" @@ -400,49 +556,52 @@ "WF_REC_TAIL_GENERAL" > "HOLLight.hollight.WF_REC_TAIL_GENERAL" "WF_REC_TAIL" > "HOLLight.hollight.WF_REC_TAIL" "WF_REC_INVARIANT" > "HOLLight.hollight.WF_REC_INVARIANT" - "WF_REC_CASES" > "HOLLight.hollight.WF_REC_CASES" "WF_REC" > "HOLLight.hollight.WF_REC" "WF_POINTWISE" > "HOLLight.hollight.WF_POINTWISE" "WF_MEASURE_GEN" > "HOLLight.hollight.WF_MEASURE_GEN" "WF_MEASURE" > "HOLLight.hollight.WF_MEASURE" "WF_LEX_DEPENDENT" > "HOLLight.hollight.WF_LEX_DEPENDENT" "WF_LEX" > "HOLLight.hollight.WF_LEX" + "WF_INT_MEASURE_2" > "HOLLight.hollight.WF_INT_MEASURE_2" + "WF_INT_MEASURE" > "HOLLight.hollight.WF_INT_MEASURE" "WF_IND" > "HOLLight.hollight.WF_IND" - "WF_FALSE" > "HOLLight.hollight.WF_FALSE" + "WF_FINITE" > "HOLLight.hollight.WF_FINITE" + "WF_FALSE" > "Wellfounded.wfP_empty" "WF_EREC" > "HOLLight.hollight.WF_EREC" "WF_EQ" > "HOLLight.hollight.WF_EQ" "WF_DCHAIN" > "HOLLight.hollight.WF_DCHAIN" "UNWIND_THM2" > "HOL.simp_thms_39" "UNWIND_THM1" > "HOL.simp_thms_40" - "UNIV_def" > "HOLLight.hollight.UNIV_def" "UNIV_SUBSET" > "HOLLight.hollight.UNIV_SUBSET" - "UNIV_NOT_EMPTY" > "HOLLight.hollight.UNIV_NOT_EMPTY" + "UNIV_NOT_EMPTY" > "Set.UNIV_not_empty" "UNIQUE_SKOLEM_THM" > "HOL.choice_eq" "UNIQUE_SKOLEM_ALT" > "HOLLight.hollight.UNIQUE_SKOLEM_ALT" - "UNION_def" > "HOLLight.hollight.UNION_def" "UNION_UNIV" > "HOLLight.hollight.UNION_UNIV" - "UNION_SUBSET" > "HOLLight.hollight.UNION_SUBSET" - "UNION_OVER_INTER" > "HOLLight.hollight.UNION_OVER_INTER" - "UNION_IDEMPOT" > "HOLLight.hollight.UNION_IDEMPOT" + "UNION_SUBSET" > "Lattices.semilattice_sup_class.le_sup_iff" + "UNION_OVER_INTER" > "Lattices.distrib_lattice_class.distrib_3" + "UNION_IDEMPOT" > "Big_Operators.lattice_class.Sup_fin.idem" "UNION_EMPTY" > "HOLLight.hollight.UNION_EMPTY" - "UNION_COMM" > "HOLLight.hollight.UNION_COMM" - "UNION_ASSOC" > "HOLLight.hollight.UNION_ASSOC" + "UNION_COMM" > "Lattices.lattice_class.inf_sup_aci_5" + "UNION_ASSOC" > "Lattices.lattice_class.inf_sup_aci_6" "UNION_ACI" > "HOLLight.hollight.UNION_ACI" - "UNIONS_def" > "HOLLight.hollight.UNIONS_def" - "UNIONS_INSERT" > "HOLLight.hollight.UNIONS_INSERT" - "UNIONS_2" > "HOLLight.hollight.UNIONS_2" - "UNIONS_1" > "HOLLight.hollight.UNIONS_1" - "UNIONS_0" > "HOLLight.hollight.UNIONS_0" + "UNIONS_UNION" > "Complete_Lattice.Union_Un_distrib" + "UNIONS_SUBSET" > "HOLLight.hollight.UNIONS_SUBSET" + "UNIONS_INTERS" > "HOLLight.hollight.UNIONS_INTERS" + "UNIONS_INSERT" > "Complete_Lattice.Union_insert" + "UNIONS_IMAGE" > "HOLLight.hollight.UNIONS_IMAGE" + "UNIONS_GSPEC" > "HOLLight.hollight.UNIONS_GSPEC" + "UNIONS_2" > "Complete_Lattice.Un_eq_Union" + "UNIONS_1" > "Complete_Lattice.complete_lattice_class.Sup_singleton" + "UNIONS_0" > "Complete_Lattice.Union_empty" "UNCURRY_def" > "HOLLight.hollight.UNCURRY_def" "TYDEF_recspace" > "HOLLight.hollight.TYDEF_recspace" "TYDEF_real" > "HOLLight.hollight.TYDEF_real" - "TYDEF_option" > "HOLLight.hollight.TYDEF_option" "TYDEF_nadd" > "HOLLight.hollight.TYDEF_nadd" - "TYDEF_list" > "HOLLight.hollight.TYDEF_list" "TYDEF_int" > "HOLLight.hollight.TYDEF_int" "TYDEF_hreal" > "HOLLight.hollight.TYDEF_hreal" "TYDEF_finite_sum" > "HOLLight.hollight.TYDEF_finite_sum" "TYDEF_finite_image" > "HOLLight.hollight.TYDEF_finite_image" + "TYDEF_char" > "HOLLight.hollight.TYDEF_char" "TYDEF_cart" > "HOLLight.hollight.TYDEF_cart" "TYDEF_3" > "HOLLight.hollight.TYDEF_3" "TYDEF_2" > "HOLLight.hollight.TYDEF_2" @@ -488,25 +647,43 @@ "TREAL_ADD_LID" > "HOLLight.hollight.TREAL_ADD_LID" "TREAL_ADD_LDISTRIB" > "HOLLight.hollight.TREAL_ADD_LDISTRIB" "TREAL_ADD_ASSOC" > "HOLLight.hollight.TREAL_ADD_ASSOC" - "TL_def" > "HOLLight.hollight.TL_def" - "TAIL_IMP_SUPERADMISSIBLE" > "HOLLight.hollight.TAIL_IMP_SUPERADMISSIBLE" - "SWAP_FORALL_THM" > "HOLLight.hollight.SWAP_FORALL_THM" - "SWAP_EXISTS_THM" > "HOLLight.hollight.SWAP_EXISTS_THM" + "TRANSITIVE_STEPWISE_LT_EQ" > "HOLLight.hollight.TRANSITIVE_STEPWISE_LT_EQ" + "TRANSITIVE_STEPWISE_LT" > "HOLLight.hollight.TRANSITIVE_STEPWISE_LT" + "TRANSITIVE_STEPWISE_LE_EQ" > "HOLLight.hollight.TRANSITIVE_STEPWISE_LE_EQ" + "TRANSITIVE_STEPWISE_LE" > "HOLLight.hollight.TRANSITIVE_STEPWISE_LE" + "TOPOLOGICAL_SORT" > "HOLLight.hollight.TOPOLOGICAL_SORT" + "SWAP_FORALL_THM" > "HOL.all_comm" + "SWAP_EXISTS_THM" > "HOL.ex_comm" "SURJ_def" > "HOLLight.hollight.SURJ_def" "SURJECTIVE_RIGHT_INVERSE" > "HOLLight.hollight.SURJECTIVE_RIGHT_INVERSE" "SURJECTIVE_ON_RIGHT_INVERSE" > "HOLLight.hollight.SURJECTIVE_ON_RIGHT_INVERSE" + "SURJECTIVE_ON_IMAGE" > "HOLLight.hollight.SURJECTIVE_ON_IMAGE" + "SURJECTIVE_MAP" > "HOLLightList.SURJECTIVE_MAP" + "SURJECTIVE_IMAGE_THM" > "HOLLight.hollight.SURJECTIVE_IMAGE_THM" + "SURJECTIVE_IMAGE_EQ" > "HOLLight.hollight.SURJECTIVE_IMAGE_EQ" + "SURJECTIVE_IMAGE" > "HOLLight.hollight.SURJECTIVE_IMAGE" "SURJECTIVE_IFF_INJECTIVE_GEN" > "HOLLight.hollight.SURJECTIVE_IFF_INJECTIVE_GEN" "SURJECTIVE_IFF_INJECTIVE" > "HOLLight.hollight.SURJECTIVE_IFF_INJECTIVE" + "SURJECTIVE_FORALL_THM" > "HOLLight.hollight.SURJECTIVE_FORALL_THM" + "SURJECTIVE_EXISTS_THM" > "HOLLight.hollight.SURJECTIVE_EXISTS_THM" "SUPPORT_SUPPORT" > "HOLLight.hollight.SUPPORT_SUPPORT" "SUPPORT_SUBSET" > "HOLLight.hollight.SUPPORT_SUBSET" "SUPPORT_EMPTY" > "HOLLight.hollight.SUPPORT_EMPTY" "SUPPORT_DELTA" > "HOLLight.hollight.SUPPORT_DELTA" "SUPPORT_CLAUSES" > "HOLLight.hollight.SUPPORT_CLAUSES" + "SUPERADMISSIBLE_TAIL" > "HOLLight.hollight.SUPERADMISSIBLE_TAIL" "SUPERADMISSIBLE_T" > "HOLLight.hollight.SUPERADMISSIBLE_T" + "SUPERADMISSIBLE_MATCH_UNGUARDED_PATTERN" > "HOLLight.hollight.SUPERADMISSIBLE_MATCH_UNGUARDED_PATTERN" + "SUPERADMISSIBLE_MATCH_SEQPATTERN" > "HOLLight.hollight.SUPERADMISSIBLE_MATCH_SEQPATTERN" + "SUPERADMISSIBLE_MATCH_GUARDED_PATTERN" > "HOLLight.hollight.SUPERADMISSIBLE_MATCH_GUARDED_PATTERN" + "SUPERADMISSIBLE_CONST" > "HOLLight.hollight.SUPERADMISSIBLE_CONST" "SUPERADMISSIBLE_COND" > "HOLLight.hollight.SUPERADMISSIBLE_COND" + "SUM_ZERO_EXISTS" > "HOLLight.hollight.SUM_ZERO_EXISTS" "SUM_UNION_RZERO" > "HOLLight.hollight.SUM_UNION_RZERO" + "SUM_UNION_NONZERO" > "HOLLight.hollight.SUM_UNION_NONZERO" "SUM_UNION_LZERO" > "HOLLight.hollight.SUM_UNION_LZERO" "SUM_UNION_EQ" > "HOLLight.hollight.SUM_UNION_EQ" + "SUM_UNIONS_NONZERO" > "HOLLight.hollight.SUM_UNIONS_NONZERO" "SUM_UNION" > "HOLLight.hollight.SUM_UNION" "SUM_TRIV_NUMSEG" > "HOLLight.hollight.SUM_TRIV_NUMSEG" "SUM_SWAP_NUMSEG" > "HOLLight.hollight.SUM_SWAP_NUMSEG" @@ -521,6 +698,7 @@ "SUM_SUB" > "HOLLight.hollight.SUM_SUB" "SUM_SING_NUMSEG" > "HOLLight.hollight.SUM_SING_NUMSEG" "SUM_SING" > "HOLLight.hollight.SUM_SING" + "SUM_RMUL" > "HOLLight.hollight.SUM_RMUL" "SUM_RESTRICT_SET" > "HOLLight.hollight.SUM_RESTRICT_SET" "SUM_RESTRICT" > "HOLLight.hollight.SUM_RESTRICT" "SUM_POS_LE_NUMSEG" > "HOLLight.hollight.SUM_POS_LE_NUMSEG" @@ -528,34 +706,51 @@ "SUM_POS_EQ_0_NUMSEG" > "HOLLight.hollight.SUM_POS_EQ_0_NUMSEG" "SUM_POS_EQ_0" > "HOLLight.hollight.SUM_POS_EQ_0" "SUM_POS_BOUND" > "HOLLight.hollight.SUM_POS_BOUND" + "SUM_PARTIAL_SUC" > "HOLLight.hollight.SUM_PARTIAL_SUC" + "SUM_PARTIAL_PRE" > "HOLLight.hollight.SUM_PARTIAL_PRE" + "SUM_PAIR" > "HOLLight.hollight.SUM_PAIR" "SUM_OFFSET_0" > "HOLLight.hollight.SUM_OFFSET_0" "SUM_OFFSET" > "HOLLight.hollight.SUM_OFFSET" - "SUM_NEG_NUMSEG" > "HOLLight.hollight.SUM_NEG_NUMSEG" "SUM_NEG" > "HOLLight.hollight.SUM_NEG" "SUM_MULTICOUNT_GEN" > "HOLLight.hollight.SUM_MULTICOUNT_GEN" "SUM_MULTICOUNT" > "HOLLight.hollight.SUM_MULTICOUNT" "SUM_LT_ALL" > "HOLLight.hollight.SUM_LT_ALL" "SUM_LT" > "HOLLight.hollight.SUM_LT" + "SUM_LMUL" > "HOLLight.hollight.SUM_LMUL" "SUM_LE_NUMSEG" > "HOLLight.hollight.SUM_LE_NUMSEG" + "SUM_LE_INCLUDED" > "HOLLight.hollight.SUM_LE_INCLUDED" "SUM_LE" > "HOLLight.hollight.SUM_LE" + "SUM_INJECTION" > "HOLLight.hollight.SUM_INJECTION" + "SUM_INCL_EXCL" > "HOLLight.hollight.SUM_INCL_EXCL" + "SUM_IMAGE_NONZERO" > "HOLLight.hollight.SUM_IMAGE_NONZERO" + "SUM_IMAGE_LE" > "HOLLight.hollight.SUM_IMAGE_LE" "SUM_IMAGE_GEN" > "HOLLight.hollight.SUM_IMAGE_GEN" "SUM_IMAGE" > "HOLLight.hollight.SUM_IMAGE" + "SUM_GROUP" > "HOLLight.hollight.SUM_GROUP" "SUM_EQ_SUPERSET" > "HOLLight.hollight.SUM_EQ_SUPERSET" "SUM_EQ_NUMSEG" > "HOLLight.hollight.SUM_EQ_NUMSEG" + "SUM_EQ_GENERAL_INVERSES" > "HOLLight.hollight.SUM_EQ_GENERAL_INVERSES" "SUM_EQ_GENERAL" > "HOLLight.hollight.SUM_EQ_GENERAL" "SUM_EQ_0_NUMSEG" > "HOLLight.hollight.SUM_EQ_0_NUMSEG" "SUM_EQ_0" > "HOLLight.hollight.SUM_EQ_0" "SUM_EQ" > "HOLLight.hollight.SUM_EQ" + "SUM_DIFFS_ALT" > "HOLLight.hollight.SUM_DIFFS_ALT" + "SUM_DIFFS" > "HOLLight.hollight.SUM_DIFFS" "SUM_DIFF" > "HOLLight.hollight.SUM_DIFF" "SUM_DELTA" > "HOLLight.hollight.SUM_DELTA" + "SUM_DELETE_CASES" > "HOLLight.hollight.SUM_DELETE_CASES" + "SUM_DELETE" > "HOLLight.hollight.SUM_DELETE" "SUM_CONST_NUMSEG" > "HOLLight.hollight.SUM_CONST_NUMSEG" "SUM_CONST" > "HOLLight.hollight.SUM_CONST" - "SUM_CMUL_NUMSEG" > "HOLLight.hollight.SUM_CMUL_NUMSEG" - "SUM_CMUL" > "HOLLight.hollight.SUM_CMUL" + "SUM_COMBINE_R" > "HOLLight.hollight.SUM_COMBINE_R" + "SUM_COMBINE_L" > "HOLLight.hollight.SUM_COMBINE_L" + "SUM_CLOSED" > "HOLLight.hollight.SUM_CLOSED" "SUM_CLAUSES_RIGHT" > "HOLLight.hollight.SUM_CLAUSES_RIGHT" "SUM_CLAUSES_NUMSEG" > "HOLLight.hollight.SUM_CLAUSES_NUMSEG" "SUM_CLAUSES_LEFT" > "HOLLight.hollight.SUM_CLAUSES_LEFT" "SUM_CLAUSES" > "HOLLight.hollight.SUM_CLAUSES" + "SUM_CASES_1" > "HOLLight.hollight.SUM_CASES_1" + "SUM_CASES" > "HOLLight.hollight.SUM_CASES" "SUM_BOUND_LT_GEN" > "HOLLight.hollight.SUM_BOUND_LT_GEN" "SUM_BOUND_LT_ALL" > "HOLLight.hollight.SUM_BOUND_LT_ALL" "SUM_BOUND_LT" > "HOLLight.hollight.SUM_BOUND_LT" @@ -564,58 +759,66 @@ "SUM_BIJECTION" > "HOLLight.hollight.SUM_BIJECTION" "SUM_ADD_SPLIT" > "HOLLight.hollight.SUM_ADD_SPLIT" "SUM_ADD_NUMSEG" > "HOLLight.hollight.SUM_ADD_NUMSEG" + "SUM_ADD_GEN" > "HOLLight.hollight.SUM_ADD_GEN" "SUM_ADD" > "HOLLight.hollight.SUM_ADD" "SUM_ABS_NUMSEG" > "HOLLight.hollight.SUM_ABS_NUMSEG" + "SUM_ABS_LE" > "HOLLight.hollight.SUM_ABS_LE" "SUM_ABS_BOUND" > "HOLLight.hollight.SUM_ABS_BOUND" "SUM_ABS" > "HOLLight.hollight.SUM_ABS" "SUM_0" > "HOLLight.hollight.SUM_0" - "SUC_SUB1" > "HOLLight.hollight.SUC_SUB1" - "SUC_INJ" > "Nat.nat.simps_3" + "SUC_SUB1" > "Nat.diff_Suc_1" + "SUC_INJ" > "HOLLightCompat.SUC_INJ" "SUB_SUC" > "Nat.diff_Suc_Suc" "SUB_REFL" > "Nat.diff_self_eq_0" "SUB_PRESUC" > "HOLLight.hollight.SUB_PRESUC" - "SUB_EQ_0" > "HOLLight.hollight.SUB_EQ_0" + "SUB_EQ_0" > "Nat.diff_is_0_eq" "SUB_ELIM_THM" > "HOLLight.hollight.SUB_ELIM_THM" "SUB_ADD_RCANCEL" > "Nat.diff_cancel2" "SUB_ADD_LCANCEL" > "Nat.diff_cancel" - "SUB_ADD" > "HOLLight.hollight.SUB_ADD" + "SUB_ADD" > "Nat.le_add_diff_inverse2" "SUB_0" > "HOLLight.hollight.SUB_0" - "SUBSET_def" > "HOLLight.hollight.SUBSET_def" - "SUBSET_UNIV" > "HOLLight.hollight.SUBSET_UNIV" - "SUBSET_UNION_ABSORPTION" > "HOLLight.hollight.SUBSET_UNION_ABSORPTION" + "SUBSET_UNIV" > "Orderings.top_class.top_greatest" + "SUBSET_UNION_ABSORPTION" > "Lattices.semilattice_sup_class.le_iff_sup" + "SUBSET_UNIONS" > "Complete_Lattice.Union_mono" "SUBSET_UNION" > "HOLLight.hollight.SUBSET_UNION" - "SUBSET_TRANS" > "HOLLight.hollight.SUBSET_TRANS" + "SUBSET_TRANS" > "Orderings.order_trans_rules_23" "SUBSET_RESTRICT" > "HOLLight.hollight.SUBSET_RESTRICT" - "SUBSET_REFL" > "HOLLight.hollight.SUBSET_REFL" - "SUBSET_PSUBSET_TRANS" > "HOLLight.hollight.SUBSET_PSUBSET_TRANS" + "SUBSET_REFL" > "Inductive.basic_monos_1" + "SUBSET_PSUBSET_TRANS" > "Orderings.order_le_less_trans" "SUBSET_NUMSEG" > "HOLLight.hollight.SUBSET_NUMSEG" - "SUBSET_INTER_ABSORPTION" > "HOLLight.hollight.SUBSET_INTER_ABSORPTION" + "SUBSET_INTER_ABSORPTION" > "Lattices.semilattice_inf_class.le_iff_inf" + "SUBSET_INTER" > "Lattices.semilattice_inf_class.le_inf_iff" "SUBSET_INSERT_DELETE" > "HOLLight.hollight.SUBSET_INSERT_DELETE" - "SUBSET_INSERT" > "HOLLight.hollight.SUBSET_INSERT" - "SUBSET_IMAGE" > "HOLLight.hollight.SUBSET_IMAGE" - "SUBSET_EMPTY" > "HOLLight.hollight.SUBSET_EMPTY" - "SUBSET_DIFF" > "HOLLight.hollight.SUBSET_DIFF" + "SUBSET_INSERT" > "Set.subset_insert" + "SUBSET_IMAGE" > "Set.subset_image_iff" + "SUBSET_EMPTY" > "Set.subset_empty" + "SUBSET_DIFF" > "Set.Diff_subset" "SUBSET_DELETE" > "HOLLight.hollight.SUBSET_DELETE" - "SUBSET_ANTISYM" > "HOLLight.hollight.SUBSET_ANTISYM" - "SOME_def" > "HOLLight.hollight.SOME_def" - "SNDCART_PASTECART" > "HOLLight.hollight.SNDCART_PASTECART" + "SUBSET_CARD_EQ" > "HOLLight.hollight.SUBSET_CARD_EQ" + "SUBSET_ANTISYM_EQ" > "Orderings.order_class.eq_iff" + "SUBSET_ANTISYM" > "Orderings.order_antisym" "SND" > "Product_Type.snd_conv" "SKOLEM_THM" > "HOLLight.hollight.SKOLEM_THM" "SING_def" > "HOLLight.hollight.SING_def" + "SING_SUBSET" > "HOLLight.hollight.SING_SUBSET" + "SING_GSPEC" > "HOLLight.hollight.SING_GSPEC" + "SIMPLE_IMAGE_GEN" > "HOLLight.hollight.SIMPLE_IMAGE_GEN" "SIMPLE_IMAGE" > "HOLLight.hollight.SIMPLE_IMAGE" "SET_RECURSION_LEMMA" > "HOLLight.hollight.SET_RECURSION_LEMMA" - "SET_OF_LIST_OF_SET" > "HOLLight.hollight.SET_OF_LIST_OF_SET" - "SET_OF_LIST_APPEND" > "HOLLight.hollight.SET_OF_LIST_APPEND" + "SET_PROVE_CASES" > "HOLLight.hollight.SET_PROVE_CASES" + "SET_PAIR_THM" > "HOLLight.hollight.SET_PAIR_THM" + "SET_OF_LIST_MAP" > "List.set_map" + "SET_OF_LIST_EQ_EMPTY" > "List.set_empty" + "SET_OF_LIST_APPEND" > "List.set_append" "SET_CASES" > "HOLLight.hollight.SET_CASES" - "SETSPEC_def" > "HOLLight.hollight.SETSPEC_def" "SEMIRING_PTHS" > "HOLLight.hollight.SEMIRING_PTHS" "SELECT_UNIQUE" > "HOLLight.hollight.SELECT_UNIQUE" "SELECT_REFL" > "Hilbert_Choice.some_eq_trivial" "SELECT_LEMMA" > "Hilbert_Choice.some_sym_eq_trivial" - "RIGHT_SUB_DISTRIB" > "Nat.nat_distrib_3" + "RIGHT_SUB_DISTRIB" > "Nat.diff_mult_distrib" "RIGHT_OR_FORALL_THM" > "HOL.all_simps_4" "RIGHT_OR_EXISTS_THM" > "HOL.ex_simps_4" - "RIGHT_OR_DISTRIB" > "HOL.conj_disj_distribR" + "RIGHT_OR_DISTRIB" > "Groebner_Basis.dnf_2" "RIGHT_IMP_FORALL_THM" > "HOL.all_simps_6" "RIGHT_IMP_EXISTS_THM" > "HOL.ex_simps_6" "RIGHT_FORALL_OR_THM" > "HOL.all_simps_4" @@ -624,16 +827,11 @@ "RIGHT_EXISTS_AND_THM" > "HOL.ex_simps_2" "RIGHT_AND_FORALL_THM" > "HOL.all_simps_2" "RIGHT_AND_EXISTS_THM" > "HOL.ex_simps_2" - "RIGHT_ADD_DISTRIB" > "Nat.nat_distrib_1" - "REVERSE_def" > "HOLLight.hollight.REVERSE_def" - "REVERSE_REVERSE" > "HOLLight.hollight.REVERSE_REVERSE" - "REVERSE_APPEND" > "HOLLight.hollight.REVERSE_APPEND" + "RIGHT_ADD_DISTRIB" > "Fields.linordered_field_class.sign_simps_26" + "REVERSE_REVERSE" > "List.rev_rev_ident" + "REVERSE_APPEND" > "List.rev_append" "REST_def" > "HOLLight.hollight.REST_def" - "REP_ABS_PAIR" > "HOLLightCompat.REP_ABS_PAIR" - "REPLICATE_def" > "HOLLight.hollight.REPLICATE_def" - "REFL_CLAUSE" > "HOL.simp_thms_6" - "RECURSION_CASEWISE_PAIRWISE" > "HOLLight.hollight.RECURSION_CASEWISE_PAIRWISE" - "RECURSION_CASEWISE" > "HOLLight.hollight.RECURSION_CASEWISE" + "REFL_CLAUSE" > "Groebner_Basis.bool_simps_6" "REAL_WLOG_LT" > "HOLLight.hollight.REAL_WLOG_LT" "REAL_WLOG_LE" > "HOLLight.hollight.REAL_WLOG_LE" "REAL_SUB_TRIANGLE" > "HOLLight.hollight.REAL_SUB_TRIANGLE" @@ -643,6 +841,9 @@ "REAL_SUB_RNEG" > "HOLLight.hollight.REAL_SUB_RNEG" "REAL_SUB_REFL" > "HOLLight.hollight.REAL_SUB_REFL" "REAL_SUB_RDISTRIB" > "HOLLight.hollight.REAL_SUB_RDISTRIB" + "REAL_SUB_POW_R1" > "HOLLight.hollight.REAL_SUB_POW_R1" + "REAL_SUB_POW_L1" > "HOLLight.hollight.REAL_SUB_POW_L1" + "REAL_SUB_POW" > "HOLLight.hollight.REAL_SUB_POW" "REAL_SUB_NEG2" > "HOLLight.hollight.REAL_SUB_NEG2" "REAL_SUB_LZERO" > "HOLLight.hollight.REAL_SUB_LZERO" "REAL_SUB_LT" > "HOLLight.hollight.REAL_SUB_LT" @@ -655,7 +856,15 @@ "REAL_SUB_ABS" > "HOLLight.hollight.REAL_SUB_ABS" "REAL_SUB_0" > "HOLLight.hollight.REAL_SUB_0" "REAL_SOS_EQ_0" > "HOLLight.hollight.REAL_SOS_EQ_0" + "REAL_SGN_NEG" > "HOLLight.hollight.REAL_SGN_NEG" + "REAL_SGN_MUL" > "HOLLight.hollight.REAL_SGN_MUL" + "REAL_SGN_INV" > "HOLLight.hollight.REAL_SGN_INV" + "REAL_SGN_DIV" > "HOLLight.hollight.REAL_SGN_DIV" + "REAL_SGN_ABS" > "HOLLight.hollight.REAL_SGN_ABS" + "REAL_SGN_0" > "HOLLight.hollight.REAL_SGN_0" + "REAL_SGN" > "HOLLight.hollight.REAL_SGN" "REAL_RNEG_UNIQ" > "HOLLight.hollight.REAL_RNEG_UNIQ" + "REAL_POW_ZERO" > "HOLLight.hollight.REAL_POW_ZERO" "REAL_POW_SUB" > "HOLLight.hollight.REAL_POW_SUB" "REAL_POW_POW" > "HOLLight.hollight.REAL_POW_POW" "REAL_POW_ONE" > "HOLLight.hollight.REAL_POW_ONE" @@ -663,17 +872,33 @@ "REAL_POW_NEG" > "HOLLight.hollight.REAL_POW_NEG" "REAL_POW_MUL" > "HOLLight.hollight.REAL_POW_MUL" "REAL_POW_MONO_LT" > "HOLLight.hollight.REAL_POW_MONO_LT" + "REAL_POW_MONO_INV" > "HOLLight.hollight.REAL_POW_MONO_INV" "REAL_POW_MONO" > "HOLLight.hollight.REAL_POW_MONO" + "REAL_POW_LT_1" > "HOLLight.hollight.REAL_POW_LT_1" + "REAL_POW_LT2_REV" > "HOLLight.hollight.REAL_POW_LT2_REV" + "REAL_POW_LT2_ODD_EQ" > "HOLLight.hollight.REAL_POW_LT2_ODD_EQ" + "REAL_POW_LT2_ODD" > "HOLLight.hollight.REAL_POW_LT2_ODD" "REAL_POW_LT2" > "HOLLight.hollight.REAL_POW_LT2" "REAL_POW_LT" > "HOLLight.hollight.REAL_POW_LT" "REAL_POW_LE_1" > "HOLLight.hollight.REAL_POW_LE_1" + "REAL_POW_LE2_REV" > "HOLLight.hollight.REAL_POW_LE2_REV" + "REAL_POW_LE2_ODD_EQ" > "HOLLight.hollight.REAL_POW_LE2_ODD_EQ" + "REAL_POW_LE2_ODD" > "HOLLight.hollight.REAL_POW_LE2_ODD" "REAL_POW_LE2" > "HOLLight.hollight.REAL_POW_LE2" "REAL_POW_LE" > "HOLLight.hollight.REAL_POW_LE" "REAL_POW_INV" > "HOLLight.hollight.REAL_POW_INV" + "REAL_POW_EQ_ODD_EQ" > "HOLLight.hollight.REAL_POW_EQ_ODD_EQ" + "REAL_POW_EQ_ODD" > "HOLLight.hollight.REAL_POW_EQ_ODD" + "REAL_POW_EQ_EQ" > "HOLLight.hollight.REAL_POW_EQ_EQ" + "REAL_POW_EQ_ABS" > "HOLLight.hollight.REAL_POW_EQ_ABS" + "REAL_POW_EQ_1_IMP" > "HOLLight.hollight.REAL_POW_EQ_1_IMP" + "REAL_POW_EQ_1" > "HOLLight.hollight.REAL_POW_EQ_1" "REAL_POW_EQ_0" > "HOLLight.hollight.REAL_POW_EQ_0" + "REAL_POW_EQ" > "HOLLight.hollight.REAL_POW_EQ" "REAL_POW_DIV" > "HOLLight.hollight.REAL_POW_DIV" "REAL_POW_ADD" > "HOLLight.hollight.REAL_POW_ADD" "REAL_POW_2" > "HOLLight.hollight.REAL_POW_2" + "REAL_POW_1_LT" > "HOLLight.hollight.REAL_POW_1_LT" "REAL_POW_1_LE" > "HOLLight.hollight.REAL_POW_1_LE" "REAL_POW_1" > "HOLLight.hollight.REAL_POW_1" "REAL_POW2_ABS" > "HOLLight.hollight.REAL_POW2_ABS" @@ -686,6 +911,8 @@ "REAL_OF_NUM_SUC" > "HOLLight.hollight.REAL_OF_NUM_SUC" "REAL_OF_NUM_SUB" > "HOLLight.hollight.REAL_OF_NUM_SUB" "REAL_OF_NUM_POW" > "HOLLight.hollight.REAL_OF_NUM_POW" + "REAL_OF_NUM_MIN" > "HOLLight.hollight.REAL_OF_NUM_MIN" + "REAL_OF_NUM_MAX" > "HOLLight.hollight.REAL_OF_NUM_MAX" "REAL_OF_NUM_LT" > "HOLLight.hollight.REAL_OF_NUM_LT" "REAL_OF_NUM_GT" > "HOLLight.hollight.REAL_OF_NUM_GT" "REAL_OF_NUM_GE" > "HOLLight.hollight.REAL_OF_NUM_GE" @@ -712,6 +939,8 @@ "REAL_MUL_RINV_UNIQ" > "HOLLight.hollight.REAL_MUL_RINV_UNIQ" "REAL_MUL_RINV" > "HOLLight.hollight.REAL_MUL_RINV" "REAL_MUL_RID" > "HOLLight.hollight.REAL_MUL_RID" + "REAL_MUL_POS_LT" > "HOLLight.hollight.REAL_MUL_POS_LT" + "REAL_MUL_POS_LE" > "HOLLight.hollight.REAL_MUL_POS_LE" "REAL_MUL_LZERO" > "HOLLight.hollight.REAL_MUL_LZERO" "REAL_MUL_LNEG" > "HOLLight.hollight.REAL_MUL_LNEG" "REAL_MUL_LINV_UNIQ" > "HOLLight.hollight.REAL_MUL_LINV_UNIQ" @@ -735,10 +964,12 @@ "REAL_LT_TOTAL" > "HOLLight.hollight.REAL_LT_TOTAL" "REAL_LT_SUB_RADD" > "HOLLight.hollight.REAL_LT_SUB_RADD" "REAL_LT_SUB_LADD" > "HOLLight.hollight.REAL_LT_SUB_LADD" + "REAL_LT_SQUARE_ABS" > "HOLLight.hollight.REAL_LT_SQUARE_ABS" "REAL_LT_SQUARE" > "HOLLight.hollight.REAL_LT_SQUARE" "REAL_LT_RNEG" > "HOLLight.hollight.REAL_LT_RNEG" "REAL_LT_RMUL_EQ" > "HOLLight.hollight.REAL_LT_RMUL_EQ" "REAL_LT_RMUL" > "HOLLight.hollight.REAL_LT_RMUL" + "REAL_LT_RINV" > "HOLLight.hollight.REAL_LT_RINV" "REAL_LT_REFL" > "HOLLight.hollight.REAL_LT_REFL" "REAL_LT_RDIV_EQ" > "HOLLight.hollight.REAL_LT_RDIV_EQ" "REAL_LT_RCANCEL_IMP" > "HOLLight.hollight.REAL_LT_RCANCEL_IMP" @@ -747,6 +978,7 @@ "REAL_LT_NEGTOTAL" > "HOLLight.hollight.REAL_LT_NEGTOTAL" "REAL_LT_NEG2" > "HOLLight.hollight.REAL_LT_NEG2" "REAL_LT_NEG" > "HOLLight.hollight.REAL_LT_NEG" + "REAL_LT_MUL_EQ" > "HOLLight.hollight.REAL_LT_MUL_EQ" "REAL_LT_MUL2" > "HOLLight.hollight.REAL_LT_MUL2" "REAL_LT_MUL" > "HOLLight.hollight.REAL_LT_MUL" "REAL_LT_MIN" > "HOLLight.hollight.REAL_LT_MIN" @@ -754,6 +986,7 @@ "REAL_LT_LNEG" > "HOLLight.hollight.REAL_LT_LNEG" "REAL_LT_LMUL_EQ" > "HOLLight.hollight.REAL_LT_LMUL_EQ" "REAL_LT_LMUL" > "HOLLight.hollight.REAL_LT_LMUL" + "REAL_LT_LINV" > "HOLLight.hollight.REAL_LT_LINV" "REAL_LT_LE" > "HOLLight.hollight.REAL_LT_LE" "REAL_LT_LDIV_EQ" > "HOLLight.hollight.REAL_LT_LDIV_EQ" "REAL_LT_LCANCEL_IMP" > "HOLLight.hollight.REAL_LT_LCANCEL_IMP" @@ -784,7 +1017,6 @@ "REAL_LTE_ADD2" > "HOLLight.hollight.REAL_LTE_ADD2" "REAL_LTE_ADD" > "HOLLight.hollight.REAL_LTE_ADD" "REAL_LNEG_UNIQ" > "HOLLight.hollight.REAL_LNEG_UNIQ" - "REAL_LE_TOTAL" > "HOLLight.hollight.REAL_LE_TOTAL" "REAL_LE_SUB_RADD" > "HOLLight.hollight.REAL_LE_SUB_RADD" "REAL_LE_SUB_LADD" > "HOLLight.hollight.REAL_LE_SUB_LADD" "REAL_LE_SQUARE_ABS" > "HOLLight.hollight.REAL_LE_SQUARE_ABS" @@ -792,16 +1024,18 @@ "REAL_LE_RNEG" > "HOLLight.hollight.REAL_LE_RNEG" "REAL_LE_RMUL_EQ" > "HOLLight.hollight.REAL_LE_RMUL_EQ" "REAL_LE_RMUL" > "HOLLight.hollight.REAL_LE_RMUL" - "REAL_LE_REFL" > "HOLLight.hollight.REAL_LE_REFL" + "REAL_LE_RINV" > "HOLLight.hollight.REAL_LE_RINV" "REAL_LE_RDIV_EQ" > "HOLLight.hollight.REAL_LE_RDIV_EQ" "REAL_LE_RCANCEL_IMP" > "HOLLight.hollight.REAL_LE_RCANCEL_IMP" "REAL_LE_RADD" > "HOLLight.hollight.REAL_LE_RADD" + "REAL_LE_POW_2" > "HOLLight.hollight.REAL_LE_POW_2" "REAL_LE_POW2" > "HOLLight.hollight.REAL_LE_POW2" "REAL_LE_NEGTOTAL" > "HOLLight.hollight.REAL_LE_NEGTOTAL" "REAL_LE_NEGR" > "HOLLight.hollight.REAL_LE_NEGR" "REAL_LE_NEGL" > "HOLLight.hollight.REAL_LE_NEGL" "REAL_LE_NEG2" > "HOLLight.hollight.REAL_LE_NEG2" "REAL_LE_NEG" > "HOLLight.hollight.REAL_LE_NEG" + "REAL_LE_MUL_EQ" > "HOLLight.hollight.REAL_LE_MUL_EQ" "REAL_LE_MUL2" > "HOLLight.hollight.REAL_LE_MUL2" "REAL_LE_MIN" > "HOLLight.hollight.REAL_LE_MIN" "REAL_LE_MAX" > "HOLLight.hollight.REAL_LE_MAX" @@ -809,6 +1043,7 @@ "REAL_LE_LNEG" > "HOLLight.hollight.REAL_LE_LNEG" "REAL_LE_LMUL_EQ" > "HOLLight.hollight.REAL_LE_LMUL_EQ" "REAL_LE_LMUL" > "HOLLight.hollight.REAL_LE_LMUL" + "REAL_LE_LINV" > "HOLLight.hollight.REAL_LE_LINV" "REAL_LE_LDIV_EQ" > "HOLLight.hollight.REAL_LE_LDIV_EQ" "REAL_LE_LCANCEL_IMP" > "HOLLight.hollight.REAL_LE_LCANCEL_IMP" "REAL_LE_LADD" > "HOLLight.hollight.REAL_LE_LADD" @@ -818,7 +1053,6 @@ "REAL_LE_DOUBLE" > "HOLLight.hollight.REAL_LE_DOUBLE" "REAL_LE_DIV2_EQ" > "HOLLight.hollight.REAL_LE_DIV2_EQ" "REAL_LE_DIV" > "HOLLight.hollight.REAL_LE_DIV" - "REAL_LE_ANTISYM" > "HOLLight.hollight.REAL_LE_ANTISYM" "REAL_LE_ADDR" > "HOLLight.hollight.REAL_LE_ADDR" "REAL_LE_ADDL" > "HOLLight.hollight.REAL_LE_ADDL" "REAL_LE_ADD2" > "HOLLight.hollight.REAL_LE_ADD2" @@ -829,18 +1063,24 @@ "REAL_LET_ANTISYM" > "HOLLight.hollight.REAL_LET_ANTISYM" "REAL_LET_ADD2" > "HOLLight.hollight.REAL_LET_ADD2" "REAL_LET_ADD" > "HOLLight.hollight.REAL_LET_ADD" + "REAL_INV_POW" > "HOLLight.hollight.REAL_INV_POW" "REAL_INV_NEG" > "HOLLight.hollight.REAL_INV_NEG" "REAL_INV_MUL" > "HOLLight.hollight.REAL_INV_MUL" + "REAL_INV_LT_1" > "HOLLight.hollight.REAL_INV_LT_1" "REAL_INV_LE_1" > "HOLLight.hollight.REAL_INV_LE_1" "REAL_INV_INV" > "HOLLight.hollight.REAL_INV_INV" + "REAL_INV_EQ_1" > "HOLLight.hollight.REAL_INV_EQ_1" "REAL_INV_EQ_0" > "HOLLight.hollight.REAL_INV_EQ_0" "REAL_INV_DIV" > "HOLLight.hollight.REAL_INV_DIV" + "REAL_INV_1_LT" > "HOLLight.hollight.REAL_INV_1_LT" "REAL_INV_1_LE" > "HOLLight.hollight.REAL_INV_1_LE" "REAL_INV_1" > "HOLLight.hollight.REAL_INV_1" + "REAL_INTEGRAL" > "HOLLight.hollight.REAL_INTEGRAL" "REAL_HREAL_LEMMA2" > "HOLLight.hollight.REAL_HREAL_LEMMA2" "REAL_HREAL_LEMMA1" > "HOLLight.hollight.REAL_HREAL_LEMMA1" "REAL_EQ_SUB_RADD" > "HOLLight.hollight.REAL_EQ_SUB_RADD" "REAL_EQ_SUB_LADD" > "HOLLight.hollight.REAL_EQ_SUB_LADD" + "REAL_EQ_SQUARE_ABS" > "HOLLight.hollight.REAL_EQ_SQUARE_ABS" "REAL_EQ_RDIV_EQ" > "HOLLight.hollight.REAL_EQ_RDIV_EQ" "REAL_EQ_RCANCEL_IMP" > "HOLLight.hollight.REAL_EQ_RCANCEL_IMP" "REAL_EQ_NEG2" > "HOLLight.hollight.REAL_EQ_NEG2" @@ -848,6 +1088,7 @@ "REAL_EQ_MUL_LCANCEL" > "HOLLight.hollight.REAL_EQ_MUL_LCANCEL" "REAL_EQ_LDIV_EQ" > "HOLLight.hollight.REAL_EQ_LDIV_EQ" "REAL_EQ_LCANCEL_IMP" > "HOLLight.hollight.REAL_EQ_LCANCEL_IMP" + "REAL_EQ_INV2" > "HOLLight.hollight.REAL_EQ_INV2" "REAL_EQ_IMP_LE" > "HOLLight.hollight.REAL_EQ_IMP_LE" "REAL_EQ_ADD_RCANCEL_0" > "HOLLight.hollight.REAL_EQ_ADD_RCANCEL_0" "REAL_EQ_ADD_RCANCEL" > "HOLLight.hollight.REAL_EQ_ADD_RCANCEL" @@ -865,6 +1106,8 @@ "REAL_DIFFSQ" > "HOLLight.hollight.REAL_DIFFSQ" "REAL_COMPLETE_SOMEPOS" > "HOLLight.hollight.REAL_COMPLETE_SOMEPOS" "REAL_COMPLETE" > "HOLLight.hollight.REAL_COMPLETE" + "REAL_BOUNDS_LT" > "HOLLight.hollight.REAL_BOUNDS_LT" + "REAL_BOUNDS_LE" > "HOLLight.hollight.REAL_BOUNDS_LE" "REAL_ADD_SUB2" > "HOLLight.hollight.REAL_ADD_SUB2" "REAL_ADD_SUB" > "HOLLight.hollight.REAL_ADD_SUB" "REAL_ADD_RINV" > "HOLLight.hollight.REAL_ADD_RINV" @@ -881,6 +1124,7 @@ "REAL_ABS_STILLNZ" > "HOLLight.hollight.REAL_ABS_STILLNZ" "REAL_ABS_SIGN2" > "HOLLight.hollight.REAL_ABS_SIGN2" "REAL_ABS_SIGN" > "HOLLight.hollight.REAL_ABS_SIGN" + "REAL_ABS_SGN" > "HOLLight.hollight.REAL_ABS_SGN" "REAL_ABS_REFL" > "HOLLight.hollight.REAL_ABS_REFL" "REAL_ABS_POW" > "HOLLight.hollight.REAL_ABS_POW" "REAL_ABS_POS" > "HOLLight.hollight.REAL_ABS_POS" @@ -906,44 +1150,50 @@ "RAT_LEMMA3" > "HOLLight.hollight.RAT_LEMMA3" "RAT_LEMMA2" > "HOLLight.hollight.RAT_LEMMA2" "RAT_LEMMA1" > "HOLLight.hollight.RAT_LEMMA1" - "PSUBSET_def" > "HOLLight.hollight.PSUBSET_def" "PSUBSET_UNIV" > "HOLLight.hollight.PSUBSET_UNIV" - "PSUBSET_TRANS" > "HOLLight.hollight.PSUBSET_TRANS" - "PSUBSET_SUBSET_TRANS" > "HOLLight.hollight.PSUBSET_SUBSET_TRANS" + "PSUBSET_TRANS" > "Orderings.order_less_trans" + "PSUBSET_SUBSET_TRANS" > "Orderings.order_less_le_trans" "PSUBSET_MEMBER" > "HOLLight.hollight.PSUBSET_MEMBER" - "PSUBSET_IRREFL" > "HOLLight.hollight.PSUBSET_IRREFL" + "PSUBSET_IRREFL" > "Orderings.order_less_irrefl" "PSUBSET_INSERT_SUBSET" > "HOLLight.hollight.PSUBSET_INSERT_SUBSET" + "PSUBSET_ALT" > "HOLLight.hollight.PSUBSET_ALT" "PRE_ELIM_THM" > "HOLLight.hollight.PRE_ELIM_THM" - "PASTECART_FST_SND" > "HOLLight.hollight.PASTECART_FST_SND" - "PASTECART_EQ" > "HOLLight.hollight.PASTECART_EQ" + "POWERSET_CLAUSES" > "HOLLight.hollight.POWERSET_CLAUSES" "PASSOC_def" > "HOLLight.hollight.PASSOC_def" - "PAIR_SURJECTIVE" > "Datatype.prod.nchotomy" + "PAIR_SURJECTIVE" > "Product_Type.prod.nchotomy" "PAIR_EXISTS_THM" > "HOLLight.hollight.PAIR_EXISTS_THM" - "PAIR_EQ" > "Datatype.prod.simps_1" + "PAIR_EQ" > "Product_Type.Pair_eq" "PAIRWISE_def" > "HOLLight.hollight.PAIRWISE_def" - "PAIR" > "HOL4Compat.PAIR" - "OUTR_def" > "HOLLight.hollight.OUTR_def" - "OUTL_def" > "HOLLight.hollight.OUTL_def" + "PAIRWISE_SING" > "HOLLight.hollight.PAIRWISE_SING" + "PAIRWISE_MONO" > "HOLLight.hollight.PAIRWISE_MONO" + "PAIRWISE_EMPTY" > "HOLLight.hollight.PAIRWISE_EMPTY" + "PAIR" > "HOLLightCompat.PAIR" "OR_EXISTS_THM" > "HOL.ex_disj_distrib" "OR_CLAUSES" > "HOLLight.hollight.OR_CLAUSES" - "ONE" > "HOLLight.hollight.ONE" - "ODD_def" > "HOLLight.hollight.ODD_def" + "ONE" > "Nat.One_nat_def" + "ODD_SUB" > "HOLLight.hollight.ODD_SUB" "ODD_MULT" > "HOLLight.hollight.ODD_MULT" "ODD_MOD" > "HOLLight.hollight.ODD_MOD" "ODD_EXP" > "HOLLight.hollight.ODD_EXP" - "ODD_EXISTS" > "HOLLight.hollight.ODD_EXISTS" + "ODD_EXISTS" > "Parity.odd_Suc_mult_two_ex" "ODD_DOUBLE" > "HOLLight.hollight.ODD_DOUBLE" - "ODD_ADD" > "HOLLight.hollight.ODD_ADD" + "ODD_ADD" > "Parity.odd_add" + "NUM_REP_def" > "HOLLight.hollight.NUM_REP_def" + "NUM_OF_INT_OF_NUM" > "HOLLight.hollight.NUM_OF_INT_OF_NUM" + "NUM_OF_INT" > "HOLLight.hollight.NUM_OF_INT" "NUM_INTEGRAL_LEMMA" > "HOLLight.hollight.NUM_INTEGRAL_LEMMA" "NUM_INTEGRAL" > "HOLLight.hollight.NUM_INTEGRAL" + "NUM_GCD" > "HOLLight.hollight.NUM_GCD" "NUMSUM_def" > "HOLLight.hollight.NUMSUM_def" "NUMSUM_INJ" > "HOLLight.hollight.NUMSUM_INJ" "NUMSND_def" > "HOLLight.hollight.NUMSND_def" - "NUMSEG_SING" > "HOLLight.hollight.NUMSEG_SING" + "NUMSEG_SING" > "SetInterval.order_class.atLeastAtMost_singleton" "NUMSEG_RREC" > "HOLLight.hollight.NUMSEG_RREC" - "NUMSEG_REC" > "HOLLight.hollight.NUMSEG_REC" - "NUMSEG_OFFSET_IMAGE" > "HOLLight.hollight.NUMSEG_OFFSET_IMAGE" + "NUMSEG_REC" > "SetInterval.atLeastAtMostSuc_conv" + "NUMSEG_OFFSET_IMAGE" > "SetInterval.image_add_atLeastAtMost" + "NUMSEG_LT" > "HOLLight.hollight.NUMSEG_LT" "NUMSEG_LREC" > "HOLLight.hollight.NUMSEG_LREC" + "NUMSEG_LE" > "HOLLight.hollight.NUMSEG_LE" "NUMSEG_EMPTY" > "HOLLight.hollight.NUMSEG_EMPTY" "NUMSEG_COMBINE_R" > "HOLLight.hollight.NUMSEG_COMBINE_R" "NUMSEG_COMBINE_L" > "HOLLight.hollight.NUMSEG_COMBINE_L" @@ -955,10 +1205,11 @@ "NUMPAIR_INJ" > "HOLLight.hollight.NUMPAIR_INJ" "NUMLEFT_def" > "HOLLight.hollight.NUMLEFT_def" "NUMFST_def" > "HOLLight.hollight.NUMFST_def" - "NULL_def" > "HOLLight.hollight.NULL_def" "NSUM_UNION_RZERO" > "HOLLight.hollight.NSUM_UNION_RZERO" + "NSUM_UNION_NONZERO" > "HOLLight.hollight.NSUM_UNION_NONZERO" "NSUM_UNION_LZERO" > "HOLLight.hollight.NSUM_UNION_LZERO" "NSUM_UNION_EQ" > "HOLLight.hollight.NSUM_UNION_EQ" + "NSUM_UNIONS_NONZERO" > "HOLLight.hollight.NSUM_UNIONS_NONZERO" "NSUM_UNION" > "HOLLight.hollight.NSUM_UNION" "NSUM_TRIV_NUMSEG" > "HOLLight.hollight.NSUM_TRIV_NUMSEG" "NSUM_SWAP_NUMSEG" > "HOLLight.hollight.NSUM_SWAP_NUMSEG" @@ -969,13 +1220,11 @@ "NSUM_SUBSET" > "HOLLight.hollight.NSUM_SUBSET" "NSUM_SING_NUMSEG" > "HOLLight.hollight.NSUM_SING_NUMSEG" "NSUM_SING" > "HOLLight.hollight.NSUM_SING" + "NSUM_RMUL" > "HOLLight.hollight.NSUM_RMUL" "NSUM_RESTRICT_SET" > "HOLLight.hollight.NSUM_RESTRICT_SET" "NSUM_RESTRICT" > "HOLLight.hollight.NSUM_RESTRICT" - "NSUM_POS_LE_NUMSEG" > "HOLLight.hollight.NSUM_POS_LE_NUMSEG" - "NSUM_POS_LE" > "HOLLight.hollight.NSUM_POS_LE" - "NSUM_POS_EQ_0_NUMSEG" > "HOLLight.hollight.NSUM_POS_EQ_0_NUMSEG" - "NSUM_POS_EQ_0" > "HOLLight.hollight.NSUM_POS_EQ_0" "NSUM_POS_BOUND" > "HOLLight.hollight.NSUM_POS_BOUND" + "NSUM_PAIR" > "HOLLight.hollight.NSUM_PAIR" "NSUM_OFFSET_0" > "HOLLight.hollight.NSUM_OFFSET_0" "NSUM_OFFSET" > "HOLLight.hollight.NSUM_OFFSET" "NSUM_NSUM_RESTRICT" > "HOLLight.hollight.NSUM_NSUM_RESTRICT" @@ -984,26 +1233,35 @@ "NSUM_MULTICOUNT" > "HOLLight.hollight.NSUM_MULTICOUNT" "NSUM_LT_ALL" > "HOLLight.hollight.NSUM_LT_ALL" "NSUM_LT" > "HOLLight.hollight.NSUM_LT" + "NSUM_LMUL" > "HOLLight.hollight.NSUM_LMUL" "NSUM_LE_NUMSEG" > "HOLLight.hollight.NSUM_LE_NUMSEG" "NSUM_LE" > "HOLLight.hollight.NSUM_LE" + "NSUM_INJECTION" > "HOLLight.hollight.NSUM_INJECTION" + "NSUM_INCL_EXCL" > "HOLLight.hollight.NSUM_INCL_EXCL" + "NSUM_IMAGE_NONZERO" > "HOLLight.hollight.NSUM_IMAGE_NONZERO" "NSUM_IMAGE_GEN" > "HOLLight.hollight.NSUM_IMAGE_GEN" "NSUM_IMAGE" > "HOLLight.hollight.NSUM_IMAGE" + "NSUM_GROUP" > "HOLLight.hollight.NSUM_GROUP" "NSUM_EQ_SUPERSET" > "HOLLight.hollight.NSUM_EQ_SUPERSET" "NSUM_EQ_NUMSEG" > "HOLLight.hollight.NSUM_EQ_NUMSEG" + "NSUM_EQ_GENERAL_INVERSES" > "HOLLight.hollight.NSUM_EQ_GENERAL_INVERSES" "NSUM_EQ_GENERAL" > "HOLLight.hollight.NSUM_EQ_GENERAL" "NSUM_EQ_0_NUMSEG" > "HOLLight.hollight.NSUM_EQ_0_NUMSEG" + "NSUM_EQ_0_IFF_NUMSEG" > "HOLLight.hollight.NSUM_EQ_0_IFF_NUMSEG" + "NSUM_EQ_0_IFF" > "HOLLight.hollight.NSUM_EQ_0_IFF" "NSUM_EQ_0" > "HOLLight.hollight.NSUM_EQ_0" "NSUM_EQ" > "HOLLight.hollight.NSUM_EQ" "NSUM_DIFF" > "HOLLight.hollight.NSUM_DIFF" "NSUM_DELTA" > "HOLLight.hollight.NSUM_DELTA" + "NSUM_DELETE" > "HOLLight.hollight.NSUM_DELETE" "NSUM_CONST_NUMSEG" > "HOLLight.hollight.NSUM_CONST_NUMSEG" "NSUM_CONST" > "HOLLight.hollight.NSUM_CONST" - "NSUM_CMUL_NUMSEG" > "HOLLight.hollight.NSUM_CMUL_NUMSEG" - "NSUM_CMUL" > "HOLLight.hollight.NSUM_CMUL" + "NSUM_CLOSED" > "HOLLight.hollight.NSUM_CLOSED" "NSUM_CLAUSES_RIGHT" > "HOLLight.hollight.NSUM_CLAUSES_RIGHT" "NSUM_CLAUSES_NUMSEG" > "HOLLight.hollight.NSUM_CLAUSES_NUMSEG" "NSUM_CLAUSES_LEFT" > "HOLLight.hollight.NSUM_CLAUSES_LEFT" "NSUM_CLAUSES" > "HOLLight.hollight.NSUM_CLAUSES" + "NSUM_CASES" > "HOLLight.hollight.NSUM_CASES" "NSUM_BOUND_LT_GEN" > "HOLLight.hollight.NSUM_BOUND_LT_GEN" "NSUM_BOUND_LT_ALL" > "HOLLight.hollight.NSUM_BOUND_LT_ALL" "NSUM_BOUND_LT" > "HOLLight.hollight.NSUM_BOUND_LT" @@ -1012,27 +1270,26 @@ "NSUM_BIJECTION" > "HOLLight.hollight.NSUM_BIJECTION" "NSUM_ADD_SPLIT" > "HOLLight.hollight.NSUM_ADD_SPLIT" "NSUM_ADD_NUMSEG" > "HOLLight.hollight.NSUM_ADD_NUMSEG" + "NSUM_ADD_GEN" > "HOLLight.hollight.NSUM_ADD_GEN" "NSUM_ADD" > "HOLLight.hollight.NSUM_ADD" "NSUM_0" > "HOLLight.hollight.NSUM_0" - "NOT_UNIV_PSUBSET" > "HOLLight.hollight.NOT_UNIV_PSUBSET" - "NOT_SUC" > "Nat.nat.simps_1" - "NOT_PSUBSET_EMPTY" > "HOLLight.hollight.NOT_PSUBSET_EMPTY" + "NOT_UNIV_PSUBSET" > "Complete_Lattice.complete_lattice_class.not_top_less" + "NOT_SUC" > "Nat.Suc_not_Zero" + "NOT_PSUBSET_EMPTY" > "Complete_Lattice.complete_lattice_class.not_less_bot" "NOT_ODD" > "HOLLight.hollight.NOT_ODD" - "NOT_LT" > "HOLLight.hollight.NOT_LT" - "NOT_LE" > "HOLLight.hollight.NOT_LE" + "NOT_LT" > "Orderings.linorder_class.not_less" + "NOT_LE" > "Orderings.linorder_class.not_le" "NOT_IN_EMPTY" > "HOLLight.hollight.NOT_IN_EMPTY" - "NOT_INSERT_EMPTY" > "HOLLight.hollight.NOT_INSERT_EMPTY" - "NOT_FORALL_THM" > "Inductive.basic_monos_15" - "NOT_EXISTS_THM" > "Inductive.basic_monos_16" - "NOT_EX" > "HOLLight.hollight.NOT_EX" + "NOT_INSERT_EMPTY" > "Set.insert_not_empty" + "NOT_FORALL_THM" > "HOL.not_all" + "NOT_EXISTS_THM" > "HOL.not_ex" + "NOT_EX" > "HOLLightList.NOT_EX" "NOT_EVEN" > "HOLLight.hollight.NOT_EVEN" "NOT_EQUAL_SETS" > "HOLLight.hollight.NOT_EQUAL_SETS" - "NOT_EMPTY_INSERT" > "HOLLight.hollight.NOT_EMPTY_INSERT" - "NOT_CONS_NIL" > "HOLLight.hollight.NOT_CONS_NIL" + "NOT_EMPTY_INSERT" > "Set.empty_not_insert" + "NOT_CONS_NIL" > "List.list.distinct_2" "NOT_CLAUSES_WEAK" > "HOLLight.hollight.NOT_CLAUSES_WEAK" - "NOT_ALL" > "HOLLight.hollight.NOT_ALL" - "NONE_def" > "HOLLight.hollight.NONE_def" - "NIL_def" > "HOLLight.hollight.NIL_def" + "NOT_ALL" > "HOLLightList.NOT_ALL" "NEUTRAL_REAL_MUL" > "HOLLight.hollight.NEUTRAL_REAL_MUL" "NEUTRAL_REAL_ADD" > "HOLLight.hollight.NEUTRAL_REAL_ADD" "NEUTRAL_MUL" > "HOLLight.hollight.NEUTRAL_MUL" @@ -1105,229 +1362,252 @@ "NADD_ADD_ASSOC" > "HOLLight.hollight.NADD_ADD_ASSOC" "NADD_ADDITIVE" > "HOLLight.hollight.NADD_ADDITIVE" "NADD_ADD" > "HOLLight.hollight.NADD_ADD" - "MULT_SYM" > "Int.zmult_ac_2" + "MULT_SYM" > "Fields.linordered_field_class.sign_simps_40" "MULT_SUC" > "Nat.mult_Suc_right" - "MULT_EXP" > "HOLLight.hollight.MULT_EXP" - "MULT_EQ_1" > "HOLLight.hollight.MULT_EQ_1" + "MULT_EXP" > "Power.comm_monoid_mult_class.power_mult_distrib" + "MULT_EQ_1" > "Nat.nat_mult_eq_1_iff" "MULT_EQ_0" > "Nat.mult_is_0" + "MULT_DIV_LE" > "HOLLight.hollight.MULT_DIV_LE" "MULT_CLAUSES" > "HOLLight.hollight.MULT_CLAUSES" - "MULT_ASSOC" > "Int.zmult_ac_1" + "MULT_ASSOC" > "Fields.linordered_field_class.sign_simps_41" "MULT_AC" > "HOLLight.hollight.MULT_AC" - "MULT_2" > "HOLLight.hollight.MULT_2" - "MULT_0" > "Nat.mult_0_right" + "MULT_2" > "Int.semiring_mult_2" + "MULT_0" > "Divides.arithmetic_simps_41" "MONO_FORALL" > "Inductive.basic_monos_6" "MONO_EXISTS" > "Inductive.basic_monos_5" "MONO_COND" > "HOLLight.hollight.MONO_COND" - "MONO_ALL2" > "HOLLight.hollight.MONO_ALL2" - "MONO_ALL" > "HOLLight.hollight.MONO_ALL" + "MONO_ALL2" > "List.list_all2_mono" + "MONO_ALL" > "HOLLightList.MONO_ALL" "MONOIDAL_REAL_MUL" > "HOLLight.hollight.MONOIDAL_REAL_MUL" "MONOIDAL_REAL_ADD" > "HOLLight.hollight.MONOIDAL_REAL_ADD" "MONOIDAL_MUL" > "HOLLight.hollight.MONOIDAL_MUL" "MONOIDAL_ADD" > "HOLLight.hollight.MONOIDAL_ADD" - "MOD_def" > "HOLLight.hollight.MOD_def" + "MONOIDAL_AC" > "HOLLight.hollight.MONOIDAL_AC" "MOD_UNIQ" > "HOLLight.hollight.MOD_UNIQ" "MOD_MULT_RMOD" > "HOLLight.hollight.MOD_MULT_RMOD" "MOD_MULT_MOD2" > "HOLLight.hollight.MOD_MULT_MOD2" "MOD_MULT_LMOD" > "HOLLight.hollight.MOD_MULT_LMOD" - "MOD_MULT_ADD" > "HOLLight.hollight.MOD_MULT_ADD" + "MOD_MULT_ADD" > "Divides.mod_mult_self3" "MOD_MULT2" > "HOLLight.hollight.MOD_MULT2" - "MOD_MULT" > "HOLLight.hollight.MOD_MULT" "MOD_MOD_REFL" > "HOLLight.hollight.MOD_MOD_REFL" "MOD_MOD" > "HOLLight.hollight.MOD_MOD" - "MOD_LT" > "HOLLight.hollight.MOD_LT" + "MOD_LT" > "Divides.mod_less" "MOD_LE" > "HOLLight.hollight.MOD_LE" "MOD_EXP_MOD" > "HOLLight.hollight.MOD_EXP_MOD" "MOD_EXISTS" > "HOLLight.hollight.MOD_EXISTS" "MOD_EQ_0" > "HOLLight.hollight.MOD_EQ_0" "MOD_EQ" > "HOLLight.hollight.MOD_EQ" "MOD_ADD_MOD" > "HOLLight.hollight.MOD_ADD_MOD" - "MOD_1" > "HOLLight.hollight.MOD_1" - "MOD_0" > "HOLLight.hollight.MOD_0" "MK_REC_INJ" > "HOLLight.hollight.MK_REC_INJ" "MINIMAL" > "HOLLight.hollight.MINIMAL" - "MEM_def" > "HOLLight.hollight.MEM_def" - "MEM_MAP" > "HOLLight.hollight.MEM_MAP" - "MEM_LIST_OF_SET" > "HOLLight.hollight.MEM_LIST_OF_SET" - "MEM_FILTER" > "HOLLight.hollight.MEM_FILTER" - "MEM_EL" > "HOLLight.hollight.MEM_EL" - "MEM_ASSOC" > "HOLLight.hollight.MEM_ASSOC" - "MEM_APPEND" > "HOLLight.hollight.MEM_APPEND" - "MEMBER_NOT_EMPTY" > "HOLLight.hollight.MEMBER_NOT_EMPTY" + "MEM_MAP" > "HOLLightList.MEM_MAP" + "MEM_FILTER" > "HOLLightList.MEM_FILTER" + "MEM_EXISTS_EL" > "HOLLightList.MEM_EXISTS_EL" + "MEM_EL" > "List.nth_mem" + "MEM_APPEND" > "HOLLightList.MEM_APPEND" + "MEMBER_NOT_EMPTY" > "Set.ex_in_conv" "MEASURE_LE" > "HOLLight.hollight.MEASURE_LE" - "MAP_o" > "HOLLight.hollight.MAP_o" - "MAP_def" > "HOLLight.hollight.MAP_def" - "MAP_SND_ZIP" > "HOLLight.hollight.MAP_SND_ZIP" - "MAP_FST_ZIP" > "HOLLight.hollight.MAP_FST_ZIP" - "MAP_EQ_DEGEN" > "HOLLight.hollight.MAP_EQ_DEGEN" - "MAP_EQ_ALL2" > "HOLLight.hollight.MAP_EQ_ALL2" - "MAP_EQ" > "HOLLight.hollight.MAP_EQ" - "MAP_APPEND" > "HOLLight.hollight.MAP_APPEND" - "MAP2_def" > "HOLLight.hollight.MAP2_def" - "MAP2" > "HOLLight.hollight.MAP2" - "LT_TRANS" > "HOLLight.hollight.LT_TRANS" - "LT_SUC_LE" > "HOLLight.hollight.LT_SUC_LE" - "LT_SUC" > "HOLLight.hollight.LT_SUC" - "LT_REFL" > "HOLLight.hollight.LT_REFL" - "LT_NZ" > "HOLLight.hollight.LT_NZ" + "MATCH_SEQPATTERN" > "HOLLight.hollight.MATCH_SEQPATTERN" + "MAP_o" > "List.map.compositionality" + "MAP_SND_ZIP" > "List.map_snd_zip" + "MAP_ID" > "List.map_ident" + "MAP_I" > "List.map.id" + "MAP_FST_ZIP" > "List.map_fst_zip" + "MAP_EQ_NIL" > "List.map_is_Nil_conv" + "MAP_EQ_DEGEN" > "HOLLightList.MAP_EQ_DEGEN" + "MAP_EQ_ALL2" > "HOLLightList.MAP_EQ_ALL2" + "MAP_EQ" > "HOLLightList.MAP_EQ" + "MAP_APPEND" > "List.map_append" + "MAP2" > "HOLLightList.MAP2" + "LT_TRANS" > "Orderings.order_less_trans" + "LT_SUC_LE" > "Nat.le_simps_2" + "LT_SUC" > "Nat.Suc_less_eq" + "LT_REFL" > "Nat.less_not_refl" + "LT_NZ" > "Nat.neq0_conv" "LT_MULT_RCANCEL" > "HOLLight.hollight.LT_MULT_RCANCEL" "LT_MULT_LCANCEL" > "HOLLight.hollight.LT_MULT_LCANCEL" "LT_MULT2" > "HOLLight.hollight.LT_MULT2" - "LT_MULT" > "HOLLight.hollight.LT_MULT" + "LT_MULT" > "Nat.nat_0_less_mult_iff" "LT_LMULT" > "HOLLight.hollight.LT_LMULT" - "LT_LE" > "HOLLight.hollight.LT_LE" - "LT_IMP_LE" > "HOLLight.hollight.LT_IMP_LE" + "LT_LE" > "Nat.nat_less_le" + "LT_IMP_LE" > "FunDef.termination_basic_simps_5" "LT_EXP" > "HOLLight.hollight.LT_EXP" "LT_EXISTS" > "HOLLight.hollight.LT_EXISTS" "LT_CASES" > "HOLLight.hollight.LT_CASES" "LT_ANTISYM" > "HOLLight.hollight.LT_ANTISYM" - "LT_ADD_RCANCEL" > "HOLLight.hollight.LT_ADD_RCANCEL" - "LT_ADD_LCANCEL" > "HOLLight.hollight.LT_ADD_LCANCEL" + "LT_ADD_RCANCEL" > "Groups.ordered_ab_semigroup_add_imp_le_class.add_less_cancel_right" + "LT_ADD_LCANCEL" > "Groups.ordered_ab_semigroup_add_imp_le_class.add_less_cancel_left" "LT_ADDR" > "HOLLight.hollight.LT_ADDR" - "LT_ADD2" > "HOLLight.hollight.LT_ADD2" + "LT_ADD2" > "Groups.add_mono_thms_linordered_field_5" "LT_ADD" > "HOLLight.hollight.LT_ADD" - "LT_0" > "HOLLight.hollight.LT_0" - "LTE_TRANS" > "HOLLight.hollight.LTE_TRANS" + "LT_0" > "Nat.zero_less_Suc" + "LTE_TRANS" > "Orderings.order_less_le_trans" "LTE_CASES" > "HOLLight.hollight.LTE_CASES" "LTE_ANTISYM" > "HOLLight.hollight.LTE_ANTISYM" - "LTE_ADD2" > "HOLLight.hollight.LTE_ADD2" - "LIST_OF_SET_PROPERTIES" > "HOLLight.hollight.LIST_OF_SET_PROPERTIES" - "LE_TRANS" > "HOLLight.hollight.LE_TRANS" - "LE_SUC_LT" > "HOLLight.hollight.LE_SUC_LT" - "LE_SUC" > "HOLLight.hollight.LE_SUC" - "LE_SQUARE_REFL" > "HOLLight.hollight.LE_SQUARE_REFL" - "LE_REFL" > "HOLLight.hollight.LE_REFL" + "LTE_ADD2" > "Groups.add_mono_thms_linordered_field_3" + "LE_TRANS" > "Nat.le_trans" + "LE_SUC_LT" > "Nat.Suc_le_eq" + "LE_SUC" > "Nat.Suc_le_mono" + "LE_SQUARE_REFL" > "Nat.le_square" + "LE_REFL" > "Nat.le_refl" "LE_RDIV_EQ" > "HOLLight.hollight.LE_RDIV_EQ" "LE_MULT_RCANCEL" > "HOLLight.hollight.LE_MULT_RCANCEL" "LE_MULT_LCANCEL" > "HOLLight.hollight.LE_MULT_LCANCEL" - "LE_MULT2" > "HOLLight.hollight.LE_MULT2" - "LE_LT" > "HOLLight.hollight.LE_LT" + "LE_MULT2" > "Nat.mult_le_mono" + "LE_LT" > "Nat.le_eq_less_or_eq" "LE_LDIV_EQ" > "HOLLight.hollight.LE_LDIV_EQ" "LE_LDIV" > "HOLLight.hollight.LE_LDIV" "LE_EXP" > "HOLLight.hollight.LE_EXP" - "LE_EXISTS" > "HOLLight.hollight.LE_EXISTS" - "LE_CASES" > "HOLLight.hollight.LE_CASES" - "LE_ANTISYM" > "HOLLight.hollight.LE_ANTISYM" - "LE_ADD_RCANCEL" > "HOLLight.hollight.LE_ADD_RCANCEL" - "LE_ADD_LCANCEL" > "HOLLight.hollight.LE_ADD_LCANCEL" - "LE_ADDR" > "HOLLight.hollight.LE_ADDR" - "LE_ADD2" > "HOLLight.hollight.LE_ADD2" - "LE_ADD" > "HOLLight.hollight.LE_ADD" - "LE_0" > "HOLLight.hollight.LE_0" - "LET_TRANS" > "HOLLight.hollight.LET_TRANS" + "LE_EXISTS" > "Nat.le_iff_add" + "LE_CASES" > "Nat.nat_le_linear" + "LE_C" > "HOLLight.hollight.LE_C" + "LE_ANTISYM" > "Orderings.order_class.eq_iff" + "LE_ADD_RCANCEL" > "Groups.ordered_ab_semigroup_add_imp_le_class.add_le_cancel_right" + "LE_ADD_LCANCEL" > "Groups.ordered_ab_semigroup_add_imp_le_class.add_le_cancel_left" + "LE_ADDR" > "Nat.le_add2" + "LE_ADD2" > "Groups.add_mono_thms_linordered_semiring_1" + "LE_ADD" > "Nat.le_add1" + "LE_1" > "HOLLight.hollight.LE_1" + "LE_0" > "Nat.le0" + "LET_TRANS" > "Orderings.order_le_less_trans" "LET_END_def" > "HOLLight.hollight.LET_END_def" - "LET_CASES" > "HOLLight.hollight.LET_CASES" + "LET_CASES" > "Orderings.linorder_class.le_less_linear" "LET_ANTISYM" > "HOLLight.hollight.LET_ANTISYM" - "LET_ADD2" > "HOLLight.hollight.LET_ADD2" - "LENGTH_def" > "HOLLight.hollight.LENGTH_def" - "LENGTH_REPLICATE" > "HOLLight.hollight.LENGTH_REPLICATE" - "LENGTH_MAP2" > "HOLLight.hollight.LENGTH_MAP2" - "LENGTH_MAP" > "HOLLight.hollight.LENGTH_MAP" - "LENGTH_LIST_OF_SET" > "HOLLight.hollight.LENGTH_LIST_OF_SET" - "LENGTH_EQ_NIL" > "HOLLight.hollight.LENGTH_EQ_NIL" - "LENGTH_EQ_CONS" > "HOLLight.hollight.LENGTH_EQ_CONS" - "LENGTH_APPEND" > "HOLLight.hollight.LENGTH_APPEND" - "LEFT_SUB_DISTRIB" > "Nat.nat_distrib_4" + "LET_ADD2" > "Groups.add_mono_thms_linordered_field_4" + "LENGTH_TL" > "HOLLightList.LENGTH_TL" + "LENGTH_REPLICATE" > "List.length_replicate" + "LENGTH_MAP2" > "HOLLightList.LENGTH_MAP2" + "LENGTH_MAP" > "List.length_map" + "LENGTH_EQ_NIL" > "List.length_0_conv" + "LENGTH_EQ_CONS" > "List.length_Suc_conv" + "LENGTH_APPEND" > "List.length_append" + "LEFT_SUB_DISTRIB" > "Nat.diff_mult_distrib2" "LEFT_OR_FORALL_THM" > "HOL.all_simps_3" "LEFT_OR_EXISTS_THM" > "HOL.ex_simps_3" - "LEFT_OR_DISTRIB" > "HOL.conj_disj_distribL" - "LEFT_IMP_FORALL_THM" > "HOL.imp_all" - "LEFT_IMP_EXISTS_THM" > "HOL.imp_ex" + "LEFT_OR_DISTRIB" > "Groebner_Basis.dnf_1" + "LEFT_IMP_FORALL_THM" > "HOL.ex_simps_5" + "LEFT_IMP_EXISTS_THM" > "HOL.all_simps_5" "LEFT_FORALL_OR_THM" > "HOL.all_simps_3" - "LEFT_FORALL_IMP_THM" > "HOL.imp_ex" - "LEFT_EXISTS_IMP_THM" > "HOL.imp_all" + "LEFT_FORALL_IMP_THM" > "HOL.all_simps_5" + "LEFT_EXISTS_IMP_THM" > "HOL.ex_simps_5" "LEFT_EXISTS_AND_THM" > "HOL.ex_simps_1" "LEFT_AND_FORALL_THM" > "HOL.all_simps_1" "LEFT_AND_EXISTS_THM" > "HOL.ex_simps_1" - "LEFT_ADD_DISTRIB" > "Nat.nat_distrib_2" - "LAST_def" > "HOLLight.hollight.LAST_def" - "LAST_CLAUSES" > "HOLLight.hollight.LAST_CLAUSES" + "LEFT_ADD_DISTRIB" > "Fields.linordered_field_class.sign_simps_25" + "LAST_EL" > "List.last_conv_nth" + "LAST_CLAUSES" > "HOLLightList.LAST_CLAUSES" + "LAST_APPEND" > "List.last_append" "LAMBDA_UNIQUE" > "HOLLight.hollight.LAMBDA_UNIQUE" + "LAMBDA_PAIR_THM" > "HOLLight.hollight.LAMBDA_PAIR_THM" "LAMBDA_ETA" > "HOLLight.hollight.LAMBDA_ETA" "LAMBDA_BETA" > "HOLLight.hollight.LAMBDA_BETA" - "I_THM" > "Fun.id_apply" + "I_THM" > "HOL.refl" "I_O_ID" > "HOLLight.hollight.I_O_ID" "ITSET_def" > "HOLLight.hollight.ITSET_def" "ITSET_EQ" > "HOLLight.hollight.ITSET_EQ" - "ITLIST_def" > "HOLLight.hollight.ITLIST_def" - "ITLIST_EXTRA" > "HOLLight.hollight.ITLIST_EXTRA" - "ITLIST2_def" > "HOLLight.hollight.ITLIST2_def" - "ITLIST2" > "HOLLight.hollight.ITLIST2" + "ITLIST_EXTRA" > "HOLLightList.ITLIST_EXTRA" + "ITLIST_APPEND" > "List.foldr_append" + "ITLIST2" > "HOLLightList.ITLIST2" + "ITERATE_UNION_NONZERO" > "HOLLight.hollight.ITERATE_UNION_NONZERO" "ITERATE_UNION_GEN" > "HOLLight.hollight.ITERATE_UNION_GEN" "ITERATE_UNION" > "HOLLight.hollight.ITERATE_UNION" "ITERATE_SUPPORT" > "HOLLight.hollight.ITERATE_SUPPORT" + "ITERATE_SUPERSET" > "HOLLight.hollight.ITERATE_SUPERSET" "ITERATE_SING" > "HOLLight.hollight.ITERATE_SING" "ITERATE_RELATED" > "HOLLight.hollight.ITERATE_RELATED" + "ITERATE_PAIR" > "HOLLight.hollight.ITERATE_PAIR" + "ITERATE_OP_GEN" > "HOLLight.hollight.ITERATE_OP_GEN" + "ITERATE_OP" > "HOLLight.hollight.ITERATE_OP" "ITERATE_ITERATE_PRODUCT" > "HOLLight.hollight.ITERATE_ITERATE_PRODUCT" + "ITERATE_INJECTION" > "HOLLight.hollight.ITERATE_INJECTION" + "ITERATE_INCL_EXCL" > "HOLLight.hollight.ITERATE_INCL_EXCL" + "ITERATE_IMAGE_NONZERO" > "HOLLight.hollight.ITERATE_IMAGE_NONZERO" "ITERATE_IMAGE" > "HOLLight.hollight.ITERATE_IMAGE" + "ITERATE_EXPAND_CASES" > "HOLLight.hollight.ITERATE_EXPAND_CASES" "ITERATE_EQ_NEUTRAL" > "HOLLight.hollight.ITERATE_EQ_NEUTRAL" + "ITERATE_EQ_GENERAL_INVERSES" > "HOLLight.hollight.ITERATE_EQ_GENERAL_INVERSES" "ITERATE_EQ_GENERAL" > "HOLLight.hollight.ITERATE_EQ_GENERAL" "ITERATE_EQ" > "HOLLight.hollight.ITERATE_EQ" "ITERATE_DIFF_GEN" > "HOLLight.hollight.ITERATE_DIFF_GEN" "ITERATE_DIFF" > "HOLLight.hollight.ITERATE_DIFF" "ITERATE_DELTA" > "HOLLight.hollight.ITERATE_DELTA" - "ITERATE_CLOSED_GEN" > "HOLLight.hollight.ITERATE_CLOSED_GEN" + "ITERATE_DELETE" > "HOLLight.hollight.ITERATE_DELETE" "ITERATE_CLOSED" > "HOLLight.hollight.ITERATE_CLOSED" + "ITERATE_CLAUSES_NUMSEG" > "HOLLight.hollight.ITERATE_CLAUSES_NUMSEG" "ITERATE_CLAUSES_GEN" > "HOLLight.hollight.ITERATE_CLAUSES_GEN" "ITERATE_CLAUSES" > "HOLLight.hollight.ITERATE_CLAUSES" + "ITERATE_CASES" > "HOLLight.hollight.ITERATE_CASES" "ITERATE_BIJECTION" > "HOLLight.hollight.ITERATE_BIJECTION" "ISO_def" > "HOLLight.hollight.ISO_def" "ISO_USAGE" > "HOLLight.hollight.ISO_USAGE" "ISO_REFL" > "HOLLight.hollight.ISO_REFL" "ISO_FUN" > "HOLLight.hollight.ISO_FUN" - "IN_def" > "HOLLight.hollight.IN_def" - "IN_UNIV" > "HOLLight.hollight.IN_UNIV" + "IN_UNIV" > "Set.UNIV_I" "IN_UNIONS" > "HOLLight.hollight.IN_UNIONS" - "IN_UNION" > "HOLLight.hollight.IN_UNION" + "IN_UNION" > "Complete_Lattice.mem_simps_3" "IN_SUPPORT" > "HOLLight.hollight.IN_SUPPORT" - "IN_SING" > "HOLLight.hollight.IN_SING" - "IN_SET_OF_LIST" > "HOLLight.hollight.IN_SET_OF_LIST" + "IN_SING" > "Set.singleton_iff" "IN_REST" > "HOLLight.hollight.IN_REST" - "IN_NUMSEG" > "HOLLight.hollight.IN_NUMSEG" + "IN_NUMSEG_0" > "HOLLight.hollight.IN_NUMSEG_0" + "IN_NUMSEG" > "SetInterval.ord_class.atLeastAtMost_iff" "IN_INTERS" > "HOLLight.hollight.IN_INTERS" - "IN_INTER" > "HOLLight.hollight.IN_INTER" - "IN_INSERT" > "HOLLight.hollight.IN_INSERT" + "IN_INTER" > "Complete_Lattice.mem_simps_4" + "IN_INSERT" > "Complete_Lattice.mem_simps_1" "IN_IMAGE" > "HOLLight.hollight.IN_IMAGE" "IN_ELIM_THM" > "HOLLight.hollight.IN_ELIM_THM" "IN_ELIM_PAIR_THM" > "HOLLight.hollight.IN_ELIM_PAIR_THM" "IN_DISJOINT" > "HOLLight.hollight.IN_DISJOINT" - "IN_DIFF" > "HOLLight.hollight.IN_DIFF" + "IN_DIFF" > "Complete_Lattice.mem_simps_6" "IN_DELETE_EQ" > "HOLLight.hollight.IN_DELETE_EQ" "IN_DELETE" > "HOLLight.hollight.IN_DELETE" + "IN_CROSS" > "HOLLight.hollight.IN_CROSS" + "INT_WOP" > "HOLLight.hollight.INT_WOP" "INT_POW" > "HOLLight.hollight.INT_POW" + "INT_OF_NUM_OF_INT" > "HOLLight.hollight.INT_OF_NUM_OF_INT" "INT_LT_DISCRETE" > "HOLLight.hollight.INT_LT_DISCRETE" "INT_LT" > "HOLLight.hollight.INT_LT" + "INT_INTEGRAL" > "HOLLight.hollight.INT_INTEGRAL" "INT_IMAGE" > "HOLLight.hollight.INT_IMAGE" "INT_GT_DISCRETE" > "HOLLight.hollight.INT_GT_DISCRETE" "INT_GT" > "HOLLight.hollight.INT_GT" "INT_GE" > "HOLLight.hollight.INT_GE" + "INT_GCD_EXISTS_POS" > "HOLLight.hollight.INT_GCD_EXISTS_POS" + "INT_GCD_EXISTS" > "HOLLight.hollight.INT_GCD_EXISTS" "INT_FORALL_POS" > "HOLLight.hollight.INT_FORALL_POS" + "INT_FORALL_ABS" > "HOLLight.hollight.INT_FORALL_ABS" + "INT_EXISTS_POS" > "HOLLight.hollight.INT_EXISTS_POS" + "INT_EXISTS_ABS" > "HOLLight.hollight.INT_EXISTS_ABS" + "INT_DIVMOD_UNIQ" > "HOLLight.hollight.INT_DIVMOD_UNIQ" + "INT_DIVMOD_EXIST_0" > "HOLLight.hollight.INT_DIVMOD_EXIST_0" + "INT_DIVISION" > "HOLLight.hollight.INT_DIVISION" "INT_ARCH" > "HOLLight.hollight.INT_ARCH" "INT_ABS_MUL_1" > "HOLLight.hollight.INT_ABS_MUL_1" "INT_ABS" > "HOLLight.hollight.INT_ABS" - "INTER_def" > "HOLLight.hollight.INTER_def" "INTER_UNIV" > "HOLLight.hollight.INTER_UNIV" + "INTER_UNIONS" > "HOLLight.hollight.INTER_UNIONS" "INTER_SUBSET" > "HOLLight.hollight.INTER_SUBSET" - "INTER_OVER_UNION" > "HOLLight.hollight.INTER_OVER_UNION" - "INTER_IDEMPOT" > "HOLLight.hollight.INTER_IDEMPOT" + "INTER_OVER_UNION" > "Lattices.distrib_lattice_class.distrib_1" + "INTER_IDEMPOT" > "Big_Operators.lattice_class.Inf_fin.idem" "INTER_EMPTY" > "HOLLight.hollight.INTER_EMPTY" - "INTER_COMM" > "HOLLight.hollight.INTER_COMM" - "INTER_ASSOC" > "HOLLight.hollight.INTER_ASSOC" + "INTER_COMM" > "Lattices.lattice_class.inf_sup_aci_1" + "INTER_ASSOC" > "Lattices.lattice_class.inf_sup_aci_2" "INTER_ACI" > "HOLLight.hollight.INTER_ACI" - "INTERS_def" > "HOLLight.hollight.INTERS_def" - "INTERS_INSERT" > "HOLLight.hollight.INTERS_INSERT" - "INTERS_2" > "HOLLight.hollight.INTERS_2" - "INTERS_1" > "HOLLight.hollight.INTERS_1" - "INTERS_0" > "HOLLight.hollight.INTERS_0" - "INSERT_def" > "HOLLight.hollight.INSERT_def" + "INTERS_UNIONS" > "HOLLight.hollight.INTERS_UNIONS" + "INTERS_INSERT" > "Complete_Lattice.Inter_insert" + "INTERS_IMAGE" > "HOLLight.hollight.INTERS_IMAGE" + "INTERS_GSPEC" > "HOLLight.hollight.INTERS_GSPEC" + "INTERS_2" > "Complete_Lattice.Int_eq_Inter" + "INTERS_1" > "Complete_Lattice.complete_lattice_class.Inf_singleton" + "INTERS_0" > "Complete_Lattice.Inter_empty" "INSERT_UNIV" > "HOLLight.hollight.INSERT_UNIV" - "INSERT_UNION_EQ" > "HOLLight.hollight.INSERT_UNION_EQ" + "INSERT_UNION_EQ" > "Set.Un_insert_left" "INSERT_UNION" > "HOLLight.hollight.INSERT_UNION" - "INSERT_SUBSET" > "HOLLight.hollight.INSERT_SUBSET" - "INSERT_INTER" > "HOLLight.hollight.INSERT_INTER" - "INSERT_INSERT" > "HOLLight.hollight.INSERT_INSERT" - "INSERT_DIFF" > "HOLLight.hollight.INSERT_DIFF" - "INSERT_DELETE" > "HOLLight.hollight.INSERT_DELETE" - "INSERT_COMM" > "HOLLight.hollight.INSERT_COMM" + "INSERT_SUBSET" > "Set.insert_subset" + "INSERT_INTER" > "Set.Int_insert_left" + "INSERT_INSERT" > "Set.insert_absorb2" + "INSERT_DIFF" > "Set.insert_Diff_if" + "INSERT_DELETE" > "Set.insert_Diff" + "INSERT_COMM" > "Set.insert_commute" "INSERT_AC" > "HOLLight.hollight.INSERT_AC" "INSERT" > "HOLLight.hollight.INSERT" "INJ_def" > "HOLLight.hollight.INJ_def" @@ -1339,26 +1619,36 @@ "INJF_def" > "HOLLight.hollight.INJF_def" "INJF_INJ" > "HOLLight.hollight.INJF_INJ" "INJECTIVE_ON_LEFT_INVERSE" > "HOLLight.hollight.INJECTIVE_ON_LEFT_INVERSE" + "INJECTIVE_ON_IMAGE" > "HOLLight.hollight.INJECTIVE_ON_IMAGE" + "INJECTIVE_MAP" > "HOLLightList.INJECTIVE_MAP" "INJECTIVE_LEFT_INVERSE" > "HOLLight.hollight.INJECTIVE_LEFT_INVERSE" + "INJECTIVE_IMAGE" > "HOLLight.hollight.INJECTIVE_IMAGE" "INJA_def" > "HOLLight.hollight.INJA_def" "INJA_INJ" > "HOLLight.hollight.INJA_INJ" - "INFINITE_def" > "HOLLight.hollight.INFINITE_def" - "INFINITE_NONEMPTY" > "HOLLight.hollight.INFINITE_NONEMPTY" + "INFINITE_NONEMPTY" > "Infinite_Set.infinite_imp_nonempty" "INFINITE_IMAGE_INJ" > "HOLLight.hollight.INFINITE_IMAGE_INJ" - "INFINITE_DIFF_FINITE" > "HOLLight.hollight.INFINITE_DIFF_FINITE" + "INFINITE_DIFF_FINITE" > "Infinite_Set.Diff_infinite_finite" + "IND_SUC_def" > "HOLLight.hollight.IND_SUC_def" + "IND_SUC_0_EXISTS" > "HOLLight.hollight.IND_SUC_0_EXISTS" + "IND_0_def" > "HOLLight.hollight.IND_0_def" "IMP_EQ_CLAUSE" > "HOLLight.hollight.IMP_EQ_CLAUSE" + "IMP_CONJ_ALT" > "HOLLight.hollight.IMP_CONJ_ALT" "IMP_CONJ" > "HOL.imp_conjL" "IMP_CLAUSES" > "HOLLight.hollight.IMP_CLAUSES" - "IMAGE_o" > "HOLLight.hollight.IMAGE_o" - "IMAGE_def" > "HOLLight.hollight.IMAGE_def" - "IMAGE_UNION" > "HOLLight.hollight.IMAGE_UNION" - "IMAGE_SUBSET" > "HOLLight.hollight.IMAGE_SUBSET" + "IMAGE_o" > "Fun.image_compose" + "IMAGE_UNIONS" > "HOLLight.hollight.IMAGE_UNIONS" + "IMAGE_UNION" > "Set.image_Un" + "IMAGE_SUBSET" > "Set.image_mono" + "IMAGE_INTER_INJ" > "HOLLight.hollight.IMAGE_INTER_INJ" + "IMAGE_INJECTIVE_IMAGE_OF_SUBSET" > "HOLLight.hollight.IMAGE_INJECTIVE_IMAGE_OF_SUBSET" "IMAGE_IMP_INJECTIVE_GEN" > "HOLLight.hollight.IMAGE_IMP_INJECTIVE_GEN" "IMAGE_IMP_INJECTIVE" > "HOLLight.hollight.IMAGE_IMP_INJECTIVE" - "IMAGE_EQ_EMPTY" > "HOLLight.hollight.IMAGE_EQ_EMPTY" + "IMAGE_ID" > "Fun.image_ident" + "IMAGE_I" > "Fun.image_id" + "IMAGE_EQ_EMPTY" > "Set.image_is_empty" "IMAGE_DIFF_INJ" > "HOLLight.hollight.IMAGE_DIFF_INJ" "IMAGE_DELETE_INJ" > "HOLLight.hollight.IMAGE_DELETE_INJ" - "IMAGE_CONST" > "HOLLight.hollight.IMAGE_CONST" + "IMAGE_CONST" > "Set.image_constant_conv" "IMAGE_CLAUSES" > "HOLLight.hollight.IMAGE_CLAUSES" "HREAL_MUL_RZERO" > "HOLLight.hollight.HREAL_MUL_RZERO" "HREAL_MUL_LZERO" > "HOLLight.hollight.HREAL_MUL_LZERO" @@ -1372,7 +1662,7 @@ "HREAL_ADD_RID" > "HOLLight.hollight.HREAL_ADD_RID" "HREAL_ADD_RDISTRIB" > "HOLLight.hollight.HREAL_ADD_RDISTRIB" "HREAL_ADD_AC" > "HOLLight.hollight.HREAL_ADD_AC" - "HD_def" > "HOLLight.hollight.HD_def" + "HD_APPEND" > "List.hd_append" "HAS_SIZE_def" > "HOLLight.hollight.HAS_SIZE_def" "HAS_SIZE_UNIONS" > "HOLLight.hollight.HAS_SIZE_UNIONS" "HAS_SIZE_UNION" > "HOLLight.hollight.HAS_SIZE_UNION" @@ -1385,31 +1675,40 @@ "HAS_SIZE_NUMSEG_1" > "HOLLight.hollight.HAS_SIZE_NUMSEG_1" "HAS_SIZE_NUMSEG" > "HOLLight.hollight.HAS_SIZE_NUMSEG" "HAS_SIZE_INDEX" > "HOLLight.hollight.HAS_SIZE_INDEX" + "HAS_SIZE_IMAGE_INJ_EQ" > "HOLLight.hollight.HAS_SIZE_IMAGE_INJ_EQ" "HAS_SIZE_IMAGE_INJ" > "HOLLight.hollight.HAS_SIZE_IMAGE_INJ" "HAS_SIZE_FUNSPACE" > "HOLLight.hollight.HAS_SIZE_FUNSPACE" "HAS_SIZE_FINITE_IMAGE" > "HOLLight.hollight.HAS_SIZE_FINITE_IMAGE" + "HAS_SIZE_DIFF" > "HOLLight.hollight.HAS_SIZE_DIFF" + "HAS_SIZE_CROSS" > "HOLLight.hollight.HAS_SIZE_CROSS" "HAS_SIZE_CLAUSES" > "HOLLight.hollight.HAS_SIZE_CLAUSES" "HAS_SIZE_CARD" > "HOLLight.hollight.HAS_SIZE_CARD" + "HAS_SIZE_1" > "HOLLight.hollight.HAS_SIZE_1" "HAS_SIZE_0" > "HOLLight.hollight.HAS_SIZE_0" - "GSPEC_def" > "HOLLight.hollight.GSPEC_def" - "GEQ_def" > "HOLLight.hollight.GEQ_def" - "GABS_def" > "HOLLight.hollight.GABS_def" + "GE_C" > "HOLLight.hollight.GE_C" + "FUN_IN_IMAGE" > "Set.imageI" "FUN_EQ_THM" > "Fun.fun_eq_iff" "FUNCTION_FACTORS_RIGHT" > "HOLLight.hollight.FUNCTION_FACTORS_RIGHT" "FUNCTION_FACTORS_LEFT" > "HOLLight.hollight.FUNCTION_FACTORS_LEFT" - "FSTCART_PASTECART" > "HOLLight.hollight.FSTCART_PASTECART" "FST" > "Product_Type.fst_conv" + "FORALL_UNWIND_THM2" > "HOL.simp_thms_41" + "FORALL_UNWIND_THM1" > "HOL.simp_thms_42" + "FORALL_UNCURRY" > "HOLLight.hollight.FORALL_UNCURRY" + "FORALL_TRIPLED_THM" > "HOLLight.hollight.FORALL_TRIPLED_THM" + "FORALL_SUBSET_IMAGE" > "HOLLight.hollight.FORALL_SUBSET_IMAGE" "FORALL_SIMP" > "HOL.simp_thms_35" - "FORALL_PASTECART" > "HOLLight.hollight.FORALL_PASTECART" "FORALL_PAIR_THM" > "Product_Type.split_paired_All" - "FORALL_NOT_THM" > "Inductive.basic_monos_16" + "FORALL_PAIRED_THM" > "HOLLight.hollight.FORALL_PAIRED_THM" + "FORALL_NOT_THM" > "HOL.not_ex" "FORALL_IN_UNIONS" > "HOLLight.hollight.FORALL_IN_UNIONS" + "FORALL_IN_INSERT" > "HOLLight.hollight.FORALL_IN_INSERT" "FORALL_IN_IMAGE" > "HOLLight.hollight.FORALL_IN_IMAGE" + "FORALL_IN_GSPEC" > "HOLLight.hollight.FORALL_IN_GSPEC" "FORALL_IN_CLAUSES" > "HOLLight.hollight.FORALL_IN_CLAUSES" "FORALL_FINITE_INDEX" > "HOLLight.hollight.FORALL_FINITE_INDEX" "FORALL_BOOL_THM" > "Set.all_bool_eq" "FORALL_AND_THM" > "HOL.all_conj_distrib" - "FORALL_ALL" > "HOLLight.hollight.FORALL_ALL" + "FORALL_ALL" > "HOLLightList.FORALL_ALL" "FNIL_def" > "HOLLight.hollight.FNIL_def" "FINREC_def" > "HOLLight.hollight.FINREC_def" "FINREC_UNIQUE_LEMMA" > "HOLLight.hollight.FINREC_UNIQUE_LEMMA" @@ -1418,139 +1717,153 @@ "FINREC_FUN" > "HOLLight.hollight.FINREC_FUN" "FINREC_EXISTS_LEMMA" > "HOLLight.hollight.FINREC_EXISTS_LEMMA" "FINREC_1_LEMMA" > "HOLLight.hollight.FINREC_1_LEMMA" - "FINITE_def" > "HOLLight.hollight.FINITE_def" - "FINITE_UNION_IMP" > "HOLLight.hollight.FINITE_UNION_IMP" + "FINITE_UNION_IMP" > "Finite_Set.finite_UnI" "FINITE_UNIONS" > "HOLLight.hollight.FINITE_UNIONS" - "FINITE_UNION" > "HOLLight.hollight.FINITE_UNION" + "FINITE_UNION" > "Finite_Set.finite_Un" "FINITE_SUPPORT_DELTA" > "HOLLight.hollight.FINITE_SUPPORT_DELTA" "FINITE_SUPPORT" > "HOLLight.hollight.FINITE_SUPPORT" + "FINITE_SUM_IMAGE" > "HOLLight.hollight.FINITE_SUM_IMAGE" "FINITE_SUBSET_IMAGE_IMP" > "HOLLight.hollight.FINITE_SUBSET_IMAGE_IMP" "FINITE_SUBSET_IMAGE" > "HOLLight.hollight.FINITE_SUBSET_IMAGE" - "FINITE_SUBSETS" > "HOLLight.hollight.FINITE_SUBSETS" - "FINITE_SUBSET" > "HOLLight.hollight.FINITE_SUBSET" - "FINITE_SET_OF_LIST" > "HOLLight.hollight.FINITE_SET_OF_LIST" + "FINITE_SUBSET" > "Finite_Set.finite_subset" + "FINITE_SING" > "HOLLight.hollight.FINITE_SING" "FINITE_RESTRICT" > "HOLLight.hollight.FINITE_RESTRICT" "FINITE_RECURSION_DELETE" > "HOLLight.hollight.FINITE_RECURSION_DELETE" "FINITE_RECURSION" > "HOLLight.hollight.FINITE_RECURSION" + "FINITE_REAL_INTERVAL" > "HOLLight.hollight.FINITE_REAL_INTERVAL" "FINITE_PRODUCT_DEPENDENT" > "HOLLight.hollight.FINITE_PRODUCT_DEPENDENT" "FINITE_PRODUCT" > "HOLLight.hollight.FINITE_PRODUCT" "FINITE_POWERSET" > "HOLLight.hollight.FINITE_POWERSET" "FINITE_NUMSEG_LT" > "HOLLight.hollight.FINITE_NUMSEG_LT" "FINITE_NUMSEG_LE" > "HOLLight.hollight.FINITE_NUMSEG_LE" - "FINITE_NUMSEG" > "HOLLight.hollight.FINITE_NUMSEG" - "FINITE_INTER" > "HOLLight.hollight.FINITE_INTER" - "FINITE_INSERT" > "HOLLight.hollight.FINITE_INSERT" - "FINITE_INDUCT_STRONG" > "HOLLight.hollight.FINITE_INDUCT_STRONG" - "FINITE_INDEX_WORKS_FINITE" > "HOLLight.hollight.FINITE_INDEX_WORKS_FINITE" + "FINITE_NUMSEG" > "SetInterval.finite_atLeastAtMost" + "FINITE_INTSEG" > "HOLLight.hollight.FINITE_INTSEG" + "FINITE_INTER" > "Finite_Set.finite_Int" + "FINITE_INSERT" > "Finite_Set.finite_insert" + "FINITE_INDUCT_STRONG" > "Finite_Set.finite_induct" + "FINITE_INDUCT_DELETE" > "HOLLight.hollight.FINITE_INDUCT_DELETE" "FINITE_INDEX_WORKS" > "HOLLight.hollight.FINITE_INDEX_WORKS" "FINITE_INDEX_NUMSEG" > "HOLLight.hollight.FINITE_INDEX_NUMSEG" "FINITE_INDEX_NUMBERS" > "HOLLight.hollight.FINITE_INDEX_NUMBERS" + "FINITE_INDEX_INRANGE" > "HOLLight.hollight.FINITE_INDEX_INRANGE" "FINITE_INDEX_INJ" > "HOLLight.hollight.FINITE_INDEX_INJ" "FINITE_IMAGE_INJ_GENERAL" > "HOLLight.hollight.FINITE_IMAGE_INJ_GENERAL" + "FINITE_IMAGE_INJ_EQ" > "HOLLight.hollight.FINITE_IMAGE_INJ_EQ" "FINITE_IMAGE_INJ" > "HOLLight.hollight.FINITE_IMAGE_INJ" "FINITE_IMAGE_IMAGE" > "HOLLight.hollight.FINITE_IMAGE_IMAGE" "FINITE_IMAGE_EXPAND" > "HOLLight.hollight.FINITE_IMAGE_EXPAND" - "FINITE_IMAGE" > "HOLLight.hollight.FINITE_IMAGE" - "FINITE_HAS_SIZE_LEMMA" > "HOLLight.hollight.FINITE_HAS_SIZE_LEMMA" + "FINITE_IMAGE" > "Finite_Set.finite_imageI" + "FINITE_HAS_SIZE" > "HOLLight.hollight.FINITE_HAS_SIZE" "FINITE_FUNSPACE" > "HOLLight.hollight.FINITE_FUNSPACE" + "FINITE_FINITE_UNIONS" > "HOLLight.hollight.FINITE_FINITE_UNIONS" + "FINITE_FINITE_PREIMAGE_GENERAL" > "HOLLight.hollight.FINITE_FINITE_PREIMAGE_GENERAL" + "FINITE_FINITE_PREIMAGE" > "HOLLight.hollight.FINITE_FINITE_PREIMAGE" "FINITE_FINITE_IMAGE" > "HOLLight.hollight.FINITE_FINITE_IMAGE" - "FINITE_DIFF" > "HOLLight.hollight.FINITE_DIFF" + "FINITE_EMPTY" > "Finite_Set.finite.emptyI" + "FINITE_DIFF" > "Finite_Set.finite_Diff" "FINITE_DELETE_IMP" > "HOLLight.hollight.FINITE_DELETE_IMP" "FINITE_DELETE" > "HOLLight.hollight.FINITE_DELETE" - "FILTER_def" > "HOLLight.hollight.FILTER_def" - "FILTER_MAP" > "HOLLight.hollight.FILTER_MAP" - "FILTER_APPEND" > "HOLLight.hollight.FILTER_APPEND" + "FINITE_CROSS" > "HOLLight.hollight.FINITE_CROSS" + "FINITE_CART" > "HOLLight.hollight.FINITE_CART" + "FILTER_MAP" > "List.filter_map" + "FILTER_APPEND" > "List.filter_append" "FCONS_def" > "HOLLight.hollight.FCONS_def" "FCONS_UNDO" > "HOLLight.hollight.FCONS_UNDO" - "FACT_def" > "HOLLight.hollight.FACT_def" - "FACT_MONO" > "HOLLight.hollight.FACT_MONO" - "FACT_LT" > "HOLLight.hollight.FACT_LT" - "FACT_LE" > "HOLLight.hollight.FACT_LE" - "EX_def" > "HOLLight.hollight.EX_def" - "EX_MEM" > "HOLLight.hollight.EX_MEM" - "EX_MAP" > "HOLLight.hollight.EX_MAP" - "EX_IMP" > "HOLLight.hollight.EX_IMP" - "EXTENSION" > "HOLLight.hollight.EXTENSION" - "EXP_def" > "HOLLight.hollight.EXP_def" - "EXP_ONE" > "HOLLight.hollight.EXP_ONE" - "EXP_MULT" > "HOLLight.hollight.EXP_MULT" + "FACT_NZ" > "Fact.fact_nonzero_nat" + "FACT_MONO" > "Fact.fact_mono_nat" + "FACT_LT" > "Fact.fact_gt_zero_nat" + "FACT_LE" > "Fact.fact_ge_one_nat" + "EX_MEM" > "HOLLightList.EX_MEM" + "EX_IMP" > "HOLLightList.EX_IMP" + "EXTENSION" > "Set.set_eq_iff" + "EXP_ZERO" > "Power.power_0_left" + "EXP_ONE" > "Power.monoid_mult_class.power_one" + "EXP_MULT" > "Power.monoid_mult_class.power_mult" + "EXP_MONO_LT_IMP" > "HOLLight.hollight.EXP_MONO_LT_IMP" + "EXP_MONO_LT" > "HOLLight.hollight.EXP_MONO_LT" + "EXP_MONO_LE_IMP" > "HOLLight.hollight.EXP_MONO_LE_IMP" + "EXP_MONO_LE" > "HOLLight.hollight.EXP_MONO_LE" + "EXP_MONO_EQ" > "HOLLight.hollight.EXP_MONO_EQ" "EXP_LT_0" > "HOLLight.hollight.EXP_LT_0" - "EXP_EQ_0" > "HOLLight.hollight.EXP_EQ_0" - "EXP_ADD" > "HOLLight.hollight.EXP_ADD" - "EXP_2" > "HOLLight.hollight.EXP_2" - "EXP_1" > "HOLLight.hollight.EXP_1" - "EXISTS_UNIQUE_THM" > "HOL4Compat.EXISTS_UNIQUE_DEF" + "EXP_EQ_1" > "HOLLight.hollight.EXP_EQ_1" + "EXP_EQ_0" > "Power.power_eq_0_iff" + "EXP_ADD" > "Power.monoid_mult_class.power_add" + "EXP_2" > "Nat_Numeral.monoid_mult_class.power2_eq_square" + "EXP_1" > "Power.monoid_mult_class.power_one_right" + "EXISTS_UNIQUE_THM" > "HOLLightCompat.EXISTS_UNIQUE_THM" "EXISTS_UNIQUE_REFL" > "HOL.ex1_eq_1" "EXISTS_UNIQUE_ALT" > "HOLLight.hollight.EXISTS_UNIQUE_ALT" "EXISTS_UNIQUE" > "HOL.Ex1_def" + "EXISTS_UNCURRY" > "HOLLight.hollight.EXISTS_UNCURRY" + "EXISTS_TRIPLED_THM" > "HOLLight.hollight.EXISTS_TRIPLED_THM" "EXISTS_THM" > "HOL4Setup.EXISTS_DEF" + "EXISTS_SUBSET_IMAGE" > "HOLLight.hollight.EXISTS_SUBSET_IMAGE" "EXISTS_SIMP" > "HOL.simp_thms_36" "EXISTS_REFL" > "HOL.simp_thms_37" - "EXISTS_PASTECART" > "HOLLight.hollight.EXISTS_PASTECART" "EXISTS_PAIR_THM" > "Product_Type.split_paired_Ex" + "EXISTS_PAIRED_THM" > "HOLLight.hollight.EXISTS_PAIRED_THM" "EXISTS_OR_THM" > "HOL.ex_disj_distrib" "EXISTS_ONE_REP" > "HOLLight.hollight.EXISTS_ONE_REP" - "EXISTS_NOT_THM" > "Inductive.basic_monos_15" + "EXISTS_NOT_THM" > "HOL.not_all" + "EXISTS_IN_UNIONS" > "HOLLight.hollight.EXISTS_IN_UNIONS" + "EXISTS_IN_INSERT" > "HOLLight.hollight.EXISTS_IN_INSERT" "EXISTS_IN_IMAGE" > "HOLLight.hollight.EXISTS_IN_IMAGE" + "EXISTS_IN_GSPEC" > "HOLLight.hollight.EXISTS_IN_GSPEC" "EXISTS_IN_CLAUSES" > "HOLLight.hollight.EXISTS_IN_CLAUSES" - "EXISTS_EX" > "HOLLight.hollight.EXISTS_EX" + "EXISTS_FINITE_SUBSET_IMAGE" > "HOLLight.hollight.EXISTS_FINITE_SUBSET_IMAGE" + "EXISTS_EX" > "HOLLightList.EXISTS_EX" "EXISTS_BOOL_THM" > "Set.ex_bool_eq" "EXCLUDED_MIDDLE" > "HOLLight.hollight.EXCLUDED_MIDDLE" - "EVEN_def" > "HOLLight.hollight.EVEN_def" + "EVEN_SUB" > "HOLLight.hollight.EVEN_SUB" "EVEN_OR_ODD" > "HOLLight.hollight.EVEN_OR_ODD" "EVEN_ODD_DECOMPOSITION" > "HOLLight.hollight.EVEN_ODD_DECOMPOSITION" - "EVEN_MULT" > "HOLLight.hollight.EVEN_MULT" + "EVEN_MULT" > "Parity.even_product_nat" "EVEN_MOD" > "HOLLight.hollight.EVEN_MOD" "EVEN_EXP" > "HOLLight.hollight.EVEN_EXP" "EVEN_EXISTS_LEMMA" > "HOLLight.hollight.EVEN_EXISTS_LEMMA" - "EVEN_EXISTS" > "HOLLight.hollight.EVEN_EXISTS" + "EVEN_EXISTS" > "Parity.even_mult_two_ex" "EVEN_DOUBLE" > "HOLLight.hollight.EVEN_DOUBLE" "EVEN_AND_ODD" > "HOLLight.hollight.EVEN_AND_ODD" - "EVEN_ADD" > "HOLLight.hollight.EVEN_ADD" + "EVEN_ADD" > "Parity.even_add" "EQ_UNIV" > "HOLLight.hollight.EQ_UNIV" - "EQ_TRANS" > "Set.basic_trans_rules_31" - "EQ_SYM_EQ" > "HOL.eq_sym_conv" - "EQ_SYM" > "HOL.meta_eq_to_obj_eq" - "EQ_SUC" > "Nat.nat.simps_3" - "EQ_REFL_T" > "HOL.simp_thms_6" - "EQ_REFL" > "Presburger.fm_modd_pinf" + "EQ_TRANS" > "HOL.trans" + "EQ_SYM_EQ" > "HOL.eq_ac_1" + "EQ_SYM" > "HOL.eq_reflection" + "EQ_REFL" > "HOL.refl" "EQ_MULT_RCANCEL" > "Nat.mult_cancel2" - "EQ_MULT_LCANCEL" > "Nat_Numeral.nat_mult_eq_cancel_disj" - "EQ_IMP_LE" > "HOLLight.hollight.EQ_IMP_LE" - "EQ_EXT" > "HOL.meta_eq_to_obj_eq" + "EQ_MULT_LCANCEL" > "Numeral_Simprocs.nat_mult_eq_cancel_disj" + "EQ_IMP_LE" > "Nat.eq_imp_le" + "EQ_EXT" > "HOL.eq_reflection" + "EQ_EXP" > "HOLLight.hollight.EQ_EXP" "EQ_CLAUSES" > "HOLLight.hollight.EQ_CLAUSES" "EQ_ADD_RCANCEL_0" > "HOLLight.hollight.EQ_ADD_RCANCEL_0" - "EQ_ADD_RCANCEL" > "Nat.nat_add_right_cancel" + "EQ_ADD_RCANCEL" > "Groups.cancel_semigroup_add_class.add_right_cancel" "EQ_ADD_LCANCEL_0" > "HOLLight.hollight.EQ_ADD_LCANCEL_0" - "EQ_ADD_LCANCEL" > "Nat.nat_add_left_cancel" - "EMPTY_def" > "HOLLight.hollight.EMPTY_def" + "EQ_ADD_LCANCEL" > "Groups.cancel_semigroup_add_class.add_left_cancel" "EMPTY_UNIONS" > "HOLLight.hollight.EMPTY_UNIONS" - "EMPTY_UNION" > "HOLLight.hollight.EMPTY_UNION" - "EMPTY_SUBSET" > "HOLLight.hollight.EMPTY_SUBSET" + "EMPTY_UNION" > "Lattices.bounded_lattice_bot_class.sup_eq_bot_iff" + "EMPTY_SUBSET" > "Orderings.bot_class.bot_least" "EMPTY_NOT_UNIV" > "HOLLight.hollight.EMPTY_NOT_UNIV" "EMPTY_GSPEC" > "HOLLight.hollight.EMPTY_GSPEC" - "EMPTY_DIFF" > "HOLLight.hollight.EMPTY_DIFF" + "EMPTY_DIFF" > "Set.empty_Diff" "EMPTY_DELETE" > "HOLLight.hollight.EMPTY_DELETE" - "EL_def" > "HOLLight.hollight.EL_def" - "DIV_def" > "HOLLight.hollight.DIV_def" + "EL_CONS" > "List.nth_Cons'" + "EL_APPEND" > "List.nth_append" "DIV_UNIQ" > "HOLLight.hollight.DIV_UNIQ" - "DIV_REFL" > "HOLLight.hollight.DIV_REFL" + "DIV_REFL" > "Divides.semiring_div_class.div_self" "DIV_MUL_LE" > "HOLLight.hollight.DIV_MUL_LE" "DIV_MULT2" > "HOLLight.hollight.DIV_MULT2" - "DIV_MULT" > "HOLLight.hollight.DIV_MULT" "DIV_MONO_LT" > "HOLLight.hollight.DIV_MONO_LT" "DIV_MONO2" > "HOLLight.hollight.DIV_MONO2" "DIV_MONO" > "HOLLight.hollight.DIV_MONO" "DIV_MOD" > "HOLLight.hollight.DIV_MOD" - "DIV_LT" > "HOLLight.hollight.DIV_LT" + "DIV_LT" > "Divides.div_less" "DIV_LE_EXCLUSION" > "HOLLight.hollight.DIV_LE_EXCLUSION" "DIV_LE" > "HOLLight.hollight.DIV_LE" "DIV_EQ_EXCLUSION" > "HOLLight.hollight.DIV_EQ_EXCLUSION" "DIV_EQ_0" > "HOLLight.hollight.DIV_EQ_0" "DIV_DIV" > "HOLLight.hollight.DIV_DIV" "DIV_ADD_MOD" > "HOLLight.hollight.DIV_ADD_MOD" - "DIV_1" > "HOLLight.hollight.DIV_1" - "DIV_0" > "HOLLight.hollight.DIV_0" "DIVMOD_UNIQ_LEMMA" > "HOLLight.hollight.DIVMOD_UNIQ_LEMMA" "DIVMOD_UNIQ" > "HOLLight.hollight.DIVMOD_UNIQ" "DIVMOD_EXIST_0" > "HOLLight.hollight.DIVMOD_EXIST_0" @@ -1559,7 +1872,6 @@ "DIVISION" > "HOLLight.hollight.DIVISION" "DIST_TRIANGLE_LE" > "HOLLight.hollight.DIST_TRIANGLE_LE" "DIST_TRIANGLES_LE" > "HOLLight.hollight.DIST_TRIANGLES_LE" - "DIST_TRIANGLE" > "HOLLight.hollight.DIST_TRIANGLE" "DIST_SYM" > "HOLLight.hollight.DIST_SYM" "DIST_RZERO" > "HOLLight.hollight.DIST_RZERO" "DIST_RMUL" > "HOLLight.hollight.DIST_RMUL" @@ -1573,13 +1885,9 @@ "DIST_LADD" > "HOLLight.hollight.DIST_LADD" "DIST_EQ_0" > "HOLLight.hollight.DIST_EQ_0" "DIST_ELIM_THM" > "HOLLight.hollight.DIST_ELIM_THM" - "DIST_ADDBOUND" > "HOLLight.hollight.DIST_ADDBOUND" - "DIST_ADD2_REV" > "HOLLight.hollight.DIST_ADD2_REV" - "DIST_ADD2" > "HOLLight.hollight.DIST_ADD2" - "DISJ_SYM" > "HOL.disj_comms_1" - "DISJ_ASSOC" > "Recdef.tfl_disj_assoc" + "DISJ_SYM" > "Groebner_Basis.dnf_4" + "DISJ_ASSOC" > "HOL.disj_ac_3" "DISJ_ACI" > "HOLLight.hollight.DISJ_ACI" - "DISJOINT_def" > "HOLLight.hollight.DISJOINT_def" "DISJOINT_UNION" > "HOLLight.hollight.DISJOINT_UNION" "DISJOINT_SYM" > "HOLLight.hollight.DISJOINT_SYM" "DISJOINT_NUMSEG" > "HOLLight.hollight.DISJOINT_NUMSEG" @@ -1587,28 +1895,26 @@ "DISJOINT_EMPTY_REFL" > "HOLLight.hollight.DISJOINT_EMPTY_REFL" "DISJOINT_EMPTY" > "HOLLight.hollight.DISJOINT_EMPTY" "DISJOINT_DELETE_SYM" > "HOLLight.hollight.DISJOINT_DELETE_SYM" + "DIMINDEX_UNIV" > "HOLLight.hollight.DIMINDEX_UNIV" + "DIMINDEX_UNIQUE" > "HOLLight.hollight.DIMINDEX_UNIQUE" "DIMINDEX_NONZERO" > "HOLLight.hollight.DIMINDEX_NONZERO" - "DIMINDEX_HAS_SIZE_FINITE_SUM" > "HOLLight.hollight.DIMINDEX_HAS_SIZE_FINITE_SUM" "DIMINDEX_GE_1" > "HOLLight.hollight.DIMINDEX_GE_1" - "DIMINDEX_FINITE_SUM" > "HOLLight.hollight.DIMINDEX_FINITE_SUM" "DIMINDEX_FINITE_IMAGE" > "HOLLight.hollight.DIMINDEX_FINITE_IMAGE" - "DIFF_def" > "HOLLight.hollight.DIFF_def" - "DIFF_UNIV" > "HOLLight.hollight.DIFF_UNIV" - "DIFF_INSERT" > "HOLLight.hollight.DIFF_INSERT" - "DIFF_EQ_EMPTY" > "HOLLight.hollight.DIFF_EQ_EMPTY" - "DIFF_EMPTY" > "HOLLight.hollight.DIFF_EMPTY" - "DIFF_DIFF" > "HOLLight.hollight.DIFF_DIFF" + "DIFF_UNIV" > "Set.Diff_UNIV" + "DIFF_INTERS" > "HOLLight.hollight.DIFF_INTERS" + "DIFF_INSERT" > "Set.Diff_insert2" + "DIFF_EQ_EMPTY" > "Set.Diff_cancel" + "DIFF_EMPTY" > "Set.Diff_empty" + "DIFF_DIFF" > "Set.Diff_idemp" "DEST_REC_INJ" > "HOLLight.hollight.recspace._dest_rec_inject" - "DELETE_def" > "HOLLight.hollight.DELETE_def" "DELETE_SUBSET" > "HOLLight.hollight.DELETE_SUBSET" "DELETE_NON_ELEMENT" > "HOLLight.hollight.DELETE_NON_ELEMENT" "DELETE_INTER" > "HOLLight.hollight.DELETE_INTER" "DELETE_INSERT" > "HOLLight.hollight.DELETE_INSERT" "DELETE_DELETE" > "HOLLight.hollight.DELETE_DELETE" "DELETE_COMM" > "HOLLight.hollight.DELETE_COMM" - "DEF_~" > "HOL.simp_thms_19" - "DEF_two_2" > "HOLLight.hollight.DEF_two_2" - "DEF_two_1" > "HOLLight.hollight.DEF_two_1" + "DEF_~" > "Groebner_Basis.bool_simps_19" + "DEF_vector" > "HOLLight.hollight.DEF_vector" "DEF_treal_of_num" > "HOLLight.hollight.DEF_treal_of_num" "DEF_treal_neg" > "HOLLight.hollight.DEF_treal_neg" "DEF_treal_mul" > "HOLLight.hollight.DEF_treal_mul" @@ -1616,20 +1922,20 @@ "DEF_treal_inv" > "HOLLight.hollight.DEF_treal_inv" "DEF_treal_eq" > "HOLLight.hollight.DEF_treal_eq" "DEF_treal_add" > "HOLLight.hollight.DEF_treal_add" - "DEF_three_3" > "HOLLight.hollight.DEF_three_3" - "DEF_three_2" > "HOLLight.hollight.DEF_three_2" - "DEF_three_1" > "HOLLight.hollight.DEF_three_1" "DEF_tailadmissible" > "HOLLight.hollight.DEF_tailadmissible" "DEF_support" > "HOLLight.hollight.DEF_support" "DEF_superadmissible" > "HOLLight.hollight.DEF_superadmissible" "DEF_sum" > "HOLLight.hollight.DEF_sum" "DEF_sndcart" > "HOLLight.hollight.DEF_sndcart" - "DEF_set_of_list" > "HOLLight.hollight.DEF_set_of_list" + "DEF_set_of_list" > "HOLLightList.DEF_set_of_list" + "DEF_rem" > "HOLLight.hollight.DEF_rem" "DEF_real_sub" > "HOLLight.hollight.DEF_real_sub" + "DEF_real_sgn" > "HOLLight.hollight.DEF_real_sgn" "DEF_real_pow" > "HOLLight.hollight.DEF_real_pow" "DEF_real_of_num" > "HOLLight.hollight.DEF_real_of_num" "DEF_real_neg" > "HOLLight.hollight.DEF_real_neg" "DEF_real_mul" > "HOLLight.hollight.DEF_real_mul" + "DEF_real_mod" > "HOLLight.hollight.DEF_real_mod" "DEF_real_min" > "HOLLight.hollight.DEF_real_min" "DEF_real_max" > "HOLLight.hollight.DEF_real_max" "DEF_real_lt" > "HOLLight.hollight.DEF_real_lt" @@ -1642,7 +1948,12 @@ "DEF_real_abs" > "HOLLight.hollight.DEF_real_abs" "DEF_pastecart" > "HOLLight.hollight.DEF_pastecart" "DEF_pairwise" > "HOLLight.hollight.DEF_pairwise" - "DEF_o" > "Fun.o_apply" + "DEF_o" > "Fun.comp_def" + "DEF_num_of_int" > "HOLLight.hollight.DEF_num_of_int" + "DEF_num_mod" > "HOLLight.hollight.DEF_num_mod" + "DEF_num_gcd" > "HOLLight.hollight.DEF_num_gcd" + "DEF_num_divides" > "HOLLight.hollight.DEF_num_divides" + "DEF_num_coprime" > "HOLLight.hollight.DEF_num_coprime" "DEF_nsum" > "HOLLight.hollight.DEF_nsum" "DEF_neutral" > "HOLLight.hollight.DEF_neutral" "DEF_nadd_rinv" > "HOLLight.hollight.DEF_nadd_rinv" @@ -1653,28 +1964,27 @@ "DEF_nadd_eq" > "HOLLight.hollight.DEF_nadd_eq" "DEF_nadd_add" > "HOLLight.hollight.DEF_nadd_add" "DEF_monoidal" > "HOLLight.hollight.DEF_monoidal" - "DEF_mod_real" > "HOLLight.hollight.DEF_mod_real" - "DEF_mod_nat" > "HOLLight.hollight.DEF_mod_nat" - "DEF_mod_int" > "HOLLight.hollight.DEF_mod_int" - "DEF_mk_pair" > "Product_Type.Pair_Rep_def" "DEF_minimal" > "HOLLight.hollight.DEF_minimal" - "DEF_measure" > "HOLLight.hollight.DEF_measure" - "DEF_list_of_set" > "HOLLight.hollight.DEF_list_of_set" "DEF_lambda" > "HOLLight.hollight.DEF_lambda" "DEF_iterate" > "HOLLight.hollight.DEF_iterate" "DEF_is_nadd" > "HOLLight.hollight.DEF_is_nadd" - "DEF_is_int" > "HOLLight.hollight.DEF_is_int" + "DEF_integer" > "HOLLight.hollight.DEF_integer" "DEF_int_sub" > "HOLLight.hollight.DEF_int_sub" + "DEF_int_sgn" > "HOLLight.hollight.DEF_int_sgn" "DEF_int_pow" > "HOLLight.hollight.DEF_int_pow" "DEF_int_of_num" > "HOLLight.hollight.DEF_int_of_num" "DEF_int_neg" > "HOLLight.hollight.DEF_int_neg" "DEF_int_mul" > "HOLLight.hollight.DEF_int_mul" + "DEF_int_mod" > "HOLLight.hollight.DEF_int_mod" "DEF_int_min" > "HOLLight.hollight.DEF_int_min" "DEF_int_max" > "HOLLight.hollight.DEF_int_max" "DEF_int_lt" > "HOLLight.hollight.DEF_int_lt" "DEF_int_le" > "HOLLight.hollight.DEF_int_le" "DEF_int_gt" > "HOLLight.hollight.DEF_int_gt" "DEF_int_ge" > "HOLLight.hollight.DEF_int_ge" + "DEF_int_gcd" > "HOLLight.hollight.DEF_int_gcd" + "DEF_int_divides" > "HOLLight.hollight.DEF_int_divides" + "DEF_int_coprime" > "HOLLight.hollight.DEF_int_coprime" "DEF_int_add" > "HOLLight.hollight.DEF_int_add" "DEF_int_abs" > "HOLLight.hollight.DEF_int_abs" "DEF_hreal_of_num" > "HOLLight.hollight.DEF_hreal_of_num" @@ -1683,178 +1993,178 @@ "DEF_hreal_inv" > "HOLLight.hollight.DEF_hreal_inv" "DEF_hreal_add" > "HOLLight.hollight.DEF_hreal_add" "DEF_fstcart" > "HOLLight.hollight.DEF_fstcart" - "DEF_finite_index" > "HOLLight.hollight.DEF_finite_index" + "DEF_div" > "HOLLight.hollight.DEF_div" "DEF_dist" > "HOLLight.hollight.DEF_dist" "DEF_dimindex" > "HOLLight.hollight.DEF_dimindex" "DEF_admissible" > "HOLLight.hollight.DEF_admissible" - "DEF__star_" > "HOLLightCompat.mult_altdef" - "DEF__slash__backslash_" > "HOLLightCompat.light_and_def" - "DEF__questionmark__exclamationmark_" > "HOL4Compat.EXISTS_UNIQUE_DEF" + "DEF__star_" > "HOLLightCompat.DEF__star_" + "DEF__slash__backslash_" > "HOLLightCompat.DEF__slash__backslash_" + "DEF__questionmark__exclamationmark_" > "HOLLightCompat.EXISTS_UNIQUE_THM" "DEF__questionmark_" > "HOL.Ex_def" - "DEF__lessthan__equal_" > "HOLLight.hollight.DEF__lessthan__equal_" - "DEF__lessthan_" > "HOLLight.hollight.DEF__lessthan_" - "DEF__greaterthan__equal_" > "HOLLight.hollight.DEF__greaterthan__equal_" - "DEF__greaterthan_" > "HOLLight.hollight.DEF__greaterthan_" + "DEF__lessthan__equal__c" > "HOLLight.hollight.DEF__lessthan__equal__c" + "DEF__lessthan__equal_" > "HOLLightCompat.DEF__lessthan__equal_" + "DEF__lessthan__c" > "HOLLight.hollight.DEF__lessthan__c" + "DEF__lessthan_" > "HOLLightCompat.DEF__lessthan_" + "DEF__greaterthan__equal__c" > "HOLLight.hollight.DEF__greaterthan__equal__c" + "DEF__greaterthan__equal_" > "HOLLightCompat.DEF__greaterthan__equal_" + "DEF__greaterthan__c" > "HOLLight.hollight.DEF__greaterthan__c" + "DEF__greaterthan_" > "HOLLightCompat.DEF__greaterthan_" "DEF__exclamationmark_" > "HOL.All_def" - "DEF__equal__equal__greaterthan_" > "HOLLightCompat.light_imp_def" + "DEF__equal__equal__greaterthan_" > "HOLLightCompat.DEF__equal__equal__greaterthan_" "DEF__equal__equal_" > "HOLLight.hollight.DEF__equal__equal_" - "DEF__dot__dot_" > "HOLLight.hollight.DEF__dot__dot_" + "DEF__equal__c" > "HOLLight.hollight.DEF__equal__c" + "DEF__dot__dot_" > "HOLLightCompat.DEF__dot__dot_" "DEF__backslash__slash_" > "HOL.or_def" + "DEF__UNGUARDED_PATTERN" > "HOLLight.hollight.DEF__UNGUARDED_PATTERN" + "DEF__SEQPATTERN" > "HOLLight.hollight.DEF__SEQPATTERN" + "DEF__MATCH" > "HOLLight.hollight.DEF__MATCH" + "DEF__GUARDED_PATTERN" > "HOLLight.hollight.DEF__GUARDED_PATTERN" + "DEF__FUNCTION" > "HOLLight.hollight.DEF__FUNCTION" "DEF__FALSITY_" > "HOLLight.hollight.DEF__FALSITY_" - "DEF__10328" > "HOLLight.hollight.DEF__10328" - "DEF__10327" > "HOLLight.hollight.DEF__10327" - "DEF__10326" > "HOLLight.hollight.DEF__10326" - "DEF__10303" > "HOLLight.hollight.DEF__10303" - "DEF__10302" > "HOLLight.hollight.DEF__10302" + "DEF__11937" > "HOLLight.hollight.DEF__11937" "DEF_ZRECSPACE" > "HOLLight.hollight.DEF_ZRECSPACE" - "DEF_ZIP" > "HOLLight.hollight.DEF_ZIP" "DEF_ZCONSTR" > "HOLLight.hollight.DEF_ZCONSTR" "DEF_ZBOT" > "HOLLight.hollight.DEF_ZBOT" - "DEF_WF" > "HOLLight.hollight.DEF_WF" - "DEF_UNIV" > "HOLLight.hollight.DEF_UNIV" - "DEF_UNIONS" > "HOLLight.hollight.DEF_UNIONS" - "DEF_UNION" > "HOLLight.hollight.DEF_UNION" + "DEF_WF" > "HOLLightCompat.DEF_WF" + "DEF_UNIV" > "HOLLightCompat.DEF_UNIV" + "DEF_UNIONS" > "HOLLightCompat.DEF_UNIONS" + "DEF_UNION" > "HOLLightCompat.DEF_UNION" "DEF_UNCURRY" > "HOLLight.hollight.DEF_UNCURRY" - "DEF_TL" > "HOLLight.hollight.DEF_TL" "DEF_T" > "HOL.True_def" "DEF_SURJ" > "HOLLight.hollight.DEF_SURJ" - "DEF_SUBSET" > "HOLLight.hollight.DEF_SUBSET" - "DEF_SOME" > "HOLLight.hollight.DEF_SOME" - "DEF_SND" > "HOLLightCompat.snd_altdef" + "DEF_SUBSET" > "HOLLightCompat.DEF_SUBSET" + "DEF_SND" > "HOLLightCompat.DEF_SND" "DEF_SING" > "HOLLight.hollight.DEF_SING" - "DEF_SETSPEC" > "HOLLight.hollight.DEF_SETSPEC" - "DEF_REVERSE" > "HOLLight.hollight.DEF_REVERSE" + "DEF_SETSPEC" > "HOLLightCompat.SETSPEC_def" + "DEF_REVERSE" > "HOLLightList.DEF_REVERSE" "DEF_REST" > "HOLLight.hollight.DEF_REST" - "DEF_REPLICATE" > "HOLLight.hollight.DEF_REPLICATE" - "DEF_PSUBSET" > "HOLLight.hollight.DEF_PSUBSET" - "DEF_PRE" > "HOLLightCompat.Pred_altdef" + "DEF_REPLICATE" > "HOLLightList.DEF_REPLICATE" + "DEF_PSUBSET" > "HOLLightCompat.DEF_PSUBSET" + "DEF_PRE" > "HOLLightCompat.DEF_PRE" "DEF_PASSOC" > "HOLLight.hollight.DEF_PASSOC" "DEF_PAIRWISE" > "HOLLight.hollight.DEF_PAIRWISE" - "DEF_OUTR" > "HOLLight.hollight.DEF_OUTR" - "DEF_OUTL" > "HOLLight.hollight.DEF_OUTL" "DEF_ONTO" > "Fun.surj_def" - "DEF_ONE_ONE" > "HOL4Setup.ONE_ONE_DEF" - "DEF_ODD" > "HOLLight.hollight.DEF_ODD" + "DEF_ONE_ONE" > "HOLLightCompat.DEF_ONE_ONE" + "DEF_ODD" > "HOLLightCompat.DEF_ODD" + "DEF_NUM_REP" > "HOLLight.hollight.DEF_NUM_REP" "DEF_NUMSUM" > "HOLLight.hollight.DEF_NUMSUM" "DEF_NUMSND" > "HOLLight.hollight.DEF_NUMSND" "DEF_NUMRIGHT" > "HOLLight.hollight.DEF_NUMRIGHT" "DEF_NUMPAIR" > "HOLLight.hollight.DEF_NUMPAIR" "DEF_NUMLEFT" > "HOLLight.hollight.DEF_NUMLEFT" "DEF_NUMFST" > "HOLLight.hollight.DEF_NUMFST" - "DEF_NUMERAL" > "HOLLightCompat.NUMERAL_rew" - "DEF_NULL" > "HOLLight.hollight.DEF_NULL" - "DEF_NONE" > "HOLLight.hollight.DEF_NONE" - "DEF_NIL" > "HOLLight.hollight.DEF_NIL" - "DEF_MOD" > "HOLLight.hollight.DEF_MOD" - "DEF_MEM" > "HOLLight.hollight.DEF_MEM" - "DEF_MAP2" > "HOLLight.hollight.DEF_MAP2" - "DEF_MAP" > "HOLLight.hollight.DEF_MAP" + "DEF_NUMERAL" > "HOLLightCompat.NUMERAL_def" + "DEF_NULL" > "HOLLightList.DEF_NULL" + "DEF_MOD" > "HOLLightCompat.DEF_MOD" + "DEF_MIN" > "Orderings.ord_class.min_def" + "DEF_MEM" > "HOLLightList.DEF_MEM" + "DEF_MEASURE" > "HOLLightCompat.MEASURE_def" + "DEF_MAX" > "Orderings.ord_class.max_def" + "DEF_MAP" > "HOLLightList.DEF_MAP" "DEF_LET_END" > "HOLLight.hollight.DEF_LET_END" - "DEF_LET" > "HOL4Compat.LET_def" - "DEF_LENGTH" > "HOLLight.hollight.DEF_LENGTH" - "DEF_LAST" > "HOLLight.hollight.DEF_LAST" + "DEF_LET" > "HOLLightCompat.LET_def" + "DEF_LENGTH" > "HOLLightList.DEF_LENGTH" + "DEF_LAST" > "HOLLightList.DEF_LAST" "DEF_ITSET" > "HOLLight.hollight.DEF_ITSET" - "DEF_ITLIST2" > "HOLLight.hollight.DEF_ITLIST2" - "DEF_ITLIST" > "HOLLight.hollight.DEF_ITLIST" + "DEF_ITLIST" > "HOLLightList.DEF_ITLIST" "DEF_ISO" > "HOLLight.hollight.DEF_ISO" - "DEF_INTERS" > "HOLLight.hollight.DEF_INTERS" - "DEF_INTER" > "HOLLight.hollight.DEF_INTER" - "DEF_INSERT" > "HOLLight.hollight.DEF_INSERT" + "DEF_INTERS" > "HOLLightCompat.DEF_INTERS" + "DEF_INTER" > "HOLLightCompat.DEF_INTER" + "DEF_INSERT" > "HOLLightCompat.DEF_INSERT" "DEF_INJP" > "HOLLight.hollight.DEF_INJP" "DEF_INJN" > "HOLLight.hollight.DEF_INJN" "DEF_INJF" > "HOLLight.hollight.DEF_INJF" "DEF_INJA" > "HOLLight.hollight.DEF_INJA" "DEF_INJ" > "HOLLight.hollight.DEF_INJ" - "DEF_INFINITE" > "HOLLight.hollight.DEF_INFINITE" - "DEF_IN" > "HOLLight.hollight.DEF_IN" - "DEF_IMAGE" > "HOLLight.hollight.DEF_IMAGE" + "DEF_INFINITE" > "HOLLightCompat.DEF_INFINITE" + "DEF_IND_SUC" > "HOLLight.hollight.DEF_IND_SUC" + "DEF_IND_0" > "HOLLight.hollight.DEF_IND_0" + "DEF_IN" > "Set.mem_def" + "DEF_IMAGE" > "HOLLightCompat.DEF_IMAGE" "DEF_I" > "Fun.id_apply" - "DEF_HD" > "HOLLight.hollight.DEF_HD" "DEF_HAS_SIZE" > "HOLLight.hollight.DEF_HAS_SIZE" - "DEF_GSPEC" > "HOLLight.hollight.DEF_GSPEC" - "DEF_GEQ" > "HOLLight.hollight.DEF_GEQ" - "DEF_GABS" > "HOLLight.hollight.DEF_GABS" - "DEF_FST" > "HOLLightCompat.fst_altdef" + "DEF_GSPEC" > "HOLLightCompat.DEF_GSPEC" + "DEF_GEQ" > "HOLLightCompat.DEF_GEQ" + "DEF_GABS" > "HOLLightCompat.DEF_GABS" + "DEF_FST" > "HOLLightCompat.DEF_FST" "DEF_FNIL" > "HOLLight.hollight.DEF_FNIL" "DEF_FINREC" > "HOLLight.hollight.DEF_FINREC" - "DEF_FINITE" > "HOLLight.hollight.DEF_FINITE" - "DEF_FILTER" > "HOLLight.hollight.DEF_FILTER" + "DEF_FINITE" > "HOLLightCompat.DEF_FINITE" + "DEF_FILTER" > "HOLLightList.DEF_FILTER" "DEF_FCONS" > "HOLLight.hollight.DEF_FCONS" - "DEF_FACT" > "HOLLight.hollight.DEF_FACT" + "DEF_FACT" > "HOLLightCompat.DEF_FACT" "DEF_F" > "HOL.False_def" - "DEF_EXP" > "HOLLight.hollight.DEF_EXP" - "DEF_EX" > "HOLLight.hollight.DEF_EX" - "DEF_EVEN" > "HOLLight.hollight.DEF_EVEN" - "DEF_EMPTY" > "HOLLight.hollight.DEF_EMPTY" - "DEF_EL" > "HOLLight.hollight.DEF_EL" - "DEF_DIV" > "HOLLight.hollight.DEF_DIV" - "DEF_DISJOINT" > "HOLLight.hollight.DEF_DISJOINT" - "DEF_DIFF" > "HOLLight.hollight.DEF_DIFF" - "DEF_DELETE" > "HOLLight.hollight.DEF_DELETE" + "DEF_EXP" > "HOLLightCompat.DEF_EXP" + "DEF_EX" > "HOLLightList.DEF_EX" + "DEF_EVEN" > "HOLLightCompat.DEF_EVEN" + "DEF_EMPTY" > "HOLLightCompat.DEF_EMPTY" + "DEF_EL" > "HOLLightList.DEF_EL" + "DEF_DIV" > "HOLLightCompat.DEF_DIV" + "DEF_DISJOINT" > "HOLLightCompat.DEF_DISJOINT" + "DEF_DIFF" > "HOLLightCompat.DEF_DIFF" + "DEF_DELETE" > "HOLLightCompat.DEF_DELETE" "DEF_DECIMAL" > "HOLLight.hollight.DEF_DECIMAL" - "DEF_CURRY" > "HOLLight.hollight.DEF_CURRY" + "DEF_CURRY" > "Product_Type.curry_conv" + "DEF_CROSS" > "HOLLight.hollight.DEF_CROSS" "DEF_COUNTABLE" > "HOLLight.hollight.DEF_COUNTABLE" "DEF_CONSTR" > "HOLLight.hollight.DEF_CONSTR" - "DEF_CONS" > "HOLLight.hollight.DEF_CONS" - "DEF_COND" > "HOLLight.hollight.DEF_COND" - "DEF_CHOICE" > "HOLLight.hollight.DEF_CHOICE" + "DEF_COND" > "HOLLightCompat.COND_DEF" + "DEF_CHOICE" > "HOLLightCompat.DEF_CHOICE" "DEF_CASEWISE" > "HOLLight.hollight.DEF_CASEWISE" - "DEF_CARD_LT" > "HOLLight.hollight.DEF_CARD_LT" - "DEF_CARD_LE" > "HOLLight.hollight.DEF_CARD_LE" - "DEF_CARD_GT" > "HOLLight.hollight.DEF_CARD_GT" - "DEF_CARD_GE" > "HOLLight.hollight.DEF_CARD_GE" - "DEF_CARD_EQ" > "HOLLight.hollight.DEF_CARD_EQ" "DEF_CARD" > "HOLLight.hollight.DEF_CARD" + "DEF_BUTLAST" > "HOLLightList.DEF_BUTLAST" "DEF_BOTTOM" > "HOLLight.hollight.DEF_BOTTOM" - "DEF_BIT1" > "HOLLightCompat.NUMERAL_BIT1_altdef" - "DEF_BIT0" > "HOLLightCompat.NUMERAL_BIT0_def" + "DEF_BIT1" > "HOLLightCompat.BIT1_DEF" + "DEF_BIT0" > "HOLLightCompat.BIT0_DEF" "DEF_BIJ" > "HOLLight.hollight.DEF_BIJ" - "DEF_ASSOC" > "HOLLight.hollight.DEF_ASSOC" - "DEF_APPEND" > "HOLLight.hollight.DEF_APPEND" - "DEF_ALL2" > "HOLLight.hollight.DEF_ALL2" - "DEF_ALL" > "HOLLight.hollight.DEF_ALL" - "DEF_-" > "HOLLightCompat.sub_altdef" - "DEF_," > "Product_Type.Pair_def" - "DEF_+" > "HOLLightCompat.add_altdef" + "DEF_ASCII" > "HOLLight.hollight.DEF_ASCII" + "DEF_APPEND" > "HOLLightList.DEF_APPEND" + "DEF_ALL2" > "HOLLightList.DEF_ALL2" + "DEF_ALL" > "HOLLightList.DEF_ALL" + "DEF_-" > "HOLLightCompat.DEF_MINUS" + "DEF_+" > "HOLLightCompat.DEF_PLUS" "DEF_$" > "HOLLight.hollight.DEF_$" "DECOMPOSITION" > "HOLLight.hollight.DECOMPOSITION" "DECIMAL_def" > "HOLLight.hollight.DECIMAL_def" - "CURRY_def" > "HOLLight.hollight.CURRY_def" + "CROSS_def" > "HOLLight.hollight.CROSS_def" + "CROSS_EQ_EMPTY" > "HOLLight.hollight.CROSS_EQ_EMPTY" "COUNTABLE_def" > "HOLLight.hollight.COUNTABLE_def" - "CONS_def" > "HOLLight.hollight.CONS_def" - "CONS_11" > "HOLLight.hollight.CONS_11" + "CONS_HD_TL" > "List.hd_Cons_tl" + "CONS_11" > "List.list.inject" "CONSTR_def" > "HOLLight.hollight.CONSTR_def" "CONSTR_REC" > "HOLLight.hollight.CONSTR_REC" "CONSTR_INJ" > "HOLLight.hollight.CONSTR_INJ" "CONSTR_IND" > "HOLLight.hollight.CONSTR_IND" "CONSTR_BOT" > "HOLLight.hollight.CONSTR_BOT" - "CONJ_SYM" > "HOL.conj_comms_1" - "CONJ_ASSOC" > "HOL.conj_assoc" + "CONJ_SYM" > "Groebner_Basis.dnf_3" + "CONJ_ASSOC" > "HOL.conj_ac_3" "CONJ_ACI" > "HOLLight.hollight.CONJ_ACI" - "COND_def" > "HOLLight.hollight.COND_def" "COND_RATOR" > "HOLLight.hollight.COND_RATOR" - "COND_RAND" > "HOLLight.hollight.COND_RAND" - "COND_ID" > "HOLLight.hollight.COND_ID" + "COND_RAND" > "HOL.if_distrib" + "COND_ID" > "HOL.if_cancel" "COND_EXPAND" > "HOLLight.hollight.COND_EXPAND" "COND_EQ_CLAUSE" > "HOLLight.hollight.COND_EQ_CLAUSE" - "COND_ELIM_THM" > "HOLLight.hollight.COND_ELIM_THM" + "COND_ELIM_THM" > "HOL.if_splits_1" "COND_CLAUSES" > "HOLLight.hollight.COND_CLAUSES" "COND_ABS" > "HOLLight.hollight.COND_ABS" - "COMPONENT" > "HOLLight.hollight.COMPONENT" + "COMPONENT" > "Set.insertI1" + "CHOOSE_SUBSET_STRONG" > "HOLLight.hollight.CHOOSE_SUBSET_STRONG" "CHOOSE_SUBSET" > "HOLLight.hollight.CHOOSE_SUBSET" - "CHOICE_def" > "HOLLight.hollight.CHOICE_def" "CHOICE_DEF" > "HOLLight.hollight.CHOICE_DEF" "CASEWISE_def" > "HOLLight.hollight.CASEWISE_def" - "CASEWISE_WORKS" > "HOLLight.hollight.CASEWISE_WORKS" - "CASEWISE_CASES" > "HOLLight.hollight.CASEWISE_CASES" - "CASEWISE" > "HOLLight.hollight.CASEWISE" + "CART_EQ_FULL" > "HOLLight.hollight.CART_EQ_FULL" "CART_EQ" > "HOLLight.hollight.CART_EQ" "CARD_def" > "HOLLight.hollight.CARD_def" + "CARD_UNION_OVERLAP_EQ" > "HOLLight.hollight.CARD_UNION_OVERLAP_EQ" + "CARD_UNION_OVERLAP" > "HOLLight.hollight.CARD_UNION_OVERLAP" "CARD_UNION_LE" > "HOLLight.hollight.CARD_UNION_LE" + "CARD_UNION_GEN" > "HOLLight.hollight.CARD_UNION_GEN" "CARD_UNION_EQ" > "HOLLight.hollight.CARD_UNION_EQ" "CARD_UNIONS_LE" > "HOLLight.hollight.CARD_UNIONS_LE" + "CARD_UNIONS" > "HOLLight.hollight.CARD_UNIONS" "CARD_UNION" > "HOLLight.hollight.CARD_UNION" "CARD_SUBSET_LE" > "HOLLight.hollight.CARD_SUBSET_LE" + "CARD_SUBSET_IMAGE" > "HOLLight.hollight.CARD_SUBSET_IMAGE" "CARD_SUBSET_EQ" > "HOLLight.hollight.CARD_SUBSET_EQ" "CARD_SUBSET" > "HOLLight.hollight.CARD_SUBSET" "CARD_PSUBSET" > "HOLLight.hollight.CARD_PSUBSET" @@ -1865,22 +2175,20 @@ "CARD_NUMSEG_LE" > "HOLLight.hollight.CARD_NUMSEG_LE" "CARD_NUMSEG_1" > "HOLLight.hollight.CARD_NUMSEG_1" "CARD_NUMSEG" > "HOLLight.hollight.CARD_NUMSEG" - "CARD_LT_def" > "HOLLight.hollight.CARD_LT_def" - "CARD_LE_def" > "HOLLight.hollight.CARD_LE_def" "CARD_LE_INJ" > "HOLLight.hollight.CARD_LE_INJ" "CARD_IMAGE_LE" > "HOLLight.hollight.CARD_IMAGE_LE" + "CARD_IMAGE_INJ_EQ" > "HOLLight.hollight.CARD_IMAGE_INJ_EQ" "CARD_IMAGE_INJ" > "HOLLight.hollight.CARD_IMAGE_INJ" - "CARD_GT_def" > "HOLLight.hollight.CARD_GT_def" - "CARD_GE_def" > "HOLLight.hollight.CARD_GE_def" - "CARD_GE_TRANS" > "HOLLight.hollight.CARD_GE_TRANS" - "CARD_GE_REFL" > "HOLLight.hollight.CARD_GE_REFL" "CARD_FUNSPACE" > "HOLLight.hollight.CARD_FUNSPACE" "CARD_FINITE_IMAGE" > "HOLLight.hollight.CARD_FINITE_IMAGE" - "CARD_EQ_def" > "HOLLight.hollight.CARD_EQ_def" "CARD_EQ_SUM" > "HOLLight.hollight.CARD_EQ_SUM" "CARD_EQ_NSUM" > "HOLLight.hollight.CARD_EQ_NSUM" + "CARD_EQ_BIJECTIONS" > "HOLLight.hollight.CARD_EQ_BIJECTIONS" + "CARD_EQ_BIJECTION" > "HOLLight.hollight.CARD_EQ_BIJECTION" "CARD_EQ_0" > "HOLLight.hollight.CARD_EQ_0" + "CARD_DIFF" > "HOLLight.hollight.CARD_DIFF" "CARD_DELETE" > "HOLLight.hollight.CARD_DELETE" + "CARD_CROSS" > "HOLLight.hollight.CARD_CROSS" "CARD_CLAUSES" > "HOLLight.hollight.CARD_CLAUSES" "BOUNDS_NOTZERO" > "HOLLight.hollight.BOUNDS_NOTZERO" "BOUNDS_LINEAR_0" > "HOLLight.hollight.BOUNDS_LINEAR_0" @@ -1888,12 +2196,19 @@ "BOUNDS_IGNORE" > "HOLLight.hollight.BOUNDS_IGNORE" "BOUNDS_DIVIDED" > "HOLLight.hollight.BOUNDS_DIVIDED" "BOTTOM_def" > "HOLLight.hollight.BOTTOM_def" - "BOOL_CASES_AX" > "Datatype.bool.nchotomy" - "BIT1_THM" > "HOLLightCompat.NUMERAL_BIT1_altdef" - "BIT0_THM" > "HOLLightCompat.NUMERAL_BIT0_def" + "BOOL_CASES_AX" > "HOL.True_or_False" + "BIT1_THM" > "HOLLight.hollight.BIT1_THM" + "BIT1" > "HOLLight.hollight.BIT1" + "BIT0_THM" > "Int.semiring_mult_2" + "BIT0" > "Int.semiring_mult_2" "BIJ_def" > "HOLLight.hollight.BIJ_def" - "BETA_THM" > "Presburger.fm_modd_pinf" - "ASSOC_def" > "HOLLight.hollight.ASSOC_def" + "BIJECTIVE_ON_LEFT_RIGHT_INVERSE" > "HOLLight.hollight.BIJECTIVE_ON_LEFT_RIGHT_INVERSE" + "BIJECTIVE_LEFT_RIGHT_INVERSE" > "HOLLight.hollight.BIJECTIVE_LEFT_RIGHT_INVERSE" + "BIJECTIONS_HAS_SIZE_EQ" > "HOLLight.hollight.BIJECTIONS_HAS_SIZE_EQ" + "BIJECTIONS_HAS_SIZE" > "HOLLight.hollight.BIJECTIONS_HAS_SIZE" + "BIJECTIONS_CARD_EQ" > "HOLLight.hollight.BIJECTIONS_CARD_EQ" + "BETA_THM" > "HOL.eta_contract_eq" + "ASCII_def" > "HOLLight.hollight.ASCII_def" "ARITH_ZERO" > "HOLLight.hollight.ARITH_ZERO" "ARITH_SUC" > "HOLLight.hollight.ARITH_SUC" "ARITH_SUB" > "HOLLight.hollight.ARITH_SUB" @@ -1906,37 +2221,42 @@ "ARITH_EVEN" > "HOLLight.hollight.ARITH_EVEN" "ARITH_EQ" > "HOLLight.hollight.ARITH_EQ" "ARITH_ADD" > "HOLLight.hollight.ARITH_ADD" - "APPEND_def" > "HOLLight.hollight.APPEND_def" - "APPEND_NIL" > "HOLLight.hollight.APPEND_NIL" - "APPEND_EQ_NIL" > "HOLLight.hollight.APPEND_EQ_NIL" - "APPEND_ASSOC" > "HOLLight.hollight.APPEND_ASSOC" + "APPEND_NIL" > "List.append_Nil2" + "APPEND_EQ_NIL" > "List.append_is_Nil_conv" + "APPEND_BUTLAST_LAST" > "List.append_butlast_last_id" + "APPEND_ASSOC" > "List.append_assoc" "AND_FORALL_THM" > "HOL.all_conj_distrib" "AND_CLAUSES" > "HOLLight.hollight.AND_CLAUSES" - "AND_ALL2" > "HOLLight.hollight.AND_ALL2" - "AND_ALL" > "HOLLight.hollight.AND_ALL" - "ALL_list_def" > "HOLLight.hollight.ALL_list_def" - "ALL_T" > "HOLLight.hollight.ALL_T" - "ALL_MP" > "HOLLight.hollight.ALL_MP" - "ALL_MEM" > "HOLLight.hollight.ALL_MEM" - "ALL_MAP" > "HOLLight.hollight.ALL_MAP" - "ALL_IMP" > "HOLLight.hollight.ALL_IMP" - "ALL_APPEND" > "HOLLight.hollight.ALL_APPEND" - "ALL2_def" > "HOLLight.hollight.ALL2_def" - "ALL2_MAP2" > "HOLLight.hollight.ALL2_MAP2" - "ALL2_MAP" > "HOLLight.hollight.ALL2_MAP" - "ALL2_AND_RIGHT" > "HOLLight.hollight.ALL2_AND_RIGHT" - "ALL2_ALL" > "HOLLight.hollight.ALL2_ALL" - "ALL2" > "HOLLight.hollight.ALL2" + "AND_ALL2" > "HOLLightList.AND_ALL2" + "AND_ALL" > "HOLLightList.AND_ALL" + "ALL_T" > "HOLLightList.ALL_T" + "ALL_MP" > "HOLLightList.ALL_MP" + "ALL_MEM" > "HOLLightList.ALL_MEM" + "ALL_IMP" > "HOLLightList.ALL_IMP" + "ALL_EL" > "List.list_all_length" + "ALL_APPEND" > "List.list_all_append" + "ALL2_MAP2" > "HOLLightList.ALL2_MAP2" + "ALL2_MAP" > "HOLLightList.ALL2_MAP" + "ALL2_AND_RIGHT" > "HOLLightList.ALL2_AND_RIGHT" + "ALL2_ALL" > "HOLLightList.ALL2_ALL" + "ALL2" > "HOLLightList.ALL2" + "ADMISSIBLE_UNGUARDED_PATTERN" > "HOLLight.hollight.ADMISSIBLE_UNGUARDED_PATTERN" "ADMISSIBLE_SUM" > "HOLLight.hollight.ADMISSIBLE_SUM" + "ADMISSIBLE_SEQPATTERN" > "HOLLight.hollight.ADMISSIBLE_SEQPATTERN" + "ADMISSIBLE_RAND" > "HOLLight.hollight.ADMISSIBLE_RAND" "ADMISSIBLE_NSUM" > "HOLLight.hollight.ADMISSIBLE_NSUM" "ADMISSIBLE_NEST" > "HOLLight.hollight.ADMISSIBLE_NEST" + "ADMISSIBLE_MATCH_SEQPATTERN" > "HOLLight.hollight.ADMISSIBLE_MATCH_SEQPATTERN" + "ADMISSIBLE_MATCH" > "HOLLight.hollight.ADMISSIBLE_MATCH" + "ADMISSIBLE_MAP" > "HOLLight.hollight.ADMISSIBLE_MAP" "ADMISSIBLE_LAMBDA" > "HOLLight.hollight.ADMISSIBLE_LAMBDA" "ADMISSIBLE_IMP_SUPERADMISSIBLE" > "HOLLight.hollight.ADMISSIBLE_IMP_SUPERADMISSIBLE" + "ADMISSIBLE_GUARDED_PATTERN" > "HOLLight.hollight.ADMISSIBLE_GUARDED_PATTERN" "ADMISSIBLE_CONST" > "HOLLight.hollight.ADMISSIBLE_CONST" "ADMISSIBLE_COND" > "HOLLight.hollight.ADMISSIBLE_COND" "ADMISSIBLE_COMB" > "HOLLight.hollight.ADMISSIBLE_COMB" "ADMISSIBLE_BASE" > "HOLLight.hollight.ADMISSIBLE_BASE" - "ADD_SYM" > "Finite_Set.AC_add.f.AC_2" + "ADD_SYM" > "Fields.linordered_field_class.sign_simps_43" "ADD_SUC" > "Nat.add_Suc_right" "ADD_SUBR2" > "Nat.diff_add_0" "ADD_SUBR" > "HOLLight.hollight.ADD_SUBR" @@ -1944,31 +2264,63 @@ "ADD_SUB" > "Nat.diff_add_inverse2" "ADD_EQ_0" > "Nat.add_is_0" "ADD_CLAUSES" > "HOLLight.hollight.ADD_CLAUSES" - "ADD_ASSOC" > "Finite_Set.AC_add.f.AC_1" + "ADD_ASSOC" > "Fields.linordered_field_class.sign_simps_44" "ADD_AC" > "HOLLight.hollight.ADD_AC" - "ADD_0" > "Presburger.add_right_zero" - "ADD1" > "HOLLight.hollight.ADD1" - "ABS_SIMP" > "Presburger.fm_modd_pinf" + "ADD_0" > "Divides.arithmetic_simps_39" + "ADD1" > "Nat_Numeral.Suc_eq_plus1" + "ABS_SIMP" > "HOL.refl" "ABSORPTION" > "HOLLight.hollight.ABSORPTION" - ">_def" > "HOLLight.hollight.>_def" - ">=_def" > "HOLLight.hollight.>=_def" - "<_def" > "HOLLight.hollight.<_def" - "<=_def" > "HOLLight.hollight.<=_def" + ">_c_def" > "HOLLight.hollight.>_c_def" + ">=_c_def" > "HOLLight.hollight.>=_c_def" + "=_c_def" > "HOLLight.hollight.=_c_def" + "<_c_def" > "HOLLight.hollight.<_c_def" + "<=_c_def" > "HOLLight.hollight.<=_c_def" "$_def" > "HOLLight.hollight.$_def" ignore_thms "TYDEF_sum" "TYDEF_prod" + "TYDEF_option" "TYDEF_num" + "TYDEF_list" "TYDEF_1" - "IND_SUC_0_EXISTS" + "SNDCART_PASTECART" + "SET_OF_LIST_OF_SET" + "REP_ABS_PAIR" + "PASTECART_FST_SND" + "PASTECART_EQ" + "MEM_LIST_OF_SET" + "MEM_ASSOC" + "LIST_OF_SET_PROPERTIES" + "LENGTH_LIST_OF_SET" + "FSTCART_PASTECART" + "FORALL_PASTECART" + "FINITE_SET_OF_LIST" + "EX_MAP" + "EXISTS_PASTECART" + "EL_TL" + "DIMINDEX_HAS_SIZE_FINITE_SUM" + "DIMINDEX_FINITE_SUM" "DEF_one" + "DEF_mk_pair" + "DEF_list_of_set" "DEF__0" + "DEF_ZIP" + "DEF_TL" "DEF_SUC" - "DEF_NUM_REP" + "DEF_SOME" + "DEF_OUTR" + "DEF_OUTL" + "DEF_NONE" + "DEF_NIL" + "DEF_MAP2" + "DEF_ITLIST2" "DEF_INR" "DEF_INL" - "DEF_IND_SUC" - "DEF_IND_0" + "DEF_HD" + "DEF_CONS" + "DEF_ASSOC" + "DEF_," + "ALL_MAP" end