# HG changeset patch # User obua # Date 1127694434 -7200 # Node ID bd59bfd4bf3794a85991e96f3046bcc392508ee7 # Parent 7e417a7cbf9f5b65262ccad0878dbdede478ab9e fixed disambiguation problem diff -r 7e417a7cbf9f -r bd59bfd4bf37 src/HOL/Import/Generate-HOLLight/GenHOLLight.thy --- a/src/HOL/Import/Generate-HOLLight/GenHOLLight.thy Mon Sep 26 02:06:44 2005 +0200 +++ b/src/HOL/Import/Generate-HOLLight/GenHOLLight.thy Mon Sep 26 02:27:14 2005 +0200 @@ -5,10 +5,6 @@ theory GenHOLLight imports "../HOLLightCompat" "../HOL4Syntax" begin; -ML {* reset ProofKernel.debug; *} -ML {* reset Shuffler.debug; *} -ML {* set show_types; set show_sorts; *} - ;import_segment "hollight"; setup_dump "../HOLLight" "HOLLight"; diff -r 7e417a7cbf9f -r bd59bfd4bf37 src/HOL/Import/HOL/HOL4Base.thy --- a/src/HOL/Import/HOL/HOL4Base.thy Mon Sep 26 02:06:44 2005 +0200 +++ b/src/HOL/Import/HOL/HOL4Base.thy Mon Sep 26 02:27:14 2005 +0200 @@ -5,215 +5,259 @@ ;setup_theory bool constdefs - ARB :: "'a" - "ARB == SOME x. True" - -lemma ARB_DEF: "ARB = (SOME x. True)" + ARB :: "'a::type" + "ARB == SOME x::'a::type. True" + +lemma ARB_DEF: "ARB = (SOME x::'a::type. True)" by (import bool ARB_DEF) constdefs - IN :: "'a => ('a => bool) => bool" - "IN == %x f. f x" - -lemma IN_DEF: "IN = (%x f. f x)" + IN :: "'a::type => ('a::type => bool) => bool" + "IN == %(x::'a::type) f::'a::type => bool. f x" + +lemma IN_DEF: "IN = (%(x::'a::type) f::'a::type => bool. f x)" by (import bool IN_DEF) constdefs - RES_FORALL :: "('a => bool) => ('a => bool) => bool" - "RES_FORALL == %p m. ALL x. IN x p --> m x" - -lemma RES_FORALL_DEF: "RES_FORALL = (%p m. ALL x. IN x p --> m x)" + RES_FORALL :: "('a::type => bool) => ('a::type => bool) => bool" + "RES_FORALL == +%(p::'a::type => bool) m::'a::type => bool. ALL x::'a::type. IN x p --> m x" + +lemma RES_FORALL_DEF: "RES_FORALL = +(%(p::'a::type => bool) m::'a::type => bool. + ALL x::'a::type. IN x p --> m x)" by (import bool RES_FORALL_DEF) constdefs - RES_EXISTS :: "('a => bool) => ('a => bool) => bool" - "RES_EXISTS == %p m. EX x. IN x p & m x" - -lemma RES_EXISTS_DEF: "RES_EXISTS = (%p m. EX x. IN x p & m x)" + RES_EXISTS :: "('a::type => bool) => ('a::type => bool) => bool" + "RES_EXISTS == +%(p::'a::type => bool) m::'a::type => bool. EX x::'a::type. IN x p & m x" + +lemma RES_EXISTS_DEF: "RES_EXISTS = +(%(p::'a::type => bool) m::'a::type => bool. EX x::'a::type. IN x p & m x)" by (import bool RES_EXISTS_DEF) constdefs - RES_EXISTS_UNIQUE :: "('a => bool) => ('a => bool) => bool" + RES_EXISTS_UNIQUE :: "('a::type => bool) => ('a::type => bool) => bool" "RES_EXISTS_UNIQUE == -%p m. RES_EXISTS p m & - RES_FORALL p (%x. RES_FORALL p (%y. m x & m y --> x = y))" +%(p::'a::type => bool) m::'a::type => bool. + RES_EXISTS p m & + RES_FORALL p + (%x::'a::type. RES_FORALL p (%y::'a::type. m x & m y --> x = y))" lemma RES_EXISTS_UNIQUE_DEF: "RES_EXISTS_UNIQUE = -(%p m. RES_EXISTS p m & - RES_FORALL p (%x. RES_FORALL p (%y. m x & m y --> x = y)))" +(%(p::'a::type => bool) m::'a::type => bool. + RES_EXISTS p m & + RES_FORALL p + (%x::'a::type. RES_FORALL p (%y::'a::type. m x & m y --> x = y)))" by (import bool RES_EXISTS_UNIQUE_DEF) constdefs - RES_SELECT :: "('a => bool) => ('a => bool) => 'a" - "RES_SELECT == %p m. SOME x. IN x p & m x" - -lemma RES_SELECT_DEF: "RES_SELECT = (%p m. SOME x. IN x p & m x)" + RES_SELECT :: "('a::type => bool) => ('a::type => bool) => 'a::type" + "RES_SELECT == +%(p::'a::type => bool) m::'a::type => bool. SOME x::'a::type. IN x p & m x" + +lemma RES_SELECT_DEF: "RES_SELECT = +(%(p::'a::type => bool) m::'a::type => bool. SOME x::'a::type. IN x p & m x)" by (import bool RES_SELECT_DEF) -lemma EXCLUDED_MIDDLE: "ALL t. t | ~ t" +lemma EXCLUDED_MIDDLE: "ALL t::bool. t | ~ t" by (import bool EXCLUDED_MIDDLE) -lemma FORALL_THM: "All f = All f" +lemma FORALL_THM: "All (f::'a::type => bool) = All f" by (import bool FORALL_THM) -lemma EXISTS_THM: "Ex f = Ex f" +lemma EXISTS_THM: "Ex (f::'a::type => bool) = Ex f" by (import bool EXISTS_THM) -lemma F_IMP: "ALL t. ~ t --> t --> False" +lemma F_IMP: "ALL t::bool. ~ t --> t --> False" by (import bool F_IMP) -lemma NOT_AND: "~ (t & ~ t)" +lemma NOT_AND: "~ ((t::bool) & ~ t)" by (import bool NOT_AND) -lemma AND_CLAUSES: "ALL t. +lemma AND_CLAUSES: "ALL t::bool. (True & t) = t & (t & True) = t & (False & t) = False & (t & False) = False & (t & t) = t" by (import bool AND_CLAUSES) -lemma OR_CLAUSES: "ALL t. +lemma OR_CLAUSES: "ALL t::bool. (True | t) = True & (t | True) = True & (False | t) = t & (t | False) = t & (t | t) = t" by (import bool OR_CLAUSES) -lemma IMP_CLAUSES: "ALL t. +lemma IMP_CLAUSES: "ALL t::bool. (True --> t) = t & (t --> True) = True & (False --> t) = True & (t --> t) = True & (t --> False) = (~ t)" by (import bool IMP_CLAUSES) -lemma NOT_CLAUSES: "(ALL t. (~ ~ t) = t) & (~ True) = False & (~ False) = True" +lemma NOT_CLAUSES: "(ALL t::bool. (~ ~ t) = t) & (~ True) = False & (~ False) = True" by (import bool NOT_CLAUSES) lemma BOOL_EQ_DISTINCT: "True ~= False & False ~= True" by (import bool BOOL_EQ_DISTINCT) -lemma EQ_CLAUSES: "ALL t. +lemma EQ_CLAUSES: "ALL t::bool. (True = t) = t & (t = True) = t & (False = t) = (~ t) & (t = False) = (~ t)" by (import bool EQ_CLAUSES) -lemma COND_CLAUSES: "ALL t1 t2. (if True then t1 else t2) = t1 & (if False then t1 else t2) = t2" +lemma COND_CLAUSES: "ALL (t1::'a::type) t2::'a::type. + (if True then t1 else t2) = t1 & (if False then t1 else t2) = t2" by (import bool COND_CLAUSES) -lemma SELECT_UNIQUE: "ALL P x. (ALL y. P y = (y = x)) --> Eps P = x" +lemma SELECT_UNIQUE: "ALL (P::'a::type => bool) x::'a::type. + (ALL y::'a::type. P y = (y = x)) --> Eps P = x" by (import bool SELECT_UNIQUE) -lemma BOTH_EXISTS_AND_THM: "ALL (P::bool) Q::bool. (EX x::'a. P & Q) = ((EX x::'a. P) & (EX x::'a. Q))" +lemma BOTH_EXISTS_AND_THM: "ALL (P::bool) Q::bool. + (EX x::'a::type. P & Q) = ((EX x::'a::type. P) & (EX x::'a::type. Q))" by (import bool BOTH_EXISTS_AND_THM) lemma BOTH_FORALL_OR_THM: "ALL (P::bool) Q::bool. - (ALL x::'a. P | Q) = ((ALL x::'a. P) | (ALL x::'a. Q))" + (ALL x::'a::type. P | Q) = ((ALL x::'a::type. P) | (ALL x::'a::type. Q))" by (import bool BOTH_FORALL_OR_THM) lemma BOTH_FORALL_IMP_THM: "ALL (P::bool) Q::bool. - (ALL x::'a. P --> Q) = ((EX x::'a. P) --> (ALL x::'a. Q))" + (ALL x::'a::type. P --> Q) = + ((EX x::'a::type. P) --> (ALL x::'a::type. Q))" by (import bool BOTH_FORALL_IMP_THM) lemma BOTH_EXISTS_IMP_THM: "ALL (P::bool) Q::bool. - (EX x::'a. P --> Q) = ((ALL x::'a. P) --> (EX x::'a. Q))" + (EX x::'a::type. P --> Q) = + ((ALL x::'a::type. P) --> (EX x::'a::type. Q))" by (import bool BOTH_EXISTS_IMP_THM) -lemma OR_IMP_THM: "ALL A B. (A = (B | A)) = (B --> A)" +lemma OR_IMP_THM: "ALL (A::bool) B::bool. (A = (B | A)) = (B --> A)" by (import bool OR_IMP_THM) -lemma DE_MORGAN_THM: "ALL A B. (~ (A & B)) = (~ A | ~ B) & (~ (A | B)) = (~ A & ~ B)" +lemma DE_MORGAN_THM: "ALL (A::bool) B::bool. (~ (A & B)) = (~ A | ~ B) & (~ (A | B)) = (~ A & ~ B)" by (import bool DE_MORGAN_THM) -lemma IMP_F_EQ_F: "ALL t. (t --> False) = (t = False)" +lemma IMP_F_EQ_F: "ALL t::bool. (t --> False) = (t = False)" by (import bool IMP_F_EQ_F) -lemma EQ_EXPAND: "ALL t1 t2. (t1 = t2) = (t1 & t2 | ~ t1 & ~ t2)" +lemma EQ_EXPAND: "ALL (t1::bool) t2::bool. (t1 = t2) = (t1 & t2 | ~ t1 & ~ t2)" by (import bool EQ_EXPAND) -lemma COND_RATOR: "ALL b f g x. (if b then f else g) x = (if b then f x else g x)" +lemma COND_RATOR: "ALL (b::bool) (f::'a::type => 'b::type) (g::'a::type => 'b::type) + x::'a::type. (if b then f else g) x = (if b then f x else g x)" by (import bool COND_RATOR) -lemma COND_ABS: "ALL b f g. (%x. if b then f x else g x) = (if b then f else g)" +lemma COND_ABS: "ALL (b::bool) (f::'a::type => 'b::type) g::'a::type => 'b::type. + (%x::'a::type. if b then f x else g x) = (if b then f else g)" by (import bool COND_ABS) -lemma COND_EXPAND: "ALL b t1 t2. (if b then t1 else t2) = ((~ b | t1) & (b | t2))" +lemma COND_EXPAND: "ALL (b::bool) (t1::bool) t2::bool. + (if b then t1 else t2) = ((~ b | t1) & (b | t2))" by (import bool COND_EXPAND) -lemma ONE_ONE_THM: "ALL f. inj f = (ALL x1 x2. f x1 = f x2 --> x1 = x2)" +lemma ONE_ONE_THM: "ALL f::'a::type => 'b::type. + inj f = (ALL (x1::'a::type) x2::'a::type. f x1 = f x2 --> x1 = x2)" by (import bool ONE_ONE_THM) -lemma ABS_REP_THM: "(All::(('a => bool) => bool) => bool) - (%P::'a => bool. +lemma ABS_REP_THM: "(All::(('a::type => bool) => bool) => bool) + (%P::'a::type => bool. (op -->::bool => bool => bool) - ((Ex::(('b => 'a) => bool) => bool) - ((TYPE_DEFINITION::('a => bool) => ('b => 'a) => bool) P)) - ((Ex::(('b => 'a) => bool) => bool) - (%x::'b => 'a. - (Ex::(('a => 'b) => bool) => bool) - (%abs::'a => 'b. + ((Ex::(('b::type => 'a::type) => bool) => bool) + ((TYPE_DEFINITION::('a::type => bool) + => ('b::type => 'a::type) => bool) + P)) + ((Ex::(('b::type => 'a::type) => bool) => bool) + (%x::'b::type => 'a::type. + (Ex::(('a::type => 'b::type) => bool) => bool) + (%abs::'a::type => 'b::type. (op &::bool => bool => bool) - ((All::('b => bool) => bool) - (%a::'b. (op =::'b => 'b => bool) (abs (x a)) a)) - ((All::('a => bool) => bool) - (%r::'a. + ((All::('b::type => bool) => bool) + (%a::'b::type. + (op =::'b::type => 'b::type => bool) (abs (x a)) a)) + ((All::('a::type => bool) => bool) + (%r::'a::type. (op =::bool => bool => bool) (P r) - ((op =::'a => 'a => bool) (x (abs r)) r)))))))" + ((op =::'a::type => 'a::type => bool) (x (abs r)) + r)))))))" by (import bool ABS_REP_THM) -lemma LET_RAND: "(P::'b => bool) (Let (M::'a) (N::'a => 'b)) = (let x::'a = M in P (N x))" +lemma LET_RAND: "(P::'b::type => bool) (Let (M::'a::type) (N::'a::type => 'b::type)) = +(let x::'a::type = M in P (N x))" by (import bool LET_RAND) -lemma LET_RATOR: "Let (M::'a) (N::'a => 'b => 'c) (b::'b) = (let x::'a = M in N x b)" +lemma LET_RATOR: "Let (M::'a::type) (N::'a::type => 'b::type => 'c::type) (b::'b::type) = +(let x::'a::type = M in N x b)" by (import bool LET_RATOR) -lemma SWAP_FORALL_THM: "ALL P. (ALL x. All (P x)) = (ALL y x. P x y)" +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 bool SWAP_FORALL_THM) -lemma SWAP_EXISTS_THM: "ALL P. (EX x. Ex (P x)) = (EX y x. P x y)" +lemma SWAP_EXISTS_THM: "ALL P::'a::type => 'b::type => bool. + (EX x::'a::type. Ex (P x)) = (EX (y::'b::type) x::'a::type. P x y)" by (import bool SWAP_EXISTS_THM) -lemma AND_CONG: "ALL P P' Q Q'. (Q --> P = P') & (P' --> Q = Q') --> (P & Q) = (P' & Q')" +lemma AND_CONG: "ALL (P::bool) (P'::bool) (Q::bool) Q'::bool. + (Q --> P = P') & (P' --> Q = Q') --> (P & Q) = (P' & Q')" by (import bool AND_CONG) -lemma OR_CONG: "ALL P P' Q Q'. (~ Q --> P = P') & (~ P' --> Q = Q') --> (P | Q) = (P' | Q')" +lemma OR_CONG: "ALL (P::bool) (P'::bool) (Q::bool) Q'::bool. + (~ Q --> P = P') & (~ P' --> Q = Q') --> (P | Q) = (P' | Q')" by (import bool OR_CONG) -lemma COND_CONG: "ALL P Q x x' y y'. +lemma COND_CONG: "ALL (P::bool) (Q::bool) (x::'a::type) (x'::'a::type) (y::'a::type) + y'::'a::type. P = Q & (Q --> x = x') & (~ Q --> y = y') --> (if P then x else y) = (if Q then x' else y')" by (import bool COND_CONG) -lemma MONO_COND: "(x --> y) --> (z --> w) --> (if b then x else z) --> (if b then y else w)" +lemma MONO_COND: "((x::bool) --> (y::bool)) --> +((z::bool) --> (w::bool)) --> +(if b::bool then x else z) --> (if b then y else w)" by (import bool MONO_COND) -lemma SKOLEM_THM: "ALL P. (ALL x. Ex (P x)) = (EX f. ALL x. P x (f x))" +lemma SKOLEM_THM: "ALL P::'a::type => 'b::type => bool. + (ALL x::'a::type. Ex (P x)) = + (EX f::'a::type => 'b::type. ALL x::'a::type. P x (f x))" by (import bool SKOLEM_THM) -lemma bool_case_thm: "(ALL (e0::'a) e1::'a. (case True of True => e0 | False => e1) = e0) & -(ALL (e0::'a) e1::'a. (case False of True => e0 | False => e1) = e1)" +lemma bool_case_thm: "(ALL (e0::'a::type) e1::'a::type. + (case True of True => e0 | False => e1) = e0) & +(ALL (e0::'a::type) e1::'a::type. + (case False of True => e0 | False => e1) = e1)" by (import bool bool_case_thm) -lemma bool_case_ID: "ALL x b. (case b of True => x | _ => x) = x" +lemma bool_case_ID: "ALL (x::'a::type) b::bool. (case b of True => x | _ => x) = x" by (import bool bool_case_ID) -lemma boolAxiom: "ALL e0 e1. EX x. x True = e0 & x False = e1" +lemma boolAxiom: "ALL (e0::'a::type) e1::'a::type. + EX x::bool => 'a::type. x True = e0 & x False = e1" by (import bool boolAxiom) -lemma UEXISTS_OR_THM: "ALL P Q. (EX! x. P x | Q x) --> Ex1 P | Ex1 Q" +lemma UEXISTS_OR_THM: "ALL (P::'a::type => bool) Q::'a::type => bool. + (EX! x::'a::type. P x | Q x) --> Ex1 P | Ex1 Q" by (import bool UEXISTS_OR_THM) -lemma UEXISTS_SIMP: "(EX! x::'a. (t::bool)) = (t & (ALL x::'a. All (op = x)))" +lemma UEXISTS_SIMP: "(EX! x::'a::type. (t::bool)) = (t & (ALL x::'a::type. All (op = x)))" by (import bool UEXISTS_SIMP) consts - RES_ABSTRACT :: "('a => bool) => ('a => 'b) => 'a => 'b" - -specification (RES_ABSTRACT) RES_ABSTRACT_DEF: "(ALL (p::'a => bool) (m::'a => 'b) x::'a. + RES_ABSTRACT :: "('a::type => bool) => ('a::type => 'b::type) => 'a::type => 'b::type" + +specification (RES_ABSTRACT) RES_ABSTRACT_DEF: "(ALL (p::'a::type => bool) (m::'a::type => 'b::type) x::'a::type. IN x p --> RES_ABSTRACT p m x = m x) & -(ALL (p::'a => bool) (m1::'a => 'b) m2::'a => 'b. - (ALL x::'a. IN x p --> m1 x = m2 x) --> +(ALL (p::'a::type => bool) (m1::'a::type => 'b::type) + m2::'a::type => 'b::type. + (ALL x::'a::type. IN x p --> m1 x = m2 x) --> RES_ABSTRACT p m1 = RES_ABSTRACT p m2)" by (import bool RES_ABSTRACT_DEF) -lemma BOOL_FUN_CASES_THM: "ALL f. f = (%b. True) | f = (%b. False) | f = (%b. b) | f = Not" +lemma BOOL_FUN_CASES_THM: "ALL f::bool => bool. + f = (%b::bool. True) | + f = (%b::bool. False) | f = (%b::bool. b) | f = Not" by (import bool BOOL_FUN_CASES_THM) -lemma BOOL_FUN_INDUCT: "ALL P. P (%b. True) & P (%b. False) & P (%b. b) & P Not --> All P" +lemma BOOL_FUN_INDUCT: "ALL P::(bool => bool) => bool. + P (%b::bool. True) & P (%b::bool. False) & P (%b::bool. b) & P Not --> + All P" by (import bool BOOL_FUN_INDUCT) ;end_setup @@ -221,67 +265,80 @@ ;setup_theory combin constdefs - K :: "'a => 'b => 'a" - "K == %x y. x" - -lemma K_DEF: "K = (%x y. x)" + K :: "'a::type => 'b::type => 'a::type" + "K == %(x::'a::type) y::'b::type. x" + +lemma K_DEF: "K = (%(x::'a::type) y::'b::type. x)" by (import combin K_DEF) constdefs - S :: "('a => 'b => 'c) => ('a => 'b) => 'a => 'c" - "S == %f g x. f x (g x)" - -lemma S_DEF: "S = (%f g x. f x (g x))" + S :: "('a::type => 'b::type => 'c::type) +=> ('a::type => 'b::type) => 'a::type => 'c::type" + "S == +%(f::'a::type => 'b::type => 'c::type) (g::'a::type => 'b::type) + x::'a::type. f x (g x)" + +lemma S_DEF: "S = +(%(f::'a::type => 'b::type => 'c::type) (g::'a::type => 'b::type) + x::'a::type. f x (g x))" by (import combin S_DEF) constdefs - I :: "'a => 'a" - "(op ==::('a => 'a) => ('a => 'a) => prop) (I::'a => 'a) - ((S::('a => ('a => 'a) => 'a) => ('a => 'a => 'a) => 'a => 'a) - (K::'a => ('a => 'a) => 'a) (K::'a => 'a => 'a))" - -lemma I_DEF: "(op =::('a => 'a) => ('a => 'a) => bool) (I::'a => 'a) - ((S::('a => ('a => 'a) => 'a) => ('a => 'a => 'a) => 'a => 'a) - (K::'a => ('a => 'a) => 'a) (K::'a => 'a => 'a))" + I :: "'a::type => 'a::type" + "(op ==::('a::type => 'a::type) => ('a::type => 'a::type) => prop) + (I::'a::type => 'a::type) + ((S::('a::type => ('a::type => 'a::type) => 'a::type) + => ('a::type => 'a::type => 'a::type) => 'a::type => 'a::type) + (K::'a::type => ('a::type => 'a::type) => 'a::type) + (K::'a::type => 'a::type => 'a::type))" + +lemma I_DEF: "(op =::('a::type => 'a::type) => ('a::type => 'a::type) => bool) + (I::'a::type => 'a::type) + ((S::('a::type => ('a::type => 'a::type) => 'a::type) + => ('a::type => 'a::type => 'a::type) => 'a::type => 'a::type) + (K::'a::type => ('a::type => 'a::type) => 'a::type) + (K::'a::type => 'a::type => 'a::type))" by (import combin I_DEF) constdefs - C :: "('a => 'b => 'c) => 'b => 'a => 'c" - "C == %f x y. f y x" - -lemma C_DEF: "C = (%f x y. f y x)" + C :: "('a::type => 'b::type => 'c::type) => 'b::type => 'a::type => 'c::type" + "C == %(f::'a::type => 'b::type => 'c::type) (x::'b::type) y::'a::type. f y x" + +lemma C_DEF: "C = +(%(f::'a::type => 'b::type => 'c::type) (x::'b::type) y::'a::type. f y x)" by (import combin C_DEF) constdefs - W :: "('a => 'a => 'b) => 'a => 'b" - "W == %f x. f x x" - -lemma W_DEF: "W = (%f x. f x x)" + W :: "('a::type => 'a::type => 'b::type) => 'a::type => 'b::type" + "W == %(f::'a::type => 'a::type => 'b::type) x::'a::type. f x x" + +lemma W_DEF: "W = (%(f::'a::type => 'a::type => 'b::type) x::'a::type. f x x)" by (import combin W_DEF) -lemma I_THM: "ALL x. I x = x" +lemma I_THM: "ALL x::'a::type. I x = x" by (import combin I_THM) -lemma I_o_ID: "ALL f. I o f = f & f o I = f" +lemma I_o_ID: "ALL f::'a::type => 'b::type. I o f = f & f o I = f" by (import combin I_o_ID) ;end_setup ;setup_theory sum -lemma ISL_OR_ISR: "ALL x. ISL x | ISR x" +lemma ISL_OR_ISR: "ALL x::'a::type + 'b::type. ISL x | ISR x" by (import sum ISL_OR_ISR) -lemma INL: "ALL x. ISL x --> Inl (OUTL x) = x" +lemma INL: "ALL x::'a::type + 'b::type. ISL x --> Inl (OUTL x) = x" by (import sum INL) -lemma INR: "ALL x. ISR x --> Inr (OUTR x) = x" +lemma INR: "ALL x::'a::type + 'b::type. ISR x --> Inr (OUTR x) = x" by (import sum INR) -lemma sum_case_cong: "ALL (M::'b + 'c) (M'::'b + 'c) (f::'b => 'a) g::'c => 'a. +lemma sum_case_cong: "ALL (M::'b::type + 'c::type) (M'::'b::type + 'c::type) + (f::'b::type => 'a::type) g::'c::type => 'a::type. M = M' & - (ALL x::'b. M' = Inl x --> f x = (f'::'b => 'a) x) & - (ALL y::'c. M' = Inr y --> g y = (g'::'c => 'a) y) --> + (ALL x::'b::type. M' = Inl x --> f x = (f'::'b::type => 'a::type) x) & + (ALL y::'c::type. M' = Inr y --> g y = (g'::'c::type => 'a::type) y) --> sum_case f g M = sum_case f' g' M'" by (import sum sum_case_cong) @@ -294,167 +351,194 @@ ;setup_theory option lemma option_CLAUSES: "(op &::bool => bool => bool) - ((All::('a => bool) => bool) - (%x::'a. - (All::('a => bool) => bool) - (%y::'a. + ((All::('a::type => bool) => bool) + (%x::'a::type. + (All::('a::type => bool) => bool) + (%y::'a::type. (op =::bool => bool => bool) - ((op =::'a option => 'a option => bool) ((Some::'a ~=> 'a) x) - ((Some::'a ~=> 'a) y)) - ((op =::'a => 'a => bool) x y)))) + ((op =::'a::type option => 'a::type option => bool) + ((Some::'a::type ~=> 'a::type) x) + ((Some::'a::type ~=> 'a::type) y)) + ((op =::'a::type => 'a::type => bool) x y)))) ((op &::bool => bool => bool) - ((All::('a => bool) => bool) - (%x::'a. - (op =::'a => 'a => bool) - ((the::'a option => 'a) ((Some::'a ~=> 'a) x)) x)) + ((All::('a::type => bool) => bool) + (%x::'a::type. + (op =::'a::type => 'a::type => bool) + ((the::'a::type option => 'a::type) + ((Some::'a::type ~=> 'a::type) x)) + x)) ((op &::bool => bool => bool) - ((All::('a => bool) => bool) - (%x::'a. + ((All::('a::type => bool) => bool) + (%x::'a::type. (Not::bool => bool) - ((op =::'a option => 'a option => bool) (None::'a option) - ((Some::'a ~=> 'a) x)))) + ((op =::'a::type option => 'a::type option => bool) + (None::'a::type option) ((Some::'a::type ~=> 'a::type) x)))) ((op &::bool => bool => bool) - ((All::('a => bool) => bool) - (%x::'a. + ((All::('a::type => bool) => bool) + (%x::'a::type. (Not::bool => bool) - ((op =::'a option => 'a option => bool) ((Some::'a ~=> 'a) x) - (None::'a option)))) + ((op =::'a::type option => 'a::type option => bool) + ((Some::'a::type ~=> 'a::type) x) (None::'a::type option)))) ((op &::bool => bool => bool) - ((All::('a => bool) => bool) - (%x::'a. + ((All::('a::type => bool) => bool) + (%x::'a::type. (op =::bool => bool => bool) - ((IS_SOME::'a option => bool) ((Some::'a ~=> 'a) x)) + ((IS_SOME::'a::type option => bool) + ((Some::'a::type ~=> 'a::type) x)) (True::bool))) ((op &::bool => bool => bool) ((op =::bool => bool => bool) - ((IS_SOME::'a option => bool) (None::'a option)) (False::bool)) + ((IS_SOME::'a::type option => bool) (None::'a::type option)) + (False::bool)) ((op &::bool => bool => bool) - ((All::('a option => bool) => bool) - (%x::'a option. + ((All::('a::type option => bool) => bool) + (%x::'a::type option. (op =::bool => bool => bool) - ((IS_NONE::'a option => bool) x) - ((op =::'a option => 'a option => bool) x - (None::'a option)))) + ((IS_NONE::'a::type option => bool) x) + ((op =::'a::type option => 'a::type option => bool) x + (None::'a::type option)))) ((op &::bool => bool => bool) - ((All::('a option => bool) => bool) - (%x::'a option. + ((All::('a::type option => bool) => bool) + (%x::'a::type option. (op =::bool => bool => bool) - ((Not::bool => bool) ((IS_SOME::'a option => bool) x)) - ((op =::'a option => 'a option => bool) x - (None::'a option)))) + ((Not::bool => bool) + ((IS_SOME::'a::type option => bool) x)) + ((op =::'a::type option => 'a::type option => bool) x + (None::'a::type option)))) ((op &::bool => bool => bool) - ((All::('a option => bool) => bool) - (%x::'a option. + ((All::('a::type option => bool) => bool) + (%x::'a::type option. (op -->::bool => bool => bool) - ((IS_SOME::'a option => bool) x) - ((op =::'a option => 'a option => bool) - ((Some::'a ~=> 'a) ((the::'a option => 'a) x)) + ((IS_SOME::'a::type option => bool) x) + ((op =::'a::type option => 'a::type option => bool) + ((Some::'a::type ~=> 'a::type) + ((the::'a::type option => 'a::type) x)) x))) ((op &::bool => bool => bool) - ((All::('a option => bool) => bool) - (%x::'a option. - (op =::'a option => 'a option => bool) - ((option_case::'a option - => ('a ~=> 'a) => 'a option ~=> 'a) - (None::'a option) (Some::'a ~=> 'a) x) + ((All::('a::type option => bool) => bool) + (%x::'a::type option. + (op =::'a::type option => 'a::type option => bool) + ((option_case::'a::type option + => ('a::type ~=> 'a::type) => 'a::type option ~=> 'a::type) + (None::'a::type option) + (Some::'a::type ~=> 'a::type) x) x)) ((op &::bool => bool => bool) - ((All::('a option => bool) => bool) - (%x::'a option. - (op =::'a option => 'a option => bool) - ((option_case::'a option - => ('a ~=> 'a) => 'a option ~=> 'a) - x (Some::'a ~=> 'a) x) + ((All::('a::type option => bool) => bool) + (%x::'a::type option. + (op =::'a::type option + => 'a::type option => bool) + ((option_case::'a::type option + => ('a::type ~=> 'a::type) => 'a::type option ~=> 'a::type) + x (Some::'a::type ~=> 'a::type) x) x)) ((op &::bool => bool => bool) - ((All::('a option => bool) => bool) - (%x::'a option. + ((All::('a::type option => bool) => bool) + (%x::'a::type option. (op -->::bool => bool => bool) - ((IS_NONE::'a option => bool) x) - ((op =::'b => 'b => bool) - ((option_case::'b - => ('a => 'b) => 'a option => 'b) - (e::'b) (f::'a => 'b) x) + ((IS_NONE::'a::type option => bool) x) + ((op =::'b::type => 'b::type => bool) + ((option_case::'b::type + => ('a::type => 'b::type) => 'a::type option => 'b::type) + (e::'b::type) (f::'a::type => 'b::type) x) e))) ((op &::bool => bool => bool) - ((All::('a option => bool) => bool) - (%x::'a option. + ((All::('a::type option => bool) => bool) + (%x::'a::type option. (op -->::bool => bool => bool) - ((IS_SOME::'a option => bool) x) - ((op =::'b => 'b => bool) - ((option_case::'b - => ('a => 'b) => 'a option => 'b) + ((IS_SOME::'a::type option => bool) x) + ((op =::'b::type => 'b::type => bool) + ((option_case::'b::type + => ('a::type => 'b::type) => 'a::type option => 'b::type) e f x) - (f ((the::'a option => 'a) x))))) + (f ((the::'a::type option => 'a::type) + x))))) ((op &::bool => bool => bool) - ((All::('a option => bool) => bool) - (%x::'a option. + ((All::('a::type option => bool) => bool) + (%x::'a::type option. (op -->::bool => bool => bool) - ((IS_SOME::'a option => bool) x) - ((op =::'a option => 'a option => bool) - ((option_case::'a option - => ('a ~=> 'a) => 'a option ~=> 'a) -(ea::'a option) (Some::'a ~=> 'a) x) + ((IS_SOME::'a::type option => bool) x) + ((op =::'a::type option + => 'a::type option => bool) + ((option_case::'a::type option + => ('a::type ~=> 'a::type) => 'a::type option ~=> 'a::type) +(ea::'a::type option) (Some::'a::type ~=> 'a::type) x) x))) ((op &::bool => bool => bool) - ((All::('b => bool) => bool) - (%u::'b. - (All::(('a => 'b) => bool) => bool) - (%f::'a => 'b. - (op =::'b => 'b => bool) - ((option_case::'b => ('a => 'b) => 'a option => 'b) u f - (None::'a option)) + ((All::('b::type => bool) => bool) + (%u::'b::type. + (All::(('a::type => 'b::type) => bool) + => bool) + (%f::'a::type => 'b::type. + (op =::'b::type => 'b::type => bool) + ((option_case::'b::type + => ('a::type => 'b::type) => 'a::type option => 'b::type) + u f (None::'a::type option)) u))) ((op &::bool => bool => bool) - ((All::('b => bool) => bool) - (%u::'b. - (All::(('a => 'b) => bool) => bool) -(%f::'a => 'b. - (All::('a => bool) => bool) - (%x::'a. - (op =::'b => 'b => bool) - ((option_case::'b => ('a => 'b) => 'a option => 'b) u f - ((Some::'a ~=> 'a) x)) + ((All::('b::type => bool) => bool) + (%u::'b::type. + (All::(('a::type => 'b::type) => bool) + => bool) +(%f::'a::type => 'b::type. + (All::('a::type => bool) => bool) + (%x::'a::type. + (op =::'b::type => 'b::type => bool) + ((option_case::'b::type + => ('a::type => 'b::type) + => 'a::type option => 'b::type) + u f ((Some::'a::type ~=> 'a::type) x)) (f x))))) ((op &::bool => bool => bool) - ((All::(('a => 'b) => bool) => bool) - (%f::'a => 'b. - (All::('a => bool) => bool) - (%x::'a. - (op =::'b option => 'b option => bool) - ((option_map::('a => 'b) => 'a option ~=> 'b) f - ((Some::'a ~=> 'a) x)) - ((Some::'b ~=> 'b) (f x))))) + ((All::(('a::type => 'b::type) => bool) + => bool) + (%f::'a::type => 'b::type. + (All::('a::type => bool) => bool) + (%x::'a::type. + (op =::'b::type option => 'b::type option => bool) + ((option_map::('a::type => 'b::type) => 'a::type option ~=> 'b::type) + f ((Some::'a::type ~=> 'a::type) x)) + ((Some::'b::type ~=> 'b::type) (f x))))) ((op &::bool => bool => bool) - ((All::(('a => 'b) => bool) => bool) - (%f::'a => 'b. - (op =::'b option => 'b option => bool) - ((option_map::('a => 'b) => 'a option ~=> 'b) f (None::'a option)) - (None::'b option))) + ((All::(('a::type => 'b::type) => bool) + => bool) + (%f::'a::type => 'b::type. + (op =::'b::type option => 'b::type option => bool) + ((option_map::('a::type => 'b::type) => 'a::type option ~=> 'b::type) f + (None::'a::type option)) + (None::'b::type option))) ((op &::bool => bool => bool) - ((op =::'a option => 'a option => bool) - ((OPTION_JOIN::'a option option ~=> 'a) (None::'a option option)) - (None::'a option)) - ((All::('a option => bool) => bool) - (%x::'a option. - (op =::'a option => 'a option => bool) - ((OPTION_JOIN::'a option option ~=> 'a) - ((Some::'a option ~=> 'a option) x)) + ((op =::'a::type option + => 'a::type option => bool) + ((OPTION_JOIN::'a::type option option ~=> 'a::type) + (None::'a::type option option)) + (None::'a::type option)) + ((All::('a::type option => bool) + => bool) + (%x::'a::type option. + (op =::'a::type option => 'a::type option => bool) + ((OPTION_JOIN::'a::type option option ~=> 'a::type) + ((Some::'a::type option ~=> 'a::type option) x)) x))))))))))))))))))))" by (import option option_CLAUSES) -lemma option_case_compute: "option_case (e::'b) (f::'a => 'b) (x::'a option) = +lemma option_case_compute: "option_case (e::'b::type) (f::'a::type => 'b::type) (x::'a::type option) = (if IS_SOME x then f (the x) else e)" by (import option option_case_compute) -lemma OPTION_MAP_EQ_SOME: "ALL f x y. (option_map f x = Some y) = (EX z. x = Some z & y = f z)" +lemma OPTION_MAP_EQ_SOME: "ALL (f::'a::type => 'b::type) (x::'a::type option) y::'b::type. + (option_map f x = Some y) = (EX z::'a::type. x = Some z & y = f z)" by (import option OPTION_MAP_EQ_SOME) -lemma OPTION_JOIN_EQ_SOME: "ALL x xa. (OPTION_JOIN x = Some xa) = (x = Some (Some xa))" +lemma OPTION_JOIN_EQ_SOME: "ALL (x::'a::type option option) xa::'a::type. + (OPTION_JOIN x = Some xa) = (x = Some (Some xa))" by (import option OPTION_JOIN_EQ_SOME) -lemma option_case_cong: "ALL M M' u f. - M = M' & (M' = None --> u = u') & (ALL x. M' = Some x --> f x = f' x) --> +lemma option_case_cong: "ALL (M::'a::type option) (M'::'a::type option) (u::'b::type) + f::'a::type => 'b::type. + M = M' & + (M' = None --> u = (u'::'b::type)) & + (ALL x::'a::type. M' = Some x --> f x = (f'::'a::type => 'b::type) x) --> option_case u f M = option_case u' f' M'" by (import option option_case_cong) @@ -463,33 +547,33 @@ ;setup_theory marker consts - stmarker :: "'a => 'a" + stmarker :: "'a::type => 'a::type" defs - stmarker_primdef: "stmarker == %x. x" - -lemma stmarker_def: "ALL x. stmarker x = x" + stmarker_primdef: "stmarker == %x::'a::type. x" + +lemma stmarker_def: "ALL x::'a::type. stmarker x = x" by (import marker stmarker_def) -lemma move_left_conj: "ALL x xa xb. +lemma move_left_conj: "ALL (x::bool) (xa::bool) xb::bool. (x & stmarker xb) = (stmarker xb & x) & ((stmarker xb & x) & xa) = (stmarker xb & x & xa) & (x & stmarker xb & xa) = (stmarker xb & x & xa)" by (import marker move_left_conj) -lemma move_right_conj: "ALL x xa xb. +lemma move_right_conj: "ALL (x::bool) (xa::bool) xb::bool. (stmarker xb & x) = (x & stmarker xb) & (x & xa & stmarker xb) = ((x & xa) & stmarker xb) & ((x & stmarker xb) & xa) = ((x & xa) & stmarker xb)" by (import marker move_right_conj) -lemma move_left_disj: "ALL x xa xb. +lemma move_left_disj: "ALL (x::bool) (xa::bool) xb::bool. (x | stmarker xb) = (stmarker xb | x) & ((stmarker xb | x) | xa) = (stmarker xb | x | xa) & (x | stmarker xb | xa) = (stmarker xb | x | xa)" by (import marker move_left_disj) -lemma move_right_disj: "ALL x xa xb. +lemma move_right_disj: "ALL (x::bool) (xa::bool) xb::bool. (stmarker xb | x) = (x | stmarker xb) & (x | xa | stmarker xb) = ((x | xa) | stmarker xb) & ((x | stmarker xb) | xa) = ((x | xa) | stmarker xb)" @@ -500,356 +584,516 @@ ;setup_theory relation constdefs - TC :: "('a => 'a => bool) => 'a => 'a => bool" + TC :: "('a::type => 'a::type => bool) => 'a::type => 'a::type => bool" "TC == -%R a b. - ALL P. - (ALL x y. R x y --> P x y) & (ALL x y z. P x y & P y z --> P x z) --> +%(R::'a::type => 'a::type => bool) (a::'a::type) b::'a::type. + ALL P::'a::type => 'a::type => bool. + (ALL (x::'a::type) y::'a::type. R x y --> P x y) & + (ALL (x::'a::type) (y::'a::type) z::'a::type. + P x y & P y z --> P x z) --> P a b" -lemma TC_DEF: "ALL R a b. +lemma TC_DEF: "ALL (R::'a::type => 'a::type => bool) (a::'a::type) b::'a::type. TC R a b = - (ALL P. - (ALL x y. R x y --> P x y) & (ALL x y z. P x y & P y z --> P x z) --> + (ALL P::'a::type => 'a::type => bool. + (ALL (x::'a::type) y::'a::type. R x y --> P x y) & + (ALL (x::'a::type) (y::'a::type) z::'a::type. + P x y & P y z --> P x z) --> P a b)" by (import relation TC_DEF) constdefs - RTC :: "('a => 'a => bool) => 'a => 'a => bool" + RTC :: "('a::type => 'a::type => bool) => 'a::type => 'a::type => bool" "RTC == -%R a b. - ALL P. (ALL x. P x x) & (ALL x y z. R x y & P y z --> P x z) --> P a b" - -lemma RTC_DEF: "ALL R a b. +%(R::'a::type => 'a::type => bool) (a::'a::type) b::'a::type. + ALL P::'a::type => 'a::type => bool. + (ALL x::'a::type. P x x) & + (ALL (x::'a::type) (y::'a::type) z::'a::type. + R x y & P y z --> P x z) --> + P a b" + +lemma RTC_DEF: "ALL (R::'a::type => 'a::type => bool) (a::'a::type) b::'a::type. RTC R a b = - (ALL P. (ALL x. P x x) & (ALL x y z. R x y & P y z --> P x z) --> P a b)" + (ALL P::'a::type => 'a::type => bool. + (ALL x::'a::type. P x x) & + (ALL (x::'a::type) (y::'a::type) z::'a::type. + R x y & P y z --> P x z) --> + P a b)" by (import relation RTC_DEF) consts - RC :: "('a => 'a => bool) => 'a => 'a => bool" + RC :: "('a::type => 'a::type => bool) => 'a::type => 'a::type => bool" defs - RC_primdef: "RC == %R x y. x = y | R x y" - -lemma RC_def: "ALL R x y. RC R x y = (x = y | R x y)" + RC_primdef: "RC == +%(R::'a::type => 'a::type => bool) (x::'a::type) y::'a::type. x = y | R x y" + +lemma RC_def: "ALL (R::'a::type => 'a::type => bool) (x::'a::type) y::'a::type. + RC R x y = (x = y | R x y)" by (import relation RC_def) consts - transitive :: "('a => 'a => bool) => bool" + transitive :: "('a::type => 'a::type => bool) => bool" defs - transitive_primdef: "transitive == %R. ALL x y z. R x y & R y z --> R x z" - -lemma transitive_def: "ALL R. transitive R = (ALL x y z. R x y & R y z --> R x z)" + transitive_primdef: "transitive == +%R::'a::type => 'a::type => bool. + ALL (x::'a::type) (y::'a::type) z::'a::type. R x y & R y z --> R x z" + +lemma transitive_def: "ALL R::'a::type => 'a::type => bool. + transitive R = + (ALL (x::'a::type) (y::'a::type) z::'a::type. R x y & R y z --> R x z)" by (import relation transitive_def) constdefs - pred_reflexive :: "('a => 'a => bool) => bool" - "pred_reflexive == %R. ALL x. R x x" - -lemma reflexive_def: "ALL R. pred_reflexive R = (ALL x. R x x)" + pred_reflexive :: "('a::type => 'a::type => bool) => bool" + "pred_reflexive == %R::'a::type => 'a::type => bool. ALL x::'a::type. R x x" + +lemma reflexive_def: "ALL R::'a::type => 'a::type => bool. + pred_reflexive R = (ALL x::'a::type. R x x)" by (import relation reflexive_def) -lemma TC_TRANSITIVE: "ALL x. transitive (TC x)" +lemma TC_TRANSITIVE: "ALL x::'a::type => 'a::type => bool. transitive (TC x)" by (import relation TC_TRANSITIVE) -lemma RTC_INDUCT: "ALL x xa. - (ALL x. xa x x) & (ALL xb y z. x xb y & xa y z --> xa xb z) --> - (ALL xb xc. RTC x xb xc --> xa xb xc)" +lemma RTC_INDUCT: "ALL (x::'a::type => 'a::type => bool) xa::'a::type => 'a::type => bool. + (ALL x::'a::type. xa x x) & + (ALL (xb::'a::type) (y::'a::type) z::'a::type. + x xb y & xa y z --> xa xb z) --> + (ALL (xb::'a::type) xc::'a::type. RTC x xb xc --> xa xb xc)" by (import relation RTC_INDUCT) -lemma TC_RULES: "ALL x. - (ALL xa xb. x xa xb --> TC x xa xb) & - (ALL xa xb xc. TC x xa xb & TC x xb xc --> TC x xa xc)" +lemma TC_RULES: "ALL x::'a::type => 'a::type => bool. + (ALL (xa::'a::type) xb::'a::type. x xa xb --> TC x xa xb) & + (ALL (xa::'a::type) (xb::'a::type) xc::'a::type. + TC x xa xb & TC x xb xc --> TC x xa xc)" by (import relation TC_RULES) -lemma RTC_RULES: "ALL x. - (ALL xa. RTC x xa xa) & - (ALL xa xb xc. x xa xb & RTC x xb xc --> RTC x xa xc)" +lemma RTC_RULES: "ALL x::'a::type => 'a::type => bool. + (ALL xa::'a::type. RTC x xa xa) & + (ALL (xa::'a::type) (xb::'a::type) xc::'a::type. + x xa xb & RTC x xb xc --> RTC x xa xc)" by (import relation RTC_RULES) -lemma RTC_STRONG_INDUCT: "ALL R P. - (ALL x. P x x) & (ALL x y z. R x y & RTC R y z & P y z --> P x z) --> - (ALL x y. RTC R x y --> P x y)" +lemma RTC_STRONG_INDUCT: "ALL (R::'a::type => 'a::type => bool) P::'a::type => 'a::type => bool. + (ALL x::'a::type. P x x) & + (ALL (x::'a::type) (y::'a::type) z::'a::type. + R x y & RTC R y z & P y z --> P x z) --> + (ALL (x::'a::type) y::'a::type. RTC R x y --> P x y)" by (import relation RTC_STRONG_INDUCT) -lemma RTC_RTC: "ALL R x y. RTC R x y --> (ALL z. RTC R y z --> RTC R x z)" +lemma RTC_RTC: "ALL (R::'a::type => 'a::type => bool) (x::'a::type) y::'a::type. + RTC R x y --> (ALL z::'a::type. RTC R y z --> RTC R x z)" by (import relation RTC_RTC) -lemma RTC_TRANSITIVE: "ALL x. transitive (RTC x)" +lemma RTC_TRANSITIVE: "ALL x::'a::type => 'a::type => bool. transitive (RTC x)" by (import relation RTC_TRANSITIVE) -lemma RTC_REFLEXIVE: "ALL R. pred_reflexive (RTC R)" +lemma RTC_REFLEXIVE: "ALL R::'a::type => 'a::type => bool. pred_reflexive (RTC R)" by (import relation RTC_REFLEXIVE) -lemma RC_REFLEXIVE: "ALL R. pred_reflexive (RC R)" +lemma RC_REFLEXIVE: "ALL R::'a::type => 'a::type => bool. pred_reflexive (RC R)" by (import relation RC_REFLEXIVE) -lemma TC_SUBSET: "ALL x xa xb. x xa xb --> TC x xa xb" +lemma TC_SUBSET: "ALL (x::'a::type => 'a::type => bool) (xa::'a::type) xb::'a::type. + x xa xb --> TC x xa xb" by (import relation TC_SUBSET) -lemma RTC_SUBSET: "ALL R x y. R x y --> RTC R x y" +lemma RTC_SUBSET: "ALL (R::'a::type => 'a::type => bool) (x::'a::type) y::'a::type. + R x y --> RTC R x y" by (import relation RTC_SUBSET) -lemma RC_SUBSET: "ALL R x y. R x y --> RC R x y" +lemma RC_SUBSET: "ALL (R::'a::type => 'a::type => bool) (x::'a::type) y::'a::type. + R x y --> RC R x y" by (import relation RC_SUBSET) -lemma RC_RTC: "ALL R x y. RC R x y --> RTC R x y" +lemma RC_RTC: "ALL (R::'a::type => 'a::type => bool) (x::'a::type) y::'a::type. + RC R x y --> RTC R x y" by (import relation RC_RTC) -lemma TC_INDUCT: "ALL x xa. - (ALL xb y. x xb y --> xa xb y) & - (ALL x y z. xa x y & xa y z --> xa x z) --> - (ALL xb xc. TC x xb xc --> xa xb xc)" +lemma TC_INDUCT: "ALL (x::'a::type => 'a::type => bool) xa::'a::type => 'a::type => bool. + (ALL (xb::'a::type) y::'a::type. x xb y --> xa xb y) & + (ALL (x::'a::type) (y::'a::type) z::'a::type. + xa x y & xa y z --> xa x z) --> + (ALL (xb::'a::type) xc::'a::type. TC x xb xc --> xa xb xc)" by (import relation TC_INDUCT) -lemma TC_INDUCT_LEFT1: "ALL x xa. - (ALL xb y. x xb y --> xa xb y) & - (ALL xb y z. x xb y & xa y z --> xa xb z) --> - (ALL xb xc. TC x xb xc --> xa xb xc)" +lemma TC_INDUCT_LEFT1: "ALL (x::'a::type => 'a::type => bool) xa::'a::type => 'a::type => bool. + (ALL (xb::'a::type) y::'a::type. x xb y --> xa xb y) & + (ALL (xb::'a::type) (y::'a::type) z::'a::type. + x xb y & xa y z --> xa xb z) --> + (ALL (xb::'a::type) xc::'a::type. TC x xb xc --> xa xb xc)" by (import relation TC_INDUCT_LEFT1) -lemma TC_STRONG_INDUCT: "ALL R P. - (ALL x y. R x y --> P x y) & - (ALL x y z. P x y & P y z & TC R x y & TC R y z --> P x z) --> - (ALL u v. TC R u v --> P u v)" +lemma TC_STRONG_INDUCT: "ALL (R::'a::type => 'a::type => bool) P::'a::type => 'a::type => bool. + (ALL (x::'a::type) y::'a::type. R x y --> P x y) & + (ALL (x::'a::type) (y::'a::type) z::'a::type. + P x y & P y z & TC R x y & TC R y z --> P x z) --> + (ALL (u::'a::type) v::'a::type. TC R u v --> P u v)" by (import relation TC_STRONG_INDUCT) -lemma TC_STRONG_INDUCT_LEFT1: "ALL R P. - (ALL x y. R x y --> P x y) & - (ALL x y z. R x y & P y z & TC R y z --> P x z) --> - (ALL u v. TC R u v --> P u v)" +lemma TC_STRONG_INDUCT_LEFT1: "ALL (R::'a::type => 'a::type => bool) P::'a::type => 'a::type => bool. + (ALL (x::'a::type) y::'a::type. R x y --> P x y) & + (ALL (x::'a::type) (y::'a::type) z::'a::type. + R x y & P y z & TC R y z --> P x z) --> + (ALL (u::'a::type) v::'a::type. TC R u v --> P u v)" by (import relation TC_STRONG_INDUCT_LEFT1) -lemma TC_RTC: "ALL R x y. TC R x y --> RTC R x y" +lemma TC_RTC: "ALL (R::'a::type => 'a::type => bool) (x::'a::type) y::'a::type. + TC R x y --> RTC R x y" by (import relation TC_RTC) -lemma RTC_TC_RC: "ALL R x y. RTC R x y --> RC R x y | TC R x y" +lemma RTC_TC_RC: "ALL (R::'a::type => 'a::type => bool) (x::'a::type) y::'a::type. + RTC R x y --> RC R x y | TC R x y" by (import relation RTC_TC_RC) -lemma TC_RC_EQNS: "ALL R. RC (TC R) = RTC R & TC (RC R) = RTC R" +lemma TC_RC_EQNS: "ALL R::'a::type => 'a::type => bool. RC (TC R) = RTC R & TC (RC R) = RTC R" by (import relation TC_RC_EQNS) -lemma RC_IDEM: "ALL R. RC (RC R) = RC R" +lemma RC_IDEM: "ALL R::'a::type => 'a::type => bool. RC (RC R) = RC R" by (import relation RC_IDEM) -lemma TC_IDEM: "ALL R. TC (TC R) = TC R" +lemma TC_IDEM: "ALL R::'a::type => 'a::type => bool. TC (TC R) = TC R" by (import relation TC_IDEM) -lemma RTC_IDEM: "ALL R. RTC (RTC R) = RTC R" +lemma RTC_IDEM: "ALL R::'a::type => 'a::type => bool. RTC (RTC R) = RTC R" by (import relation RTC_IDEM) -lemma RTC_CASES1: "ALL x xa xb. RTC x xa xb = (xa = xb | (EX u. x xa u & RTC x u xb))" +lemma RTC_CASES1: "ALL (x::'a::type => 'a::type => bool) (xa::'a::type) xb::'a::type. + RTC x xa xb = (xa = xb | (EX u::'a::type. x xa u & RTC x u xb))" by (import relation RTC_CASES1) -lemma RTC_CASES2: "ALL x xa xb. RTC x xa xb = (xa = xb | (EX u. RTC x xa u & x u xb))" +lemma RTC_CASES2: "ALL (x::'a::type => 'a::type => bool) (xa::'a::type) xb::'a::type. + RTC x xa xb = (xa = xb | (EX u::'a::type. RTC x xa u & x u xb))" by (import relation RTC_CASES2) -lemma RTC_CASES_RTC_TWICE: "ALL x xa xb. RTC x xa xb = (EX u. RTC x xa u & RTC x u xb)" +lemma RTC_CASES_RTC_TWICE: "ALL (x::'a::type => 'a::type => bool) (xa::'a::type) xb::'a::type. + RTC x xa xb = (EX u::'a::type. RTC x xa u & RTC x u xb)" by (import relation RTC_CASES_RTC_TWICE) -lemma TC_CASES1: "ALL R x z. TC R x z --> R x z | (EX y. R x y & TC R y z)" +lemma TC_CASES1: "ALL (R::'a::type => 'a::type => bool) (x::'a::type) z::'a::type. + TC R x z --> R x z | (EX y::'a::type. R x y & TC R y z)" by (import relation TC_CASES1) -lemma TC_CASES2: "ALL R x z. TC R x z --> R x z | (EX y. TC R x y & R y z)" +lemma TC_CASES2: "ALL (R::'a::type => 'a::type => bool) (x::'a::type) z::'a::type. + TC R x z --> R x z | (EX y::'a::type. TC R x y & R y z)" by (import relation TC_CASES2) -lemma TC_MONOTONE: "ALL R Q. (ALL x y. R x y --> Q x y) --> (ALL x y. TC R x y --> TC Q x y)" +lemma TC_MONOTONE: "ALL (R::'a::type => 'a::type => bool) Q::'a::type => 'a::type => bool. + (ALL (x::'a::type) y::'a::type. R x y --> Q x y) --> + (ALL (x::'a::type) y::'a::type. TC R x y --> TC Q x y)" by (import relation TC_MONOTONE) -lemma RTC_MONOTONE: "ALL R Q. (ALL x y. R x y --> Q x y) --> (ALL x y. RTC R x y --> RTC Q x y)" +lemma RTC_MONOTONE: "ALL (R::'a::type => 'a::type => bool) Q::'a::type => 'a::type => bool. + (ALL (x::'a::type) y::'a::type. R x y --> Q x y) --> + (ALL (x::'a::type) y::'a::type. RTC R x y --> RTC Q x y)" by (import relation RTC_MONOTONE) constdefs - WF :: "('a => 'a => bool) => bool" - "WF == %R. ALL B. Ex B --> (EX min. B min & (ALL b. R b min --> ~ B b))" - -lemma WF_DEF: "ALL R. WF R = (ALL B. Ex B --> (EX min. B min & (ALL b. R b min --> ~ B b)))" + WF :: "('a::type => 'a::type => bool) => bool" + "WF == +%R::'a::type => 'a::type => bool. + ALL B::'a::type => bool. + Ex B --> + (EX min::'a::type. B min & (ALL b::'a::type. R b min --> ~ B b))" + +lemma WF_DEF: "ALL R::'a::type => 'a::type => bool. + WF R = + (ALL B::'a::type => bool. + Ex B --> + (EX min::'a::type. B min & (ALL b::'a::type. R b min --> ~ B b)))" by (import relation WF_DEF) -lemma WF_INDUCTION_THM: "ALL R. WF R --> (ALL P. (ALL x. (ALL y. R y x --> P y) --> P x) --> All P)" +lemma WF_INDUCTION_THM: "ALL R::'a::type => 'a::type => bool. + WF R --> + (ALL P::'a::type => bool. + (ALL x::'a::type. (ALL y::'a::type. R y x --> P y) --> P x) --> + All P)" by (import relation WF_INDUCTION_THM) -lemma WF_NOT_REFL: "ALL x xa xb. WF x --> x xa xb --> xa ~= xb" +lemma WF_NOT_REFL: "ALL (x::'a::type => 'a::type => bool) (xa::'a::type) xb::'a::type. + WF x --> x xa xb --> xa ~= xb" by (import relation WF_NOT_REFL) constdefs - EMPTY_REL :: "'a => 'a => bool" - "EMPTY_REL == %x y. False" - -lemma EMPTY_REL_DEF: "ALL x y. EMPTY_REL x y = False" + EMPTY_REL :: "'a::type => 'a::type => bool" + "EMPTY_REL == %(x::'a::type) y::'a::type. False" + +lemma EMPTY_REL_DEF: "ALL (x::'a::type) y::'a::type. EMPTY_REL x y = False" by (import relation EMPTY_REL_DEF) lemma WF_EMPTY_REL: "WF EMPTY_REL" by (import relation WF_EMPTY_REL) -lemma WF_SUBSET: "ALL x xa. WF x & (ALL xb y. xa xb y --> x xb y) --> WF xa" +lemma WF_SUBSET: "ALL (x::'a::type => 'a::type => bool) xa::'a::type => 'a::type => bool. + WF x & (ALL (xb::'a::type) y::'a::type. xa xb y --> x xb y) --> WF xa" by (import relation WF_SUBSET) -lemma WF_TC: "ALL R. WF R --> WF (TC R)" +lemma WF_TC: "ALL R::'a::type => 'a::type => bool. WF R --> WF (TC R)" by (import relation WF_TC) consts - inv_image :: "('b => 'b => bool) => ('a => 'b) => 'a => 'a => bool" + inv_image :: "('b::type => 'b::type => bool) +=> ('a::type => 'b::type) => 'a::type => 'a::type => bool" defs inv_image_primdef: "relation.inv_image == -%(R::'b => 'b => bool) (f::'a => 'b) (x::'a) y::'a. R (f x) (f y)" - -lemma inv_image_def: "ALL (R::'b => 'b => bool) f::'a => 'b. - relation.inv_image R f = (%(x::'a) y::'a. R (f x) (f y))" +%(R::'b::type => 'b::type => bool) (f::'a::type => 'b::type) (x::'a::type) + y::'a::type. R (f x) (f y)" + +lemma inv_image_def: "ALL (R::'b::type => 'b::type => bool) f::'a::type => 'b::type. + relation.inv_image R f = (%(x::'a::type) y::'a::type. R (f x) (f y))" by (import relation inv_image_def) -lemma WF_inv_image: "ALL (R::'b => 'b => bool) f::'a => 'b. WF R --> WF (relation.inv_image R f)" +lemma WF_inv_image: "ALL (R::'b::type => 'b::type => bool) f::'a::type => 'b::type. + WF R --> WF (relation.inv_image R f)" by (import relation WF_inv_image) constdefs - RESTRICT :: "('a => 'b) => ('a => 'a => bool) => 'a => 'a => 'b" - "RESTRICT == %f R x y. if R y x then f y else ARB" - -lemma RESTRICT_DEF: "ALL f R x. RESTRICT f R x = (%y. if R y x then f y else ARB)" + RESTRICT :: "('a::type => 'b::type) +=> ('a::type => 'a::type => bool) => 'a::type => 'a::type => 'b::type" + "RESTRICT == +%(f::'a::type => 'b::type) (R::'a::type => 'a::type => bool) (x::'a::type) + y::'a::type. if R y x then f y else ARB" + +lemma RESTRICT_DEF: "ALL (f::'a::type => 'b::type) (R::'a::type => 'a::type => bool) x::'a::type. + RESTRICT f R x = (%y::'a::type. if R y x then f y else ARB)" by (import relation RESTRICT_DEF) -lemma RESTRICT_LEMMA: "ALL x xa xb xc. xa xb xc --> RESTRICT x xa xc xb = x xb" +lemma RESTRICT_LEMMA: "ALL (x::'a::type => 'b::type) (xa::'a::type => 'a::type => bool) + (xb::'a::type) xc::'a::type. xa xb xc --> RESTRICT x xa xc xb = x xb" by (import relation RESTRICT_LEMMA) consts - approx :: "('a => 'a => bool) => (('a => 'b) => 'a => 'b) => 'a => ('a => 'b) => bool" + approx :: "('a::type => 'a::type => bool) +=> (('a::type => 'b::type) => 'a::type => 'b::type) + => 'a::type => ('a::type => 'b::type) => bool" defs - approx_primdef: "approx == %R M x f. f = RESTRICT (%y. M (RESTRICT f R y) y) R x" - -lemma approx_def: "ALL R M x f. approx R M x f = (f = RESTRICT (%y. M (RESTRICT f R y) y) R x)" + approx_primdef: "approx == +%(R::'a::type => 'a::type => bool) + (M::('a::type => 'b::type) => 'a::type => 'b::type) (x::'a::type) + f::'a::type => 'b::type. + f = RESTRICT (%y::'a::type. M (RESTRICT f R y) y) R x" + +lemma approx_def: "ALL (R::'a::type => 'a::type => bool) + (M::('a::type => 'b::type) => 'a::type => 'b::type) (x::'a::type) + f::'a::type => 'b::type. + approx R M x f = (f = RESTRICT (%y::'a::type. M (RESTRICT f R y) y) R x)" by (import relation approx_def) consts - the_fun :: "('a => 'a => bool) => (('a => 'b) => 'a => 'b) => 'a => 'a => 'b" + the_fun :: "('a::type => 'a::type => bool) +=> (('a::type => 'b::type) => 'a::type => 'b::type) + => 'a::type => 'a::type => 'b::type" defs - the_fun_primdef: "the_fun == %R M x. Eps (approx R M x)" - -lemma the_fun_def: "ALL R M x. the_fun R M x = Eps (approx R M x)" + the_fun_primdef: "the_fun == +%(R::'a::type => 'a::type => bool) + (M::('a::type => 'b::type) => 'a::type => 'b::type) x::'a::type. + Eps (approx R M x)" + +lemma the_fun_def: "ALL (R::'a::type => 'a::type => bool) + (M::('a::type => 'b::type) => 'a::type => 'b::type) x::'a::type. + the_fun R M x = Eps (approx R M x)" by (import relation the_fun_def) constdefs - WFREC :: "('a => 'a => bool) => (('a => 'b) => 'a => 'b) => 'a => 'b" + WFREC :: "('a::type => 'a::type => bool) +=> (('a::type => 'b::type) => 'a::type => 'b::type) => 'a::type => 'b::type" "WFREC == -%R M x. M (RESTRICT (the_fun (TC R) (%f v. M (RESTRICT f R v) v) x) R x) x" - -lemma WFREC_DEF: "ALL R M. +%(R::'a::type => 'a::type => bool) + (M::('a::type => 'b::type) => 'a::type => 'b::type) x::'a::type. + M (RESTRICT + (the_fun (TC R) + (%(f::'a::type => 'b::type) v::'a::type. M (RESTRICT f R v) v) x) + R x) + x" + +lemma WFREC_DEF: "ALL (R::'a::type => 'a::type => bool) + M::('a::type => 'b::type) => 'a::type => 'b::type. WFREC R M = - (%x. M (RESTRICT (the_fun (TC R) (%f v. M (RESTRICT f R v) v) x) R x) x)" + (%x::'a::type. + M (RESTRICT + (the_fun (TC R) + (%(f::'a::type => 'b::type) v::'a::type. M (RESTRICT f R v) v) + x) + R x) + x)" by (import relation WFREC_DEF) -lemma WFREC_THM: "ALL R M. WF R --> (ALL x. WFREC R M x = M (RESTRICT (WFREC R M) R x) x)" +lemma WFREC_THM: "ALL (R::'a::type => 'a::type => bool) + M::('a::type => 'b::type) => 'a::type => 'b::type. + WF R --> (ALL x::'a::type. WFREC R M x = M (RESTRICT (WFREC R M) R x) x)" by (import relation WFREC_THM) -lemma WFREC_COROLLARY: "ALL M R f. f = WFREC R M --> WF R --> (ALL x. f x = M (RESTRICT f R x) x)" +lemma WFREC_COROLLARY: "ALL (M::('a::type => 'b::type) => 'a::type => 'b::type) + (R::'a::type => 'a::type => bool) f::'a::type => 'b::type. + f = WFREC R M --> WF R --> (ALL x::'a::type. f x = M (RESTRICT f R x) x)" by (import relation WFREC_COROLLARY) -lemma WF_RECURSION_THM: "ALL R. WF R --> (ALL M. EX! f. ALL x. f x = M (RESTRICT f R x) x)" +lemma WF_RECURSION_THM: "ALL R::'a::type => 'a::type => bool. + WF R --> + (ALL M::('a::type => 'b::type) => 'a::type => 'b::type. + EX! f::'a::type => 'b::type. + ALL x::'a::type. f x = M (RESTRICT f R x) x)" by (import relation WF_RECURSION_THM) ;end_setup ;setup_theory pair -lemma CURRY_ONE_ONE_THM: "(curry f = curry g) = (f = g)" +lemma CURRY_ONE_ONE_THM: "(curry (f::'a::type * 'b::type => 'c::type) = + curry (g::'a::type * 'b::type => 'c::type)) = +(f = g)" by (import pair CURRY_ONE_ONE_THM) lemma UNCURRY_ONE_ONE_THM: "(op =::bool => bool => bool) - ((op =::('a * 'b => 'c) => ('a * 'b => 'c) => bool) - ((split::('a => 'b => 'c) => 'a * 'b => 'c) (f::'a => 'b => 'c)) - ((split::('a => 'b => 'c) => 'a * 'b => 'c) (g::'a => 'b => 'c))) - ((op =::('a => 'b => 'c) => ('a => 'b => 'c) => bool) f g)" + ((op =::('a::type * 'b::type => 'c::type) + => ('a::type * 'b::type => 'c::type) => bool) + ((split::('a::type => 'b::type => 'c::type) + => 'a::type * 'b::type => 'c::type) + (f::'a::type => 'b::type => 'c::type)) + ((split::('a::type => 'b::type => 'c::type) + => 'a::type * 'b::type => 'c::type) + (g::'a::type => 'b::type => 'c::type))) + ((op =::('a::type => 'b::type => 'c::type) + => ('a::type => 'b::type => 'c::type) => bool) + f g)" by (import pair UNCURRY_ONE_ONE_THM) -lemma pair_Axiom: "ALL f. EX x. ALL xa y. x (xa, y) = f xa y" +lemma pair_Axiom: "ALL f::'a::type => 'b::type => 'c::type. + EX x::'a::type * 'b::type => 'c::type. + ALL (xa::'a::type) y::'b::type. x (xa, y) = f xa y" by (import pair pair_Axiom) -lemma UNCURRY_CONG: "ALL M M' f. - M = M' & (ALL x y. M' = (x, y) --> f x y = f' x y) --> +lemma UNCURRY_CONG: "ALL (M::'a::type * 'b::type) (M'::'a::type * 'b::type) + f::'a::type => 'b::type => 'c::type. + M = M' & + (ALL (x::'a::type) y::'b::type. + M' = (x, y) --> + f x y = (f'::'a::type => 'b::type => 'c::type) x y) --> split f M = split f' M'" by (import pair UNCURRY_CONG) -lemma ELIM_PEXISTS: "(EX p. P (fst p) (snd p)) = (EX p1. Ex (P p1))" +lemma ELIM_PEXISTS: "(EX p::'a::type * 'b::type. + (P::'a::type => 'b::type => bool) (fst p) (snd p)) = +(EX p1::'a::type. Ex (P p1))" by (import pair ELIM_PEXISTS) -lemma ELIM_PFORALL: "(ALL p. P (fst p) (snd p)) = (ALL p1. All (P p1))" +lemma ELIM_PFORALL: "(ALL p::'a::type * 'b::type. + (P::'a::type => 'b::type => bool) (fst p) (snd p)) = +(ALL p1::'a::type. All (P p1))" by (import pair ELIM_PFORALL) -lemma PFORALL_THM: "(All::(('a => 'b => bool) => bool) => bool) - (%x::'a => 'b => bool. +lemma PFORALL_THM: "(All::(('a::type => 'b::type => bool) => bool) => bool) + (%x::'a::type => 'b::type => bool. (op =::bool => bool => bool) - ((All::('a => bool) => bool) - (%xa::'a. (All::('b => bool) => bool) (x xa))) - ((All::('a * 'b => bool) => bool) - ((split::('a => 'b => bool) => 'a * 'b => bool) x)))" + ((All::('a::type => bool) => bool) + (%xa::'a::type. (All::('b::type => bool) => bool) (x xa))) + ((All::('a::type * 'b::type => bool) => bool) + ((split::('a::type => 'b::type => bool) + => 'a::type * 'b::type => bool) + x)))" by (import pair PFORALL_THM) -lemma PEXISTS_THM: "(All::(('a => 'b => bool) => bool) => bool) - (%x::'a => 'b => bool. +lemma PEXISTS_THM: "(All::(('a::type => 'b::type => bool) => bool) => bool) + (%x::'a::type => 'b::type => bool. (op =::bool => bool => bool) - ((Ex::('a => bool) => bool) - (%xa::'a. (Ex::('b => bool) => bool) (x xa))) - ((Ex::('a * 'b => bool) => bool) - ((split::('a => 'b => bool) => 'a * 'b => bool) x)))" + ((Ex::('a::type => bool) => bool) + (%xa::'a::type. (Ex::('b::type => bool) => bool) (x xa))) + ((Ex::('a::type * 'b::type => bool) => bool) + ((split::('a::type => 'b::type => bool) + => 'a::type * 'b::type => bool) + x)))" by (import pair PEXISTS_THM) -lemma LET2_RAND: "(All::(('c => 'd) => bool) => bool) - (%x::'c => 'd. - (All::('a * 'b => bool) => bool) - (%xa::'a * 'b. - (All::(('a => 'b => 'c) => bool) => bool) - (%xb::'a => 'b => 'c. - (op =::'d => 'd => bool) - (x ((Let::'a * 'b => ('a * 'b => 'c) => 'c) xa - ((split::('a => 'b => 'c) => 'a * 'b => 'c) xb))) - ((Let::'a * 'b => ('a * 'b => 'd) => 'd) xa - ((split::('a => 'b => 'd) => 'a * 'b => 'd) - (%(xa::'a) y::'b. x (xb xa y)))))))" +lemma LET2_RAND: "(All::(('c::type => 'd::type) => bool) => bool) + (%x::'c::type => 'd::type. + (All::('a::type * 'b::type => bool) => bool) + (%xa::'a::type * 'b::type. + (All::(('a::type => 'b::type => 'c::type) => bool) => bool) + (%xb::'a::type => 'b::type => 'c::type. + (op =::'d::type => 'd::type => bool) + (x ((Let::'a::type * 'b::type + => ('a::type * 'b::type => 'c::type) => 'c::type) + xa ((split::('a::type => 'b::type => 'c::type) + => 'a::type * 'b::type => 'c::type) + xb))) + ((Let::'a::type * 'b::type + => ('a::type * 'b::type => 'd::type) => 'd::type) + xa ((split::('a::type => 'b::type => 'd::type) + => 'a::type * 'b::type => 'd::type) + (%(xa::'a::type) y::'b::type. x (xb xa y)))))))" by (import pair LET2_RAND) -lemma LET2_RATOR: "(All::('a1 * 'a2 => bool) => bool) - (%x::'a1 * 'a2. - (All::(('a1 => 'a2 => 'b => 'c) => bool) => bool) - (%xa::'a1 => 'a2 => 'b => 'c. - (All::('b => bool) => bool) - (%xb::'b. - (op =::'c => 'c => bool) - ((Let::'a1 * 'a2 => ('a1 * 'a2 => 'b => 'c) => 'b => 'c) x - ((split::('a1 => 'a2 => 'b => 'c) - => 'a1 * 'a2 => 'b => 'c) - xa) +lemma LET2_RATOR: "(All::('a1::type * 'a2::type => bool) => bool) + (%x::'a1::type * 'a2::type. + (All::(('a1::type => 'a2::type => 'b::type => 'c::type) => bool) + => bool) + (%xa::'a1::type => 'a2::type => 'b::type => 'c::type. + (All::('b::type => bool) => bool) + (%xb::'b::type. + (op =::'c::type => 'c::type => bool) + ((Let::'a1::type * 'a2::type + => ('a1::type * 'a2::type => 'b::type => 'c::type) + => 'b::type => 'c::type) + x ((split::('a1::type + => 'a2::type => 'b::type => 'c::type) + => 'a1::type * 'a2::type + => 'b::type => 'c::type) + xa) xb) - ((Let::'a1 * 'a2 => ('a1 * 'a2 => 'c) => 'c) x - ((split::('a1 => 'a2 => 'c) => 'a1 * 'a2 => 'c) - (%(x::'a1) y::'a2. xa x y xb))))))" + ((Let::'a1::type * 'a2::type + => ('a1::type * 'a2::type => 'c::type) => 'c::type) + x ((split::('a1::type => 'a2::type => 'c::type) + => 'a1::type * 'a2::type => 'c::type) + (%(x::'a1::type) y::'a2::type. xa x y xb))))))" by (import pair LET2_RATOR) -lemma pair_case_cong: "ALL x xa xb. - x = xa & (ALL x y. xa = (x, y) --> xb x y = f' x y) --> +lemma pair_case_cong: "ALL (x::'a::type * 'b::type) (xa::'a::type * 'b::type) + xb::'a::type => 'b::type => 'c::type. + x = xa & + (ALL (x::'a::type) y::'b::type. + xa = (x, y) --> + xb x y = (f'::'a::type => 'b::type => 'c::type) x y) --> split xb x = split f' xa" by (import pair pair_case_cong) constdefs - LEX :: "('a => 'a => bool) => ('b => 'b => bool) => 'a * 'b => 'a * 'b => bool" - "LEX == %R1 R2 (s, t) (u, v). R1 s u | s = u & R2 t v" - -lemma LEX_DEF: "ALL R1 R2. LEX R1 R2 = (%(s, t) (u, v). R1 s u | s = u & R2 t v)" + LEX :: "('a::type => 'a::type => bool) +=> ('b::type => 'b::type => bool) + => 'a::type * 'b::type => 'a::type * 'b::type => bool" + "LEX == +%(R1::'a::type => 'a::type => bool) (R2::'b::type => 'b::type => bool) + (s::'a::type, t::'b::type) (u::'a::type, v::'b::type). + R1 s u | s = u & R2 t v" + +lemma LEX_DEF: "ALL (R1::'a::type => 'a::type => bool) R2::'b::type => 'b::type => bool. + LEX R1 R2 = + (%(s::'a::type, t::'b::type) (u::'a::type, v::'b::type). + R1 s u | s = u & R2 t v)" by (import pair LEX_DEF) -lemma WF_LEX: "ALL x xa. WF x & WF xa --> WF (LEX x xa)" +lemma WF_LEX: "ALL (x::'a::type => 'a::type => bool) xa::'b::type => 'b::type => bool. + WF x & WF xa --> WF (LEX x xa)" by (import pair WF_LEX) constdefs - RPROD :: "('a => 'a => bool) => ('b => 'b => bool) => 'a * 'b => 'a * 'b => bool" - "RPROD == %R1 R2 (s, t) (u, v). R1 s u & R2 t v" - -lemma RPROD_DEF: "ALL R1 R2. RPROD R1 R2 = (%(s, t) (u, v). R1 s u & R2 t v)" + RPROD :: "('a::type => 'a::type => bool) +=> ('b::type => 'b::type => bool) + => 'a::type * 'b::type => 'a::type * 'b::type => bool" + "RPROD == +%(R1::'a::type => 'a::type => bool) (R2::'b::type => 'b::type => bool) + (s::'a::type, t::'b::type) (u::'a::type, v::'b::type). R1 s u & R2 t v" + +lemma RPROD_DEF: "ALL (R1::'a::type => 'a::type => bool) R2::'b::type => 'b::type => bool. + RPROD R1 R2 = + (%(s::'a::type, t::'b::type) (u::'a::type, v::'b::type). R1 s u & R2 t v)" by (import pair RPROD_DEF) -lemma WF_RPROD: "ALL R Q. WF R & WF Q --> WF (RPROD R Q)" +lemma WF_RPROD: "ALL (R::'a::type => 'a::type => bool) Q::'b::type => 'b::type => bool. + WF R & WF Q --> WF (RPROD R Q)" by (import pair WF_RPROD) ;end_setup @@ -860,152 +1104,184 @@ ;setup_theory prim_rec -lemma LESS_0_0: "0 < Suc 0" +lemma LESS_0_0: "(0::nat) < Suc (0::nat)" by (import prim_rec LESS_0_0) -lemma LESS_LEMMA1: "ALL x xa. x < Suc xa --> x = xa | x < xa" +lemma LESS_LEMMA1: "ALL (x::nat) xa::nat. x < Suc xa --> x = xa | x < xa" by (import prim_rec LESS_LEMMA1) -lemma LESS_LEMMA2: "ALL m n. m = n | m < n --> m < Suc n" +lemma LESS_LEMMA2: "ALL (m::nat) n::nat. m = n | m < n --> m < Suc n" by (import prim_rec LESS_LEMMA2) -lemma LESS_THM: "ALL m n. (m < Suc n) = (m = n | m < n)" +lemma LESS_THM: "ALL (m::nat) n::nat. (m < Suc n) = (m = n | m < n)" by (import prim_rec LESS_THM) -lemma LESS_SUC_IMP: "ALL x xa. x < Suc xa --> x ~= xa --> x < xa" +lemma LESS_SUC_IMP: "ALL (x::nat) xa::nat. x < Suc xa --> x ~= xa --> x < xa" by (import prim_rec LESS_SUC_IMP) -lemma EQ_LESS: "ALL n. Suc m = n --> m < n" +lemma EQ_LESS: "ALL n::nat. Suc (m::nat) = n --> m < n" by (import prim_rec EQ_LESS) lemma NOT_LESS_EQ: "ALL (m::nat) n::nat. m = n --> ~ m < n" by (import prim_rec NOT_LESS_EQ) constdefs - SIMP_REC_REL :: "(nat => 'a) => 'a => ('a => 'a) => nat => bool" - "(op ==::((nat => 'a) => 'a => ('a => 'a) => nat => bool) - => ((nat => 'a) => 'a => ('a => 'a) => nat => bool) => prop) - (SIMP_REC_REL::(nat => 'a) => 'a => ('a => 'a) => nat => bool) - (%(fun::nat => 'a) (x::'a) (f::'a => 'a) n::nat. + SIMP_REC_REL :: "(nat => 'a::type) => 'a::type => ('a::type => 'a::type) => nat => bool" + "(op ==::((nat => 'a::type) + => 'a::type => ('a::type => 'a::type) => nat => bool) + => ((nat => 'a::type) + => 'a::type => ('a::type => 'a::type) => nat => bool) + => prop) + (SIMP_REC_REL::(nat => 'a::type) + => 'a::type => ('a::type => 'a::type) => nat => bool) + (%(fun::nat => 'a::type) (x::'a::type) (f::'a::type => 'a::type) n::nat. (op &::bool => bool => bool) - ((op =::'a => 'a => bool) (fun (0::nat)) x) + ((op =::'a::type => 'a::type => bool) (fun (0::nat)) x) ((All::(nat => bool) => bool) (%m::nat. (op -->::bool => bool => bool) ((op <::nat => nat => bool) m n) - ((op =::'a => 'a => bool) (fun ((Suc::nat => nat) m)) - (f (fun m))))))" - -lemma SIMP_REC_REL: "(All::((nat => 'a) => bool) => bool) - (%fun::nat => 'a. - (All::('a => bool) => bool) - (%x::'a. - (All::(('a => 'a) => bool) => bool) - (%f::'a => 'a. + ((op =::'a::type => 'a::type => bool) + (fun ((Suc::nat => nat) m)) (f (fun m))))))" + +lemma SIMP_REC_REL: "(All::((nat => 'a::type) => bool) => bool) + (%fun::nat => 'a::type. + (All::('a::type => bool) => bool) + (%x::'a::type. + (All::(('a::type => 'a::type) => bool) => bool) + (%f::'a::type => 'a::type. (All::(nat => bool) => bool) (%n::nat. (op =::bool => bool => bool) - ((SIMP_REC_REL::(nat => 'a) - => 'a => ('a => 'a) => nat => bool) + ((SIMP_REC_REL::(nat => 'a::type) + => 'a::type + => ('a::type => 'a::type) => nat => bool) fun x f n) ((op &::bool => bool => bool) - ((op =::'a => 'a => bool) (fun (0::nat)) x) + ((op =::'a::type => 'a::type => bool) (fun (0::nat)) + x) ((All::(nat => bool) => bool) (%m::nat. (op -->::bool => bool => bool) ((op <::nat => nat => bool) m n) - ((op =::'a => 'a => bool) + ((op =::'a::type => 'a::type => bool) (fun ((Suc::nat => nat) m)) (f (fun m))))))))))" by (import prim_rec SIMP_REC_REL) -lemma SIMP_REC_EXISTS: "ALL x f n. EX fun. SIMP_REC_REL fun x f n" +lemma SIMP_REC_EXISTS: "ALL (x::'a::type) (f::'a::type => 'a::type) n::nat. + EX fun::nat => 'a::type. SIMP_REC_REL fun x f n" by (import prim_rec SIMP_REC_EXISTS) -lemma SIMP_REC_REL_UNIQUE: "ALL x xa xb xc xd xe. +lemma SIMP_REC_REL_UNIQUE: "ALL (x::'a::type) (xa::'a::type => 'a::type) (xb::nat => 'a::type) + (xc::nat => 'a::type) (xd::nat) xe::nat. SIMP_REC_REL xb x xa xd & SIMP_REC_REL xc x xa xe --> - (ALL n. n < xd & n < xe --> xb n = xc n)" + (ALL n::nat. n < xd & n < xe --> xb n = xc n)" by (import prim_rec SIMP_REC_REL_UNIQUE) -lemma SIMP_REC_REL_UNIQUE_RESULT: "ALL x f n. EX! y. EX g. SIMP_REC_REL g x f (Suc n) & y = g n" +lemma SIMP_REC_REL_UNIQUE_RESULT: "ALL (x::'a::type) (f::'a::type => 'a::type) n::nat. + EX! y::'a::type. + EX g::nat => 'a::type. SIMP_REC_REL g x f (Suc n) & y = g n" by (import prim_rec SIMP_REC_REL_UNIQUE_RESULT) consts - SIMP_REC :: "'a => ('a => 'a) => nat => 'a" - -specification (SIMP_REC) SIMP_REC: "ALL x f' n. EX g. SIMP_REC_REL g x f' (Suc n) & SIMP_REC x f' n = g n" + SIMP_REC :: "'a::type => ('a::type => 'a::type) => nat => 'a::type" + +specification (SIMP_REC) SIMP_REC: "ALL (x::'a::type) (f'::'a::type => 'a::type) n::nat. + EX g::nat => 'a::type. + SIMP_REC_REL g x f' (Suc n) & SIMP_REC x f' n = g n" by (import prim_rec SIMP_REC) -lemma LESS_SUC_SUC: "ALL m. m < Suc m & m < Suc (Suc m)" +lemma LESS_SUC_SUC: "ALL m::nat. m < Suc m & m < Suc (Suc m)" by (import prim_rec LESS_SUC_SUC) -lemma SIMP_REC_THM: "ALL x f. - SIMP_REC x f 0 = x & (ALL m. SIMP_REC x f (Suc m) = f (SIMP_REC x f m))" +lemma SIMP_REC_THM: "ALL (x::'a::type) f::'a::type => 'a::type. + SIMP_REC x f (0::nat) = x & + (ALL m::nat. SIMP_REC x f (Suc m) = f (SIMP_REC x f m))" by (import prim_rec SIMP_REC_THM) constdefs PRE :: "nat => nat" - "PRE == %m. if m = 0 then 0 else SOME n. m = Suc n" - -lemma PRE_DEF: "ALL m. PRE m = (if m = 0 then 0 else SOME n. m = Suc n)" + "PRE == %m::nat. if m = (0::nat) then 0::nat else SOME n::nat. m = Suc n" + +lemma PRE_DEF: "ALL m::nat. + PRE m = (if m = (0::nat) then 0::nat else SOME n::nat. m = Suc n)" by (import prim_rec PRE_DEF) -lemma PRE: "PRE 0 = 0 & (ALL m. PRE (Suc m) = m)" +lemma PRE: "PRE (0::nat) = (0::nat) & (ALL m::nat. PRE (Suc m) = m)" by (import prim_rec PRE) constdefs - PRIM_REC_FUN :: "'a => ('a => nat => 'a) => nat => nat => 'a" - "PRIM_REC_FUN == %x f. SIMP_REC (%n. x) (%fun n. f (fun (PRE n)) n)" - -lemma PRIM_REC_FUN: "ALL x f. PRIM_REC_FUN x f = SIMP_REC (%n. x) (%fun n. f (fun (PRE n)) n)" + PRIM_REC_FUN :: "'a::type => ('a::type => nat => 'a::type) => nat => nat => 'a::type" + "PRIM_REC_FUN == +%(x::'a::type) f::'a::type => nat => 'a::type. + SIMP_REC (%n::nat. x) (%(fun::nat => 'a::type) n::nat. f (fun (PRE n)) n)" + +lemma PRIM_REC_FUN: "ALL (x::'a::type) f::'a::type => nat => 'a::type. + PRIM_REC_FUN x f = + SIMP_REC (%n::nat. x) (%(fun::nat => 'a::type) n::nat. f (fun (PRE n)) n)" by (import prim_rec PRIM_REC_FUN) -lemma PRIM_REC_EQN: "ALL x f. - (ALL n. PRIM_REC_FUN x f 0 n = x) & - (ALL m n. PRIM_REC_FUN x f (Suc m) n = f (PRIM_REC_FUN x f m (PRE n)) n)" +lemma PRIM_REC_EQN: "ALL (x::'a::type) f::'a::type => nat => 'a::type. + (ALL n::nat. PRIM_REC_FUN x f (0::nat) n = x) & + (ALL (m::nat) n::nat. + PRIM_REC_FUN x f (Suc m) n = f (PRIM_REC_FUN x f m (PRE n)) n)" by (import prim_rec PRIM_REC_EQN) constdefs - PRIM_REC :: "'a => ('a => nat => 'a) => nat => 'a" - "PRIM_REC == %x f m. PRIM_REC_FUN x f m (PRE m)" - -lemma PRIM_REC: "ALL x f m. PRIM_REC x f m = PRIM_REC_FUN x f m (PRE m)" + PRIM_REC :: "'a::type => ('a::type => nat => 'a::type) => nat => 'a::type" + "PRIM_REC == +%(x::'a::type) (f::'a::type => nat => 'a::type) m::nat. + PRIM_REC_FUN x f m (PRE m)" + +lemma PRIM_REC: "ALL (x::'a::type) (f::'a::type => nat => 'a::type) m::nat. + PRIM_REC x f m = PRIM_REC_FUN x f m (PRE m)" by (import prim_rec PRIM_REC) -lemma PRIM_REC_THM: "ALL x f. - PRIM_REC x f 0 = x & (ALL m. PRIM_REC x f (Suc m) = f (PRIM_REC x f m) m)" +lemma PRIM_REC_THM: "ALL (x::'a::type) f::'a::type => nat => 'a::type. + PRIM_REC x f (0::nat) = x & + (ALL m::nat. PRIM_REC x f (Suc m) = f (PRIM_REC x f m) m)" by (import prim_rec PRIM_REC_THM) -lemma DC: "ALL P R a. - P a & (ALL x. P x --> (EX y. P y & R x y)) --> - (EX x. x 0 = a & (ALL n. P (x n) & R (x n) (x (Suc n))))" +lemma DC: "ALL (P::'a::type => bool) (R::'a::type => 'a::type => bool) a::'a::type. + P a & (ALL x::'a::type. P x --> (EX y::'a::type. P y & R x y)) --> + (EX x::nat => 'a::type. + x (0::nat) = a & (ALL n::nat. P (x n) & R (x n) (x (Suc n))))" by (import prim_rec DC) -lemma num_Axiom_old: "ALL e f. EX! fn1. fn1 0 = e & (ALL n. fn1 (Suc n) = f (fn1 n) n)" +lemma num_Axiom_old: "ALL (e::'a::type) f::'a::type => nat => 'a::type. + EX! fn1::nat => 'a::type. + fn1 (0::nat) = e & (ALL n::nat. fn1 (Suc n) = f (fn1 n) n)" by (import prim_rec num_Axiom_old) -lemma num_Axiom: "ALL e f. EX x. x 0 = e & (ALL n. x (Suc n) = f n (x n))" +lemma num_Axiom: "ALL (e::'a::type) f::nat => 'a::type => 'a::type. + EX x::nat => 'a::type. + x (0::nat) = e & (ALL n::nat. x (Suc n) = f n (x n))" by (import prim_rec num_Axiom) consts - wellfounded :: "('a => 'a => bool) => bool" + wellfounded :: "('a::type => 'a::type => bool) => bool" defs - wellfounded_primdef: "wellfounded == %R. ~ (EX f. ALL n. R (f (Suc n)) (f n))" - -lemma wellfounded_def: "ALL R. wellfounded R = (~ (EX f. ALL n. R (f (Suc n)) (f n)))" + wellfounded_primdef: "wellfounded == +%R::'a::type => 'a::type => bool. + ~ (EX f::nat => 'a::type. ALL n::nat. R (f (Suc n)) (f n))" + +lemma wellfounded_def: "ALL R::'a::type => 'a::type => bool. + wellfounded R = + (~ (EX f::nat => 'a::type. ALL n::nat. R (f (Suc n)) (f n)))" by (import prim_rec wellfounded_def) -lemma WF_IFF_WELLFOUNDED: "ALL R. WF R = wellfounded R" +lemma WF_IFF_WELLFOUNDED: "ALL R::'a::type => 'a::type => bool. WF R = wellfounded R" by (import prim_rec WF_IFF_WELLFOUNDED) -lemma WF_PRED: "WF (%x y. y = Suc x)" +lemma WF_PRED: "WF (%(x::nat) y::nat. y = Suc x)" by (import prim_rec WF_PRED) lemma WF_LESS: "(WF::(nat => nat => bool) => bool) (op <::nat => nat => bool)" by (import prim_rec WF_LESS) consts - measure :: "('a => nat) => 'a => 'a => bool" + measure :: "('a::type => nat) => 'a::type => 'a::type => bool" defs measure_primdef: "prim_rec.measure == relation.inv_image op <" @@ -1013,10 +1289,11 @@ lemma measure_def: "prim_rec.measure = relation.inv_image op <" by (import prim_rec measure_def) -lemma WF_measure: "ALL x. WF (prim_rec.measure x)" +lemma WF_measure: "ALL x::'a::type => nat. WF (prim_rec.measure x)" by (import prim_rec WF_measure) -lemma measure_thm: "ALL x xa xb. prim_rec.measure x xa xb = (x xa < x xb)" +lemma measure_thm: "ALL (x::'a::type => nat) (xa::'a::type) xb::'a::type. + prim_rec.measure x xa xb = (x xa < x xb)" by (import prim_rec measure_thm) ;end_setup @@ -1025,33 +1302,36 @@ constdefs nat_elim__magic :: "nat => nat" - "nat_elim__magic == %n. n" - -lemma nat_elim__magic: "ALL n. nat_elim__magic n = n" + "nat_elim__magic == %n::nat. n" + +lemma nat_elim__magic: "ALL n::nat. nat_elim__magic n = n" by (import arithmetic nat_elim__magic) consts EVEN :: "nat => bool" -specification (EVEN) EVEN: "EVEN 0 = True & (ALL n. EVEN (Suc n) = (~ EVEN n))" +specification (EVEN) EVEN: "EVEN (0::nat) = True & (ALL n::nat. EVEN (Suc n) = (~ EVEN n))" by (import arithmetic EVEN) consts ODD :: "nat => bool" -specification (ODD) ODD: "ODD 0 = False & (ALL n. ODD (Suc n) = (~ ODD n))" +specification (ODD) ODD: "ODD (0::nat) = False & (ALL n::nat. ODD (Suc n) = (~ ODD n))" by (import arithmetic ODD) -lemma TWO: "2 = Suc 1" +lemma TWO: "(2::nat) = Suc (1::nat)" by (import arithmetic TWO) lemma NORM_0: "(0::nat) = (0::nat)" by (import arithmetic NORM_0) -lemma num_case_compute: "ALL n. nat_case f g n = (if n = 0 then f else g (PRE n))" +lemma num_case_compute: "ALL n::nat. + nat_case (f::'a::type) (g::nat => 'a::type) n = + (if n = (0::nat) then f else g (PRE n))" by (import arithmetic num_case_compute) -lemma ADD_CLAUSES: "0 + m = m & m + 0 = m & Suc m + n = Suc (m + n) & m + Suc n = Suc (m + n)" +lemma ADD_CLAUSES: "(0::nat) + (m::nat) = m & +m + (0::nat) = m & Suc m + (n::nat) = Suc (m + n) & m + Suc n = Suc (m + n)" by (import arithmetic ADD_CLAUSES) lemma LESS_ADD: "ALL (m::nat) n::nat. n < m --> (EX p::nat. p + n = m)" @@ -1060,13 +1340,14 @@ lemma LESS_ANTISYM: "ALL (m::nat) n::nat. ~ (m < n & n < m)" by (import arithmetic LESS_ANTISYM) -lemma LESS_LESS_SUC: "ALL x xa. ~ (x < xa & xa < Suc x)" +lemma LESS_LESS_SUC: "ALL (x::nat) xa::nat. ~ (x < xa & xa < Suc x)" by (import arithmetic LESS_LESS_SUC) -lemma FUN_EQ_LEMMA: "ALL f x1 x2. f x1 & ~ f x2 --> x1 ~= x2" +lemma FUN_EQ_LEMMA: "ALL (f::'a::type => bool) (x1::'a::type) x2::'a::type. + f x1 & ~ f x2 --> x1 ~= x2" by (import arithmetic FUN_EQ_LEMMA) -lemma LESS_NOT_SUC: "ALL m n. m < n & n ~= Suc m --> Suc m < n" +lemma LESS_NOT_SUC: "ALL (m::nat) n::nat. m < n & n ~= Suc m --> Suc m < n" by (import arithmetic LESS_NOT_SUC) lemma LESS_0_CASES: "ALL m::nat. (0::nat) = m | (0::nat) < m" @@ -1078,7 +1359,7 @@ lemma LESS_CASES: "ALL (m::nat) n::nat. m < n | n <= m" by (import arithmetic LESS_CASES) -lemma LESS_EQ_SUC_REFL: "ALL m. m <= Suc m" +lemma LESS_EQ_SUC_REFL: "ALL m::nat. m <= Suc m" by (import arithmetic LESS_EQ_SUC_REFL) lemma LESS_ADD_NONZERO: "ALL (m::nat) n::nat. n ~= (0::nat) --> m < m + n" @@ -1090,20 +1371,20 @@ lemma SUB_0: "ALL m::nat. (0::nat) - m = (0::nat) & m - (0::nat) = m" by (import arithmetic SUB_0) -lemma SUC_SUB1: "ALL m. Suc m - 1 = m" +lemma SUC_SUB1: "ALL m::nat. Suc m - (1::nat) = m" by (import arithmetic SUC_SUB1) -lemma PRE_SUB1: "ALL m. PRE m = m - 1" +lemma PRE_SUB1: "ALL m::nat. PRE m = m - (1::nat)" by (import arithmetic PRE_SUB1) -lemma MULT_CLAUSES: "ALL x xa. - 0 * x = 0 & - x * 0 = 0 & - 1 * x = x & - x * 1 = x & Suc x * xa = x * xa + xa & x * Suc xa = x + x * xa" +lemma MULT_CLAUSES: "ALL (x::nat) xa::nat. + (0::nat) * x = (0::nat) & + x * (0::nat) = (0::nat) & + (1::nat) * x = x & + x * (1::nat) = x & Suc x * xa = x * xa + xa & x * Suc xa = x + x * xa" by (import arithmetic MULT_CLAUSES) -lemma PRE_SUB: "ALL m n. PRE (m - n) = PRE m - n" +lemma PRE_SUB: "ALL (m::nat) n::nat. PRE (m - n) = PRE m - n" by (import arithmetic PRE_SUB) lemma ADD_EQ_1: "ALL (m::nat) n::nat. @@ -1114,13 +1395,14 @@ lemma ADD_INV_0_EQ: "ALL (m::nat) n::nat. (m + n = m) = (n = (0::nat))" by (import arithmetic ADD_INV_0_EQ) -lemma PRE_SUC_EQ: "ALL m n. 0 < n --> (m = PRE n) = (Suc m = n)" +lemma PRE_SUC_EQ: "ALL (m::nat) n::nat. (0::nat) < n --> (m = PRE n) = (Suc m = n)" by (import arithmetic PRE_SUC_EQ) -lemma INV_PRE_EQ: "ALL m n. 0 < m & 0 < n --> (PRE m = PRE n) = (m = n)" +lemma INV_PRE_EQ: "ALL (m::nat) n::nat. + (0::nat) < m & (0::nat) < n --> (PRE m = PRE n) = (m = n)" by (import arithmetic INV_PRE_EQ) -lemma LESS_SUC_NOT: "ALL m n. m < n --> ~ n < Suc m" +lemma LESS_SUC_NOT: "ALL (m::nat) n::nat. m < n --> ~ n < Suc m" by (import arithmetic LESS_SUC_NOT) lemma ADD_EQ_SUB: "ALL (m::nat) (n::nat) p::nat. n <= p --> (m + n = p) = (m = p - n)" @@ -1129,16 +1411,17 @@ lemma LESS_ADD_1: "ALL (x::nat) xa::nat. xa < x --> (EX xb::nat. x = xa + (xb + (1::nat)))" by (import arithmetic LESS_ADD_1) -lemma NOT_ODD_EQ_EVEN: "ALL n m. Suc (n + n) ~= m + m" +lemma NOT_ODD_EQ_EVEN: "ALL (n::nat) m::nat. Suc (n + n) ~= m + m" by (import arithmetic NOT_ODD_EQ_EVEN) -lemma MULT_SUC_EQ: "ALL p m n. (n * Suc p = m * Suc p) = (n = m)" +lemma MULT_SUC_EQ: "ALL (p::nat) (m::nat) n::nat. (n * Suc p = m * Suc p) = (n = m)" by (import arithmetic MULT_SUC_EQ) -lemma MULT_EXP_MONO: "ALL p q n m. (n * Suc q ^ p = m * Suc q ^ p) = (n = m)" +lemma MULT_EXP_MONO: "ALL (p::nat) (q::nat) (n::nat) m::nat. + (n * Suc q ^ p = m * Suc q ^ p) = (n = m)" by (import arithmetic MULT_EXP_MONO) -lemma LESS_ADD_SUC: "ALL m n. m < m + Suc n" +lemma LESS_ADD_SUC: "ALL (m::nat) n::nat. m < m + Suc n" by (import arithmetic LESS_ADD_SUC) lemma LESS_OR_EQ_ADD: "ALL (n::nat) m::nat. n < m | (EX p::nat. n = p + m)" @@ -1157,10 +1440,10 @@ ((Not::bool => bool) (P m)))))))" by (import arithmetic WOP) -lemma INV_PRE_LESS: "ALL m>0. ALL n. (PRE m < PRE n) = (m < n)" +lemma INV_PRE_LESS: "ALL m>0::nat. ALL n::nat. (PRE m < PRE n) = (m < n)" by (import arithmetic INV_PRE_LESS) -lemma INV_PRE_LESS_EQ: "ALL n>0. ALL m. (PRE m <= PRE n) = (m <= n)" +lemma INV_PRE_LESS_EQ: "ALL n>0::nat. ALL m::nat. (PRE m <= PRE n) = (m <= n)" by (import arithmetic INV_PRE_LESS_EQ) lemma SUB_EQ_EQ_0: "ALL (m::nat) n::nat. (m - n = m) = (m = (0::nat) | n = (0::nat))" @@ -1175,7 +1458,7 @@ lemma LESS_EQ_SUB_LESS: "ALL (x::nat) xa::nat. xa <= x --> (ALL c::nat. (x - xa < c) = (x < xa + c))" by (import arithmetic LESS_EQ_SUB_LESS) -lemma NOT_SUC_LESS_EQ: "ALL x xa. (~ Suc x <= xa) = (xa <= x)" +lemma NOT_SUC_LESS_EQ: "ALL (x::nat) xa::nat. (~ Suc x <= xa) = (xa <= x)" by (import arithmetic NOT_SUC_LESS_EQ) lemma SUB_LESS_EQ_ADD: "ALL (m::nat) p::nat. m <= p --> (ALL n::nat. (p - m <= n) = (p <= m + n))" @@ -1185,16 +1468,18 @@ xa <= x & xb <= x --> (x - xa = x - xb) = (xa = xb)" by (import arithmetic SUB_CANCEL) -lemma NOT_EXP_0: "ALL m n. Suc n ^ m ~= 0" +lemma NOT_EXP_0: "ALL (m::nat) n::nat. Suc n ^ m ~= (0::nat)" by (import arithmetic NOT_EXP_0) -lemma ZERO_LESS_EXP: "ALL m n. 0 < Suc n ^ m" +lemma ZERO_LESS_EXP: "ALL (m::nat) n::nat. (0::nat) < Suc n ^ m" by (import arithmetic ZERO_LESS_EXP) -lemma ODD_OR_EVEN: "ALL x. EX xa. x = Suc (Suc 0) * xa | x = Suc (Suc 0) * xa + 1" +lemma ODD_OR_EVEN: "ALL x::nat. + EX xa::nat. + x = Suc (Suc (0::nat)) * xa | x = Suc (Suc (0::nat)) * xa + (1::nat)" by (import arithmetic ODD_OR_EVEN) -lemma LESS_EXP_SUC_MONO: "ALL n m. Suc (Suc m) ^ n < Suc (Suc m) ^ Suc n" +lemma LESS_EXP_SUC_MONO: "ALL (n::nat) m::nat. Suc (Suc m) ^ n < Suc (Suc m) ^ Suc n" by (import arithmetic LESS_EXP_SUC_MONO) lemma LESS_LESS_CASES: "ALL (m::nat) n::nat. m = n | m < n | n < m" @@ -1212,67 +1497,69 @@ consts FACT :: "nat => nat" -specification (FACT) FACT: "FACT 0 = 1 & (ALL n. FACT (Suc n) = Suc n * FACT n)" +specification (FACT) FACT: "FACT (0::nat) = (1::nat) & (ALL n::nat. FACT (Suc n) = Suc n * FACT n)" by (import arithmetic FACT) -lemma FACT_LESS: "ALL n. 0 < FACT n" +lemma FACT_LESS: "ALL n::nat. (0::nat) < FACT n" by (import arithmetic FACT_LESS) -lemma EVEN_ODD: "ALL n. EVEN n = (~ ODD n)" +lemma EVEN_ODD: "ALL n::nat. EVEN n = (~ ODD n)" by (import arithmetic EVEN_ODD) -lemma ODD_EVEN: "ALL x. ODD x = (~ EVEN x)" +lemma ODD_EVEN: "ALL x::nat. ODD x = (~ EVEN x)" by (import arithmetic ODD_EVEN) -lemma EVEN_OR_ODD: "ALL x. EVEN x | ODD x" +lemma EVEN_OR_ODD: "ALL x::nat. EVEN x | ODD x" by (import arithmetic EVEN_OR_ODD) -lemma EVEN_AND_ODD: "ALL x. ~ (EVEN x & ODD x)" +lemma EVEN_AND_ODD: "ALL x::nat. ~ (EVEN x & ODD x)" by (import arithmetic EVEN_AND_ODD) -lemma EVEN_ADD: "ALL m n. EVEN (m + n) = (EVEN m = EVEN n)" +lemma EVEN_ADD: "ALL (m::nat) n::nat. EVEN (m + n) = (EVEN m = EVEN n)" by (import arithmetic EVEN_ADD) -lemma EVEN_MULT: "ALL m n. EVEN (m * n) = (EVEN m | EVEN n)" +lemma EVEN_MULT: "ALL (m::nat) n::nat. EVEN (m * n) = (EVEN m | EVEN n)" by (import arithmetic EVEN_MULT) -lemma ODD_ADD: "ALL m n. ODD (m + n) = (ODD m ~= ODD n)" +lemma ODD_ADD: "ALL (m::nat) n::nat. ODD (m + n) = (ODD m ~= ODD n)" by (import arithmetic ODD_ADD) -lemma ODD_MULT: "ALL m n. ODD (m * n) = (ODD m & ODD n)" +lemma ODD_MULT: "ALL (m::nat) n::nat. ODD (m * n) = (ODD m & ODD n)" by (import arithmetic ODD_MULT) -lemma EVEN_DOUBLE: "ALL n. EVEN (2 * n)" +lemma EVEN_DOUBLE: "ALL n::nat. EVEN ((2::nat) * n)" by (import arithmetic EVEN_DOUBLE) -lemma ODD_DOUBLE: "ALL x. ODD (Suc (2 * x))" +lemma ODD_DOUBLE: "ALL x::nat. ODD (Suc ((2::nat) * x))" by (import arithmetic ODD_DOUBLE) -lemma EVEN_ODD_EXISTS: "ALL x. (EVEN x --> (EX m. x = 2 * m)) & (ODD x --> (EX m. x = Suc (2 * m)))" +lemma EVEN_ODD_EXISTS: "ALL x::nat. + (EVEN x --> (EX m::nat. x = (2::nat) * m)) & + (ODD x --> (EX m::nat. x = Suc ((2::nat) * m)))" by (import arithmetic EVEN_ODD_EXISTS) -lemma EVEN_EXISTS: "ALL n. EVEN n = (EX m. n = 2 * m)" +lemma EVEN_EXISTS: "ALL n::nat. EVEN n = (EX m::nat. n = (2::nat) * m)" by (import arithmetic EVEN_EXISTS) -lemma ODD_EXISTS: "ALL n. ODD n = (EX m. n = Suc (2 * m))" +lemma ODD_EXISTS: "ALL n::nat. ODD n = (EX m::nat. n = Suc ((2::nat) * m))" by (import arithmetic ODD_EXISTS) -lemma NOT_SUC_LESS_EQ_0: "ALL x. ~ Suc x <= 0" +lemma NOT_SUC_LESS_EQ_0: "ALL x::nat. ~ Suc x <= (0::nat)" by (import arithmetic NOT_SUC_LESS_EQ_0) -lemma NOT_LEQ: "ALL x xa. (~ x <= xa) = (Suc xa <= x)" +lemma NOT_LEQ: "ALL (x::nat) xa::nat. (~ x <= xa) = (Suc xa <= x)" by (import arithmetic NOT_LEQ) -lemma NOT_NUM_EQ: "ALL x xa. (x ~= xa) = (Suc x <= xa | Suc xa <= x)" +lemma NOT_NUM_EQ: "ALL (x::nat) xa::nat. (x ~= xa) = (Suc x <= xa | Suc xa <= x)" by (import arithmetic NOT_NUM_EQ) -lemma NOT_GREATER_EQ: "ALL x xa. (~ xa <= x) = (Suc x <= xa)" +lemma NOT_GREATER_EQ: "ALL (x::nat) xa::nat. (~ xa <= x) = (Suc x <= xa)" by (import arithmetic NOT_GREATER_EQ) -lemma SUC_ADD_SYM: "ALL m n. Suc (m + n) = Suc n + m" +lemma SUC_ADD_SYM: "ALL (m::nat) n::nat. Suc (m + n) = Suc n + m" by (import arithmetic SUC_ADD_SYM) -lemma NOT_SUC_ADD_LESS_EQ: "ALL m n. ~ Suc (m + n) <= m" +lemma NOT_SUC_ADD_LESS_EQ: "ALL (m::nat) n::nat. ~ Suc (m + n) <= m" by (import arithmetic NOT_SUC_ADD_LESS_EQ) lemma SUB_LEFT_ADD: "ALL (m::nat) (n::nat) p::nat. @@ -1286,7 +1573,8 @@ m - (n - p) = (if n <= p then m else m + p - n)" by (import arithmetic SUB_LEFT_SUB) -lemma SUB_LEFT_SUC: "ALL m n. Suc (m - n) = (if m <= n then Suc 0 else Suc m - n)" +lemma SUB_LEFT_SUC: "ALL (m::nat) n::nat. + Suc (m - n) = (if m <= n then Suc (0::nat) else Suc m - n)" by (import arithmetic SUB_LEFT_SUC) lemma SUB_LEFT_LESS_EQ: "ALL (m::nat) (n::nat) p::nat. (m <= n - p) = (m + p <= n | m <= (0::nat))" @@ -1391,22 +1679,26 @@ (0::nat) < n & (0::nat) < q --> n * (p mod q) = n * p mod (n * q)" by (import arithmetic MOD_COMMON_FACTOR) -lemma num_case_cong: "ALL M M' b f. - M = M' & (M' = 0 --> b = b') & (ALL n. M' = Suc n --> f n = f' n) --> +lemma num_case_cong: "ALL (M::nat) (M'::nat) (b::'a::type) f::nat => 'a::type. + M = M' & + (M' = (0::nat) --> b = (b'::'a::type)) & + (ALL n::nat. M' = Suc n --> f n = (f'::nat => 'a::type) n) --> nat_case b f M = nat_case b' f' M'" by (import arithmetic num_case_cong) -lemma SUC_ELIM_THM: "ALL P. (ALL n. P (Suc n) n) = (ALL n>0. P n (n - 1))" +lemma SUC_ELIM_THM: "ALL P::nat => nat => bool. + (ALL n::nat. P (Suc n) n) = (ALL n>0::nat. P n (n - (1::nat)))" by (import arithmetic SUC_ELIM_THM) lemma SUB_ELIM_THM: "(P::nat => bool) ((a::nat) - (b::nat)) = (ALL x::nat. (b = a + x --> P (0::nat)) & (a = b + x --> P x))" by (import arithmetic SUB_ELIM_THM) -lemma PRE_ELIM_THM: "P (PRE n) = (ALL m. (n = 0 --> P 0) & (n = Suc m --> P m))" +lemma PRE_ELIM_THM: "(P::nat => bool) (PRE (n::nat)) = +(ALL m::nat. (n = (0::nat) --> P (0::nat)) & (n = Suc m --> P m))" by (import arithmetic PRE_ELIM_THM) -lemma MULT_INCREASES: "ALL m n. 1 < m & 0 < n --> Suc n <= m * n" +lemma MULT_INCREASES: "ALL (m::nat) n::nat. (1::nat) < m & (0::nat) < n --> Suc n <= m * n" by (import arithmetic MULT_INCREASES) lemma EXP_ALWAYS_BIG_ENOUGH: "ALL b>1::nat. ALL n::nat. EX m::nat. n <= b ^ m" @@ -1466,133 +1758,146 @@ constdefs trat_1 :: "nat * nat" - "trat_1 == (0, 0)" - -lemma trat_1: "trat_1 = (0, 0)" + "trat_1 == (0::nat, 0::nat)" + +lemma trat_1: "trat_1 = (0::nat, 0::nat)" by (import hrat trat_1) constdefs trat_inv :: "nat * nat => nat * nat" - "trat_inv == %(x, y). (y, x)" - -lemma trat_inv: "ALL x y. trat_inv (x, y) = (y, x)" + "trat_inv == %(x::nat, y::nat). (y, x)" + +lemma trat_inv: "ALL (x::nat) y::nat. trat_inv (x, y) = (y, x)" by (import hrat trat_inv) constdefs trat_add :: "nat * nat => nat * nat => nat * nat" "trat_add == -%(x, y) (x', y'). +%(x::nat, y::nat) (x'::nat, y'::nat). (PRE (Suc x * Suc y' + Suc x' * Suc y), PRE (Suc y * Suc y'))" -lemma trat_add: "ALL x y x' y'. +lemma trat_add: "ALL (x::nat) (y::nat) (x'::nat) y'::nat. trat_add (x, y) (x', y') = (PRE (Suc x * Suc y' + Suc x' * Suc y), PRE (Suc y * Suc y'))" by (import hrat trat_add) constdefs trat_mul :: "nat * nat => nat * nat => nat * nat" - "trat_mul == %(x, y) (x', y'). (PRE (Suc x * Suc x'), PRE (Suc y * Suc y'))" - -lemma trat_mul: "ALL x y x' y'. + "trat_mul == +%(x::nat, y::nat) (x'::nat, y'::nat). + (PRE (Suc x * Suc x'), PRE (Suc y * Suc y'))" + +lemma trat_mul: "ALL (x::nat) (y::nat) (x'::nat) y'::nat. trat_mul (x, y) (x', y') = (PRE (Suc x * Suc x'), PRE (Suc y * Suc y'))" by (import hrat trat_mul) consts trat_sucint :: "nat => nat * nat" -specification (trat_sucint) trat_sucint: "trat_sucint 0 = trat_1 & -(ALL n. trat_sucint (Suc n) = trat_add (trat_sucint n) trat_1)" +specification (trat_sucint) trat_sucint: "trat_sucint (0::nat) = trat_1 & +(ALL n::nat. trat_sucint (Suc n) = trat_add (trat_sucint n) trat_1)" by (import hrat trat_sucint) constdefs trat_eq :: "nat * nat => nat * nat => bool" - "trat_eq == %(x, y) (x', y'). Suc x * Suc y' = Suc x' * Suc y" - -lemma trat_eq: "ALL x y x' y'. trat_eq (x, y) (x', y') = (Suc x * Suc y' = Suc x' * Suc y)" + "trat_eq == +%(x::nat, y::nat) (x'::nat, y'::nat). Suc x * Suc y' = Suc x' * Suc y" + +lemma trat_eq: "ALL (x::nat) (y::nat) (x'::nat) y'::nat. + trat_eq (x, y) (x', y') = (Suc x * Suc y' = Suc x' * Suc y)" by (import hrat trat_eq) -lemma TRAT_EQ_REFL: "ALL p. trat_eq p p" +lemma TRAT_EQ_REFL: "ALL p::nat * nat. trat_eq p p" by (import hrat TRAT_EQ_REFL) -lemma TRAT_EQ_SYM: "ALL p q. trat_eq p q = trat_eq q p" +lemma TRAT_EQ_SYM: "ALL (p::nat * nat) q::nat * nat. trat_eq p q = trat_eq q p" by (import hrat TRAT_EQ_SYM) -lemma TRAT_EQ_TRANS: "ALL p q r. trat_eq p q & trat_eq q r --> trat_eq p r" +lemma TRAT_EQ_TRANS: "ALL (p::nat * nat) (q::nat * nat) r::nat * nat. + trat_eq p q & trat_eq q r --> trat_eq p r" by (import hrat TRAT_EQ_TRANS) -lemma TRAT_EQ_AP: "ALL p q. p = q --> trat_eq p q" +lemma TRAT_EQ_AP: "ALL (p::nat * nat) q::nat * nat. p = q --> trat_eq p q" by (import hrat TRAT_EQ_AP) -lemma TRAT_ADD_SYM_EQ: "ALL h i. trat_add h i = trat_add i h" +lemma TRAT_ADD_SYM_EQ: "ALL (h::nat * nat) i::nat * nat. trat_add h i = trat_add i h" by (import hrat TRAT_ADD_SYM_EQ) -lemma TRAT_MUL_SYM_EQ: "ALL h i. trat_mul h i = trat_mul i h" +lemma TRAT_MUL_SYM_EQ: "ALL (h::nat * nat) i::nat * nat. trat_mul h i = trat_mul i h" by (import hrat TRAT_MUL_SYM_EQ) -lemma TRAT_INV_WELLDEFINED: "ALL p q. trat_eq p q --> trat_eq (trat_inv p) (trat_inv q)" +lemma TRAT_INV_WELLDEFINED: "ALL (p::nat * nat) q::nat * nat. + trat_eq p q --> trat_eq (trat_inv p) (trat_inv q)" by (import hrat TRAT_INV_WELLDEFINED) -lemma TRAT_ADD_WELLDEFINED: "ALL p q r. trat_eq p q --> trat_eq (trat_add p r) (trat_add q r)" +lemma TRAT_ADD_WELLDEFINED: "ALL (p::nat * nat) (q::nat * nat) r::nat * nat. + trat_eq p q --> trat_eq (trat_add p r) (trat_add q r)" by (import hrat TRAT_ADD_WELLDEFINED) -lemma TRAT_ADD_WELLDEFINED2: "ALL p1 p2 q1 q2. +lemma TRAT_ADD_WELLDEFINED2: "ALL (p1::nat * nat) (p2::nat * nat) (q1::nat * nat) q2::nat * nat. trat_eq p1 p2 & trat_eq q1 q2 --> trat_eq (trat_add p1 q1) (trat_add p2 q2)" by (import hrat TRAT_ADD_WELLDEFINED2) -lemma TRAT_MUL_WELLDEFINED: "ALL p q r. trat_eq p q --> trat_eq (trat_mul p r) (trat_mul q r)" +lemma TRAT_MUL_WELLDEFINED: "ALL (p::nat * nat) (q::nat * nat) r::nat * nat. + trat_eq p q --> trat_eq (trat_mul p r) (trat_mul q r)" by (import hrat TRAT_MUL_WELLDEFINED) -lemma TRAT_MUL_WELLDEFINED2: "ALL p1 p2 q1 q2. +lemma TRAT_MUL_WELLDEFINED2: "ALL (p1::nat * nat) (p2::nat * nat) (q1::nat * nat) q2::nat * nat. trat_eq p1 p2 & trat_eq q1 q2 --> trat_eq (trat_mul p1 q1) (trat_mul p2 q2)" by (import hrat TRAT_MUL_WELLDEFINED2) -lemma TRAT_ADD_SYM: "ALL h i. trat_eq (trat_add h i) (trat_add i h)" +lemma TRAT_ADD_SYM: "ALL (h::nat * nat) i::nat * nat. trat_eq (trat_add h i) (trat_add i h)" by (import hrat TRAT_ADD_SYM) -lemma TRAT_ADD_ASSOC: "ALL h i j. trat_eq (trat_add h (trat_add i j)) (trat_add (trat_add h i) j)" +lemma TRAT_ADD_ASSOC: "ALL (h::nat * nat) (i::nat * nat) j::nat * nat. + trat_eq (trat_add h (trat_add i j)) (trat_add (trat_add h i) j)" by (import hrat TRAT_ADD_ASSOC) -lemma TRAT_MUL_SYM: "ALL h i. trat_eq (trat_mul h i) (trat_mul i h)" +lemma TRAT_MUL_SYM: "ALL (h::nat * nat) i::nat * nat. trat_eq (trat_mul h i) (trat_mul i h)" by (import hrat TRAT_MUL_SYM) -lemma TRAT_MUL_ASSOC: "ALL h i j. trat_eq (trat_mul h (trat_mul i j)) (trat_mul (trat_mul h i) j)" +lemma TRAT_MUL_ASSOC: "ALL (h::nat * nat) (i::nat * nat) j::nat * nat. + trat_eq (trat_mul h (trat_mul i j)) (trat_mul (trat_mul h i) j)" by (import hrat TRAT_MUL_ASSOC) -lemma TRAT_LDISTRIB: "ALL h i j. +lemma TRAT_LDISTRIB: "ALL (h::nat * nat) (i::nat * nat) j::nat * nat. trat_eq (trat_mul h (trat_add i j)) (trat_add (trat_mul h i) (trat_mul h j))" by (import hrat TRAT_LDISTRIB) -lemma TRAT_MUL_LID: "ALL h. trat_eq (trat_mul trat_1 h) h" +lemma TRAT_MUL_LID: "ALL h::nat * nat. trat_eq (trat_mul trat_1 h) h" by (import hrat TRAT_MUL_LID) -lemma TRAT_MUL_LINV: "ALL h. trat_eq (trat_mul (trat_inv h) h) trat_1" +lemma TRAT_MUL_LINV: "ALL h::nat * nat. trat_eq (trat_mul (trat_inv h) h) trat_1" by (import hrat TRAT_MUL_LINV) -lemma TRAT_NOZERO: "ALL h i. ~ trat_eq (trat_add h i) h" +lemma TRAT_NOZERO: "ALL (h::nat * nat) i::nat * nat. ~ trat_eq (trat_add h i) h" by (import hrat TRAT_NOZERO) -lemma TRAT_ADD_TOTAL: "ALL h i. +lemma TRAT_ADD_TOTAL: "ALL (h::nat * nat) i::nat * nat. trat_eq h i | - (EX d. trat_eq h (trat_add i d)) | (EX d. trat_eq i (trat_add h d))" + (EX d::nat * nat. trat_eq h (trat_add i d)) | + (EX d::nat * nat. trat_eq i (trat_add h d))" by (import hrat TRAT_ADD_TOTAL) -lemma TRAT_SUCINT_0: "ALL n. trat_eq (trat_sucint n) (n, 0)" +lemma TRAT_SUCINT_0: "ALL n::nat. trat_eq (trat_sucint n) (n, 0::nat)" by (import hrat TRAT_SUCINT_0) -lemma TRAT_ARCH: "ALL h. EX n d. trat_eq (trat_sucint n) (trat_add h d)" +lemma TRAT_ARCH: "ALL h::nat * nat. + EX (n::nat) d::nat * nat. trat_eq (trat_sucint n) (trat_add h d)" by (import hrat TRAT_ARCH) -lemma TRAT_SUCINT: "trat_eq (trat_sucint 0) trat_1 & -(ALL n. trat_eq (trat_sucint (Suc n)) (trat_add (trat_sucint n) trat_1))" +lemma TRAT_SUCINT: "trat_eq (trat_sucint (0::nat)) trat_1 & +(ALL n::nat. + trat_eq (trat_sucint (Suc n)) (trat_add (trat_sucint n) trat_1))" by (import hrat TRAT_SUCINT) -lemma TRAT_EQ_EQUIV: "ALL p q. trat_eq p q = (trat_eq p = trat_eq q)" +lemma TRAT_EQ_EQUIV: "ALL (p::nat * nat) q::nat * nat. trat_eq p q = (trat_eq p = trat_eq q)" by (import hrat TRAT_EQ_EQUIV) -typedef (open) hrat = "{x. EX xa. x = trat_eq xa}" +typedef (open) hrat = "{x::nat * nat => bool. EX xa::nat * nat. x = trat_eq xa}" by (rule typedef_helper,import hrat hrat_TY_DEF) lemmas hrat_TY_DEF = typedef_hol2hol4 [OF type_definition_hrat] @@ -1601,8 +1906,9 @@ mk_hrat :: "(nat * nat => bool) => hrat" dest_hrat :: "hrat => nat * nat => bool" -specification (dest_hrat mk_hrat) hrat_tybij: "(ALL a. mk_hrat (dest_hrat a) = a) & -(ALL r. (EX x. r = trat_eq x) = (dest_hrat (mk_hrat r) = r))" +specification (dest_hrat mk_hrat) hrat_tybij: "(ALL a::hrat. mk_hrat (dest_hrat a) = a) & +(ALL r::nat * nat => bool. + (EX x::nat * nat. r = trat_eq x) = (dest_hrat (mk_hrat r) = r))" by (import hrat hrat_tybij) constdefs @@ -1614,18 +1920,19 @@ constdefs hrat_inv :: "hrat => hrat" - "hrat_inv == %T1. mk_hrat (trat_eq (trat_inv (Eps (dest_hrat T1))))" - -lemma hrat_inv: "ALL T1. hrat_inv T1 = mk_hrat (trat_eq (trat_inv (Eps (dest_hrat T1))))" + "hrat_inv == %T1::hrat. mk_hrat (trat_eq (trat_inv (Eps (dest_hrat T1))))" + +lemma hrat_inv: "ALL T1::hrat. + hrat_inv T1 = mk_hrat (trat_eq (trat_inv (Eps (dest_hrat T1))))" by (import hrat hrat_inv) constdefs hrat_add :: "hrat => hrat => hrat" "hrat_add == -%T1 T2. +%(T1::hrat) T2::hrat. mk_hrat (trat_eq (trat_add (Eps (dest_hrat T1)) (Eps (dest_hrat T2))))" -lemma hrat_add: "ALL T1 T2. +lemma hrat_add: "ALL (T1::hrat) T2::hrat. hrat_add T1 T2 = mk_hrat (trat_eq (trat_add (Eps (dest_hrat T1)) (Eps (dest_hrat T2))))" by (import hrat hrat_add) @@ -1633,54 +1940,57 @@ constdefs hrat_mul :: "hrat => hrat => hrat" "hrat_mul == -%T1 T2. +%(T1::hrat) T2::hrat. mk_hrat (trat_eq (trat_mul (Eps (dest_hrat T1)) (Eps (dest_hrat T2))))" -lemma hrat_mul: "ALL T1 T2. +lemma hrat_mul: "ALL (T1::hrat) T2::hrat. hrat_mul T1 T2 = mk_hrat (trat_eq (trat_mul (Eps (dest_hrat T1)) (Eps (dest_hrat T2))))" by (import hrat hrat_mul) constdefs hrat_sucint :: "nat => hrat" - "hrat_sucint == %T1. mk_hrat (trat_eq (trat_sucint T1))" - -lemma hrat_sucint: "ALL T1. hrat_sucint T1 = mk_hrat (trat_eq (trat_sucint T1))" + "hrat_sucint == %T1::nat. mk_hrat (trat_eq (trat_sucint T1))" + +lemma hrat_sucint: "ALL T1::nat. hrat_sucint T1 = mk_hrat (trat_eq (trat_sucint T1))" by (import hrat hrat_sucint) -lemma HRAT_ADD_SYM: "ALL h i. hrat_add h i = hrat_add i h" +lemma HRAT_ADD_SYM: "ALL (h::hrat) i::hrat. hrat_add h i = hrat_add i h" by (import hrat HRAT_ADD_SYM) -lemma HRAT_ADD_ASSOC: "ALL h i j. hrat_add h (hrat_add i j) = hrat_add (hrat_add h i) j" +lemma HRAT_ADD_ASSOC: "ALL (h::hrat) (i::hrat) j::hrat. + hrat_add h (hrat_add i j) = hrat_add (hrat_add h i) j" by (import hrat HRAT_ADD_ASSOC) -lemma HRAT_MUL_SYM: "ALL h i. hrat_mul h i = hrat_mul i h" +lemma HRAT_MUL_SYM: "ALL (h::hrat) i::hrat. hrat_mul h i = hrat_mul i h" by (import hrat HRAT_MUL_SYM) -lemma HRAT_MUL_ASSOC: "ALL h i j. hrat_mul h (hrat_mul i j) = hrat_mul (hrat_mul h i) j" +lemma HRAT_MUL_ASSOC: "ALL (h::hrat) (i::hrat) j::hrat. + hrat_mul h (hrat_mul i j) = hrat_mul (hrat_mul h i) j" by (import hrat HRAT_MUL_ASSOC) -lemma HRAT_LDISTRIB: "ALL h i j. +lemma HRAT_LDISTRIB: "ALL (h::hrat) (i::hrat) j::hrat. hrat_mul h (hrat_add i j) = hrat_add (hrat_mul h i) (hrat_mul h j)" by (import hrat HRAT_LDISTRIB) -lemma HRAT_MUL_LID: "ALL h. hrat_mul hrat_1 h = h" +lemma HRAT_MUL_LID: "ALL h::hrat. hrat_mul hrat_1 h = h" by (import hrat HRAT_MUL_LID) -lemma HRAT_MUL_LINV: "ALL h. hrat_mul (hrat_inv h) h = hrat_1" +lemma HRAT_MUL_LINV: "ALL h::hrat. hrat_mul (hrat_inv h) h = hrat_1" by (import hrat HRAT_MUL_LINV) -lemma HRAT_NOZERO: "ALL h i. hrat_add h i ~= h" +lemma HRAT_NOZERO: "ALL (h::hrat) i::hrat. hrat_add h i ~= h" by (import hrat HRAT_NOZERO) -lemma HRAT_ADD_TOTAL: "ALL h i. h = i | (EX x. h = hrat_add i x) | (EX x. i = hrat_add h x)" +lemma HRAT_ADD_TOTAL: "ALL (h::hrat) i::hrat. + h = i | (EX x::hrat. h = hrat_add i x) | (EX x::hrat. i = hrat_add h x)" by (import hrat HRAT_ADD_TOTAL) -lemma HRAT_ARCH: "ALL h. EX x xa. hrat_sucint x = hrat_add h xa" +lemma HRAT_ARCH: "ALL h::hrat. EX (x::nat) xa::hrat. hrat_sucint x = hrat_add h xa" by (import hrat HRAT_ARCH) -lemma HRAT_SUCINT: "hrat_sucint 0 = hrat_1 & -(ALL x. hrat_sucint (Suc x) = hrat_add (hrat_sucint x) hrat_1)" +lemma HRAT_SUCINT: "hrat_sucint (0::nat) = hrat_1 & +(ALL x::nat. hrat_sucint (Suc x) = hrat_add (hrat_sucint x) hrat_1)" by (import hrat HRAT_SUCINT) ;end_setup @@ -1689,128 +1999,138 @@ constdefs hrat_lt :: "hrat => hrat => bool" - "hrat_lt == %x y. EX d. y = hrat_add x d" - -lemma hrat_lt: "ALL x y. hrat_lt x y = (EX d. y = hrat_add x d)" + "hrat_lt == %(x::hrat) y::hrat. EX d::hrat. y = hrat_add x d" + +lemma hrat_lt: "ALL (x::hrat) y::hrat. hrat_lt x y = (EX d::hrat. y = hrat_add x d)" by (import hreal hrat_lt) -lemma HRAT_LT_REFL: "ALL x. ~ hrat_lt x x" +lemma HRAT_LT_REFL: "ALL x::hrat. ~ hrat_lt x x" by (import hreal HRAT_LT_REFL) -lemma HRAT_LT_TRANS: "ALL x y z. hrat_lt x y & hrat_lt y z --> hrat_lt x z" +lemma HRAT_LT_TRANS: "ALL (x::hrat) (y::hrat) z::hrat. hrat_lt x y & hrat_lt y z --> hrat_lt x z" by (import hreal HRAT_LT_TRANS) -lemma HRAT_LT_ANTISYM: "ALL x y. ~ (hrat_lt x y & hrat_lt y x)" +lemma HRAT_LT_ANTISYM: "ALL (x::hrat) y::hrat. ~ (hrat_lt x y & hrat_lt y x)" by (import hreal HRAT_LT_ANTISYM) -lemma HRAT_LT_TOTAL: "ALL x y. x = y | hrat_lt x y | hrat_lt y x" +lemma HRAT_LT_TOTAL: "ALL (x::hrat) y::hrat. x = y | hrat_lt x y | hrat_lt y x" by (import hreal HRAT_LT_TOTAL) -lemma HRAT_MUL_RID: "ALL x. hrat_mul x hrat_1 = x" +lemma HRAT_MUL_RID: "ALL x::hrat. hrat_mul x hrat_1 = x" by (import hreal HRAT_MUL_RID) -lemma HRAT_MUL_RINV: "ALL x. hrat_mul x (hrat_inv x) = hrat_1" +lemma HRAT_MUL_RINV: "ALL x::hrat. hrat_mul x (hrat_inv x) = hrat_1" by (import hreal HRAT_MUL_RINV) -lemma HRAT_RDISTRIB: "ALL x y z. +lemma HRAT_RDISTRIB: "ALL (x::hrat) (y::hrat) z::hrat. hrat_mul (hrat_add x y) z = hrat_add (hrat_mul x z) (hrat_mul y z)" by (import hreal HRAT_RDISTRIB) -lemma HRAT_LT_ADDL: "ALL x y. hrat_lt x (hrat_add x y)" +lemma HRAT_LT_ADDL: "ALL (x::hrat) y::hrat. hrat_lt x (hrat_add x y)" by (import hreal HRAT_LT_ADDL) -lemma HRAT_LT_ADDR: "ALL x xa. hrat_lt xa (hrat_add x xa)" +lemma HRAT_LT_ADDR: "ALL (x::hrat) xa::hrat. hrat_lt xa (hrat_add x xa)" by (import hreal HRAT_LT_ADDR) -lemma HRAT_LT_GT: "ALL x y. hrat_lt x y --> ~ hrat_lt y x" +lemma HRAT_LT_GT: "ALL (x::hrat) y::hrat. hrat_lt x y --> ~ hrat_lt y x" by (import hreal HRAT_LT_GT) -lemma HRAT_LT_NE: "ALL x y. hrat_lt x y --> x ~= y" +lemma HRAT_LT_NE: "ALL (x::hrat) y::hrat. hrat_lt x y --> x ~= y" by (import hreal HRAT_LT_NE) -lemma HRAT_EQ_LADD: "ALL x y z. (hrat_add x y = hrat_add x z) = (y = z)" +lemma HRAT_EQ_LADD: "ALL (x::hrat) (y::hrat) z::hrat. (hrat_add x y = hrat_add x z) = (y = z)" by (import hreal HRAT_EQ_LADD) -lemma HRAT_EQ_LMUL: "ALL x y z. (hrat_mul x y = hrat_mul x z) = (y = z)" +lemma HRAT_EQ_LMUL: "ALL (x::hrat) (y::hrat) z::hrat. (hrat_mul x y = hrat_mul x z) = (y = z)" by (import hreal HRAT_EQ_LMUL) -lemma HRAT_LT_ADD2: "ALL u v x y. +lemma HRAT_LT_ADD2: "ALL (u::hrat) (v::hrat) (x::hrat) y::hrat. hrat_lt u x & hrat_lt v y --> hrat_lt (hrat_add u v) (hrat_add x y)" by (import hreal HRAT_LT_ADD2) -lemma HRAT_LT_LADD: "ALL x y z. hrat_lt (hrat_add z x) (hrat_add z y) = hrat_lt x y" +lemma HRAT_LT_LADD: "ALL (x::hrat) (y::hrat) z::hrat. + hrat_lt (hrat_add z x) (hrat_add z y) = hrat_lt x y" by (import hreal HRAT_LT_LADD) -lemma HRAT_LT_RADD: "ALL x y z. hrat_lt (hrat_add x z) (hrat_add y z) = hrat_lt x y" +lemma HRAT_LT_RADD: "ALL (x::hrat) (y::hrat) z::hrat. + hrat_lt (hrat_add x z) (hrat_add y z) = hrat_lt x y" by (import hreal HRAT_LT_RADD) -lemma HRAT_LT_MUL2: "ALL u v x y. +lemma HRAT_LT_MUL2: "ALL (u::hrat) (v::hrat) (x::hrat) y::hrat. hrat_lt u x & hrat_lt v y --> hrat_lt (hrat_mul u v) (hrat_mul x y)" by (import hreal HRAT_LT_MUL2) -lemma HRAT_LT_LMUL: "ALL x y z. hrat_lt (hrat_mul z x) (hrat_mul z y) = hrat_lt x y" +lemma HRAT_LT_LMUL: "ALL (x::hrat) (y::hrat) z::hrat. + hrat_lt (hrat_mul z x) (hrat_mul z y) = hrat_lt x y" by (import hreal HRAT_LT_LMUL) -lemma HRAT_LT_RMUL: "ALL x y z. hrat_lt (hrat_mul x z) (hrat_mul y z) = hrat_lt x y" +lemma HRAT_LT_RMUL: "ALL (x::hrat) (y::hrat) z::hrat. + hrat_lt (hrat_mul x z) (hrat_mul y z) = hrat_lt x y" by (import hreal HRAT_LT_RMUL) -lemma HRAT_LT_LMUL1: "ALL x y. hrat_lt (hrat_mul x y) y = hrat_lt x hrat_1" +lemma HRAT_LT_LMUL1: "ALL (x::hrat) y::hrat. hrat_lt (hrat_mul x y) y = hrat_lt x hrat_1" by (import hreal HRAT_LT_LMUL1) -lemma HRAT_LT_RMUL1: "ALL x y. hrat_lt (hrat_mul x y) x = hrat_lt y hrat_1" +lemma HRAT_LT_RMUL1: "ALL (x::hrat) y::hrat. hrat_lt (hrat_mul x y) x = hrat_lt y hrat_1" by (import hreal HRAT_LT_RMUL1) -lemma HRAT_GT_LMUL1: "ALL x y. hrat_lt y (hrat_mul x y) = hrat_lt hrat_1 x" +lemma HRAT_GT_LMUL1: "ALL (x::hrat) y::hrat. hrat_lt y (hrat_mul x y) = hrat_lt hrat_1 x" by (import hreal HRAT_GT_LMUL1) -lemma HRAT_LT_L1: "ALL x y. hrat_lt (hrat_mul (hrat_inv x) y) hrat_1 = hrat_lt y x" +lemma HRAT_LT_L1: "ALL (x::hrat) y::hrat. + hrat_lt (hrat_mul (hrat_inv x) y) hrat_1 = hrat_lt y x" by (import hreal HRAT_LT_L1) -lemma HRAT_LT_R1: "ALL x y. hrat_lt (hrat_mul x (hrat_inv y)) hrat_1 = hrat_lt x y" +lemma HRAT_LT_R1: "ALL (x::hrat) y::hrat. + hrat_lt (hrat_mul x (hrat_inv y)) hrat_1 = hrat_lt x y" by (import hreal HRAT_LT_R1) -lemma HRAT_GT_L1: "ALL x y. hrat_lt hrat_1 (hrat_mul (hrat_inv x) y) = hrat_lt x y" +lemma HRAT_GT_L1: "ALL (x::hrat) y::hrat. + hrat_lt hrat_1 (hrat_mul (hrat_inv x) y) = hrat_lt x y" by (import hreal HRAT_GT_L1) -lemma HRAT_INV_MUL: "ALL x y. hrat_inv (hrat_mul x y) = hrat_mul (hrat_inv x) (hrat_inv y)" +lemma HRAT_INV_MUL: "ALL (x::hrat) y::hrat. + hrat_inv (hrat_mul x y) = hrat_mul (hrat_inv x) (hrat_inv y)" by (import hreal HRAT_INV_MUL) -lemma HRAT_UP: "ALL x. Ex (hrat_lt x)" +lemma HRAT_UP: "ALL x::hrat. Ex (hrat_lt x)" by (import hreal HRAT_UP) -lemma HRAT_DOWN: "ALL x. EX xa. hrat_lt xa x" +lemma HRAT_DOWN: "ALL x::hrat. EX xa::hrat. hrat_lt xa x" by (import hreal HRAT_DOWN) -lemma HRAT_DOWN2: "ALL x y. EX xa. hrat_lt xa x & hrat_lt xa y" +lemma HRAT_DOWN2: "ALL (x::hrat) y::hrat. EX xa::hrat. hrat_lt xa x & hrat_lt xa y" by (import hreal HRAT_DOWN2) -lemma HRAT_MEAN: "ALL x y. hrat_lt x y --> (EX xa. hrat_lt x xa & hrat_lt xa y)" +lemma HRAT_MEAN: "ALL (x::hrat) y::hrat. + hrat_lt x y --> (EX xa::hrat. hrat_lt x xa & hrat_lt xa y)" by (import hreal HRAT_MEAN) constdefs isacut :: "(hrat => bool) => bool" "isacut == -%C. Ex C & - (EX x. ~ C x) & - (ALL x y. C x & hrat_lt y x --> C y) & - (ALL x. C x --> (EX y. C y & hrat_lt x y))" - -lemma isacut: "ALL C. +%C::hrat => bool. + Ex C & + (EX x::hrat. ~ C x) & + (ALL (x::hrat) y::hrat. C x & hrat_lt y x --> C y) & + (ALL x::hrat. C x --> (EX y::hrat. C y & hrat_lt x y))" + +lemma isacut: "ALL C::hrat => bool. isacut C = (Ex C & - (EX x. ~ C x) & - (ALL x y. C x & hrat_lt y x --> C y) & - (ALL x. C x --> (EX y. C y & hrat_lt x y)))" + (EX x::hrat. ~ C x) & + (ALL (x::hrat) y::hrat. C x & hrat_lt y x --> C y) & + (ALL x::hrat. C x --> (EX y::hrat. C y & hrat_lt x y)))" by (import hreal isacut) constdefs cut_of_hrat :: "hrat => hrat => bool" - "cut_of_hrat == %x y. hrat_lt y x" - -lemma cut_of_hrat: "ALL x. cut_of_hrat x = (%y. hrat_lt y x)" + "cut_of_hrat == %(x::hrat) y::hrat. hrat_lt y x" + +lemma cut_of_hrat: "ALL x::hrat. cut_of_hrat x = (%y::hrat. hrat_lt y x)" by (import hreal cut_of_hrat) -lemma ISACUT_HRAT: "ALL h. isacut (cut_of_hrat h)" +lemma ISACUT_HRAT: "ALL h::hrat. isacut (cut_of_hrat h)" by (import hreal ISACUT_HRAT) typedef (open) hreal = "Collect isacut" @@ -1822,39 +2142,45 @@ hreal :: "(hrat => bool) => hreal" cut :: "hreal => hrat => bool" -specification (cut hreal) hreal_tybij: "(ALL a. hreal (hreal.cut a) = a) & -(ALL r. isacut r = (hreal.cut (hreal r) = r))" +specification (cut hreal) hreal_tybij: "(ALL a::hreal. hreal (hreal.cut a) = a) & +(ALL r::hrat => bool. isacut r = (hreal.cut (hreal r) = r))" by (import hreal hreal_tybij) -lemma EQUAL_CUTS: "ALL X Y. hreal.cut X = hreal.cut Y --> X = Y" +lemma EQUAL_CUTS: "ALL (X::hreal) Y::hreal. hreal.cut X = hreal.cut Y --> X = Y" by (import hreal EQUAL_CUTS) -lemma CUT_ISACUT: "ALL x. isacut (hreal.cut x)" +lemma CUT_ISACUT: "ALL x::hreal. isacut (hreal.cut x)" by (import hreal CUT_ISACUT) -lemma CUT_NONEMPTY: "ALL x. Ex (hreal.cut x)" +lemma CUT_NONEMPTY: "ALL x::hreal. Ex (hreal.cut x)" by (import hreal CUT_NONEMPTY) -lemma CUT_BOUNDED: "ALL x. EX xa. ~ hreal.cut x xa" +lemma CUT_BOUNDED: "ALL x::hreal. EX xa::hrat. ~ hreal.cut x xa" by (import hreal CUT_BOUNDED) -lemma CUT_DOWN: "ALL x xa xb. hreal.cut x xa & hrat_lt xb xa --> hreal.cut x xb" +lemma CUT_DOWN: "ALL (x::hreal) (xa::hrat) xb::hrat. + hreal.cut x xa & hrat_lt xb xa --> hreal.cut x xb" by (import hreal CUT_DOWN) -lemma CUT_UP: "ALL x xa. hreal.cut x xa --> (EX y. hreal.cut x y & hrat_lt xa y)" +lemma CUT_UP: "ALL (x::hreal) xa::hrat. + hreal.cut x xa --> (EX y::hrat. hreal.cut x y & hrat_lt xa y)" by (import hreal CUT_UP) -lemma CUT_UBOUND: "ALL x xa xb. ~ hreal.cut x xa & hrat_lt xa xb --> ~ hreal.cut x xb" +lemma CUT_UBOUND: "ALL (x::hreal) (xa::hrat) xb::hrat. + ~ hreal.cut x xa & hrat_lt xa xb --> ~ hreal.cut x xb" by (import hreal CUT_UBOUND) -lemma CUT_STRADDLE: "ALL X x y. hreal.cut X x & ~ hreal.cut X y --> hrat_lt x y" +lemma CUT_STRADDLE: "ALL (X::hreal) (x::hrat) y::hrat. + hreal.cut X x & ~ hreal.cut X y --> hrat_lt x y" by (import hreal CUT_STRADDLE) -lemma CUT_NEARTOP_ADD: "ALL X e. EX x. hreal.cut X x & ~ hreal.cut X (hrat_add x e)" +lemma CUT_NEARTOP_ADD: "ALL (X::hreal) e::hrat. + EX x::hrat. hreal.cut X x & ~ hreal.cut X (hrat_add x e)" by (import hreal CUT_NEARTOP_ADD) -lemma CUT_NEARTOP_MUL: "ALL X u. - hrat_lt hrat_1 u --> (EX x. hreal.cut X x & ~ hreal.cut X (hrat_mul u x))" +lemma CUT_NEARTOP_MUL: "ALL (X::hreal) u::hrat. + hrat_lt hrat_1 u --> + (EX x::hrat. hreal.cut X x & ~ hreal.cut X (hrat_mul u x))" by (import hreal CUT_NEARTOP_MUL) constdefs @@ -1867,128 +2193,170 @@ constdefs hreal_add :: "hreal => hreal => hreal" "hreal_add == -%X Y. hreal (%w. EX x y. w = hrat_add x y & hreal.cut X x & hreal.cut Y y)" - -lemma hreal_add: "ALL X Y. +%(X::hreal) Y::hreal. + hreal + (%w::hrat. + EX (x::hrat) y::hrat. + w = hrat_add x y & hreal.cut X x & hreal.cut Y y)" + +lemma hreal_add: "ALL (X::hreal) Y::hreal. hreal_add X Y = - hreal (%w. EX x y. w = hrat_add x y & hreal.cut X x & hreal.cut Y y)" + hreal + (%w::hrat. + EX (x::hrat) y::hrat. + w = hrat_add x y & hreal.cut X x & hreal.cut Y y)" by (import hreal hreal_add) constdefs hreal_mul :: "hreal => hreal => hreal" "hreal_mul == -%X Y. hreal (%w. EX x y. w = hrat_mul x y & hreal.cut X x & hreal.cut Y y)" - -lemma hreal_mul: "ALL X Y. +%(X::hreal) Y::hreal. + hreal + (%w::hrat. + EX (x::hrat) y::hrat. + w = hrat_mul x y & hreal.cut X x & hreal.cut Y y)" + +lemma hreal_mul: "ALL (X::hreal) Y::hreal. hreal_mul X Y = - hreal (%w. EX x y. w = hrat_mul x y & hreal.cut X x & hreal.cut Y y)" + hreal + (%w::hrat. + EX (x::hrat) y::hrat. + w = hrat_mul x y & hreal.cut X x & hreal.cut Y y)" by (import hreal hreal_mul) constdefs hreal_inv :: "hreal => hreal" "hreal_inv == -%X. hreal - (%w. EX d. hrat_lt d hrat_1 & - (ALL x. hreal.cut X x --> hrat_lt (hrat_mul w x) d))" - -lemma hreal_inv: "ALL X. +%X::hreal. + hreal + (%w::hrat. + EX d::hrat. + hrat_lt d hrat_1 & + (ALL x::hrat. hreal.cut X x --> hrat_lt (hrat_mul w x) d))" + +lemma hreal_inv: "ALL X::hreal. hreal_inv X = hreal - (%w. EX d. hrat_lt d hrat_1 & - (ALL x. hreal.cut X x --> hrat_lt (hrat_mul w x) d))" + (%w::hrat. + EX d::hrat. + hrat_lt d hrat_1 & + (ALL x::hrat. hreal.cut X x --> hrat_lt (hrat_mul w x) d))" by (import hreal hreal_inv) constdefs hreal_sup :: "(hreal => bool) => hreal" - "hreal_sup == %P. hreal (%w. EX X. P X & hreal.cut X w)" - -lemma hreal_sup: "ALL P. hreal_sup P = hreal (%w. EX X. P X & hreal.cut X w)" + "hreal_sup == +%P::hreal => bool. hreal (%w::hrat. EX X::hreal. P X & hreal.cut X w)" + +lemma hreal_sup: "ALL P::hreal => bool. + hreal_sup P = hreal (%w::hrat. EX X::hreal. P X & hreal.cut X w)" by (import hreal hreal_sup) constdefs hreal_lt :: "hreal => hreal => bool" - "hreal_lt == %X Y. X ~= Y & (ALL x. hreal.cut X x --> hreal.cut Y x)" - -lemma hreal_lt: "ALL X Y. hreal_lt X Y = (X ~= Y & (ALL x. hreal.cut X x --> hreal.cut Y x))" + "hreal_lt == +%(X::hreal) Y::hreal. + X ~= Y & (ALL x::hrat. hreal.cut X x --> hreal.cut Y x)" + +lemma hreal_lt: "ALL (X::hreal) Y::hreal. + hreal_lt X Y = (X ~= Y & (ALL x::hrat. hreal.cut X x --> hreal.cut Y x))" by (import hreal hreal_lt) -lemma HREAL_INV_ISACUT: "ALL X. +lemma HREAL_INV_ISACUT: "ALL X::hreal. isacut - (%w. EX d. hrat_lt d hrat_1 & - (ALL x. hreal.cut X x --> hrat_lt (hrat_mul w x) d))" + (%w::hrat. + EX d::hrat. + hrat_lt d hrat_1 & + (ALL x::hrat. hreal.cut X x --> hrat_lt (hrat_mul w x) d))" by (import hreal HREAL_INV_ISACUT) -lemma HREAL_ADD_ISACUT: "ALL X Y. - isacut (%w. EX x y. w = hrat_add x y & hreal.cut X x & hreal.cut Y y)" +lemma HREAL_ADD_ISACUT: "ALL (X::hreal) Y::hreal. + isacut + (%w::hrat. + EX (x::hrat) y::hrat. + w = hrat_add x y & hreal.cut X x & hreal.cut Y y)" by (import hreal HREAL_ADD_ISACUT) -lemma HREAL_MUL_ISACUT: "ALL X Y. - isacut (%w. EX x y. w = hrat_mul x y & hreal.cut X x & hreal.cut Y y)" +lemma HREAL_MUL_ISACUT: "ALL (X::hreal) Y::hreal. + isacut + (%w::hrat. + EX (x::hrat) y::hrat. + w = hrat_mul x y & hreal.cut X x & hreal.cut Y y)" by (import hreal HREAL_MUL_ISACUT) -lemma HREAL_ADD_SYM: "ALL X Y. hreal_add X Y = hreal_add Y X" +lemma HREAL_ADD_SYM: "ALL (X::hreal) Y::hreal. hreal_add X Y = hreal_add Y X" by (import hreal HREAL_ADD_SYM) -lemma HREAL_MUL_SYM: "ALL X Y. hreal_mul X Y = hreal_mul Y X" +lemma HREAL_MUL_SYM: "ALL (X::hreal) Y::hreal. hreal_mul X Y = hreal_mul Y X" by (import hreal HREAL_MUL_SYM) -lemma HREAL_ADD_ASSOC: "ALL X Y Z. hreal_add X (hreal_add Y Z) = hreal_add (hreal_add X Y) Z" +lemma HREAL_ADD_ASSOC: "ALL (X::hreal) (Y::hreal) Z::hreal. + hreal_add X (hreal_add Y Z) = hreal_add (hreal_add X Y) Z" by (import hreal HREAL_ADD_ASSOC) -lemma HREAL_MUL_ASSOC: "ALL X Y Z. hreal_mul X (hreal_mul Y Z) = hreal_mul (hreal_mul X Y) Z" +lemma HREAL_MUL_ASSOC: "ALL (X::hreal) (Y::hreal) Z::hreal. + hreal_mul X (hreal_mul Y Z) = hreal_mul (hreal_mul X Y) Z" by (import hreal HREAL_MUL_ASSOC) -lemma HREAL_LDISTRIB: "ALL X Y Z. +lemma HREAL_LDISTRIB: "ALL (X::hreal) (Y::hreal) Z::hreal. hreal_mul X (hreal_add Y Z) = hreal_add (hreal_mul X Y) (hreal_mul X Z)" by (import hreal HREAL_LDISTRIB) -lemma HREAL_MUL_LID: "ALL X. hreal_mul hreal_1 X = X" +lemma HREAL_MUL_LID: "ALL X::hreal. hreal_mul hreal_1 X = X" by (import hreal HREAL_MUL_LID) -lemma HREAL_MUL_LINV: "ALL X. hreal_mul (hreal_inv X) X = hreal_1" +lemma HREAL_MUL_LINV: "ALL X::hreal. hreal_mul (hreal_inv X) X = hreal_1" by (import hreal HREAL_MUL_LINV) -lemma HREAL_NOZERO: "ALL X Y. hreal_add X Y ~= X" +lemma HREAL_NOZERO: "ALL (X::hreal) Y::hreal. hreal_add X Y ~= X" by (import hreal HREAL_NOZERO) constdefs hreal_sub :: "hreal => hreal => hreal" "hreal_sub == -%Y X. hreal (%w. EX x. ~ hreal.cut X x & hreal.cut Y (hrat_add x w))" - -lemma hreal_sub: "ALL Y X. +%(Y::hreal) X::hreal. + hreal + (%w::hrat. EX x::hrat. ~ hreal.cut X x & hreal.cut Y (hrat_add x w))" + +lemma hreal_sub: "ALL (Y::hreal) X::hreal. hreal_sub Y X = - hreal (%w. EX x. ~ hreal.cut X x & hreal.cut Y (hrat_add x w))" + hreal + (%w::hrat. EX x::hrat. ~ hreal.cut X x & hreal.cut Y (hrat_add x w))" by (import hreal hreal_sub) -lemma HREAL_LT_LEMMA: "ALL X Y. hreal_lt X Y --> (EX x. ~ hreal.cut X x & hreal.cut Y x)" +lemma HREAL_LT_LEMMA: "ALL (X::hreal) Y::hreal. + hreal_lt X Y --> (EX x::hrat. ~ hreal.cut X x & hreal.cut Y x)" by (import hreal HREAL_LT_LEMMA) -lemma HREAL_SUB_ISACUT: "ALL X Y. +lemma HREAL_SUB_ISACUT: "ALL (X::hreal) Y::hreal. hreal_lt X Y --> - isacut (%w. EX x. ~ hreal.cut X x & hreal.cut Y (hrat_add x w))" + isacut + (%w::hrat. EX x::hrat. ~ hreal.cut X x & hreal.cut Y (hrat_add x w))" by (import hreal HREAL_SUB_ISACUT) -lemma HREAL_SUB_ADD: "ALL X Y. hreal_lt X Y --> hreal_add (hreal_sub Y X) X = Y" +lemma HREAL_SUB_ADD: "ALL (X::hreal) Y::hreal. hreal_lt X Y --> hreal_add (hreal_sub Y X) X = Y" by (import hreal HREAL_SUB_ADD) -lemma HREAL_LT_TOTAL: "ALL X Y. X = Y | hreal_lt X Y | hreal_lt Y X" +lemma HREAL_LT_TOTAL: "ALL (X::hreal) Y::hreal. X = Y | hreal_lt X Y | hreal_lt Y X" by (import hreal HREAL_LT_TOTAL) -lemma HREAL_LT: "ALL X Y. hreal_lt X Y = (EX D. Y = hreal_add X D)" +lemma HREAL_LT: "ALL (X::hreal) Y::hreal. hreal_lt X Y = (EX D::hreal. Y = hreal_add X D)" by (import hreal HREAL_LT) -lemma HREAL_ADD_TOTAL: "ALL X Y. X = Y | (EX D. Y = hreal_add X D) | (EX D. X = hreal_add Y D)" +lemma HREAL_ADD_TOTAL: "ALL (X::hreal) Y::hreal. + X = Y | + (EX D::hreal. Y = hreal_add X D) | (EX D::hreal. X = hreal_add Y D)" by (import hreal HREAL_ADD_TOTAL) -lemma HREAL_SUP_ISACUT: "ALL P. - Ex P & (EX Y. ALL X. P X --> hreal_lt X Y) --> - isacut (%w. EX X. P X & hreal.cut X w)" +lemma HREAL_SUP_ISACUT: "ALL P::hreal => bool. + Ex P & (EX Y::hreal. ALL X::hreal. P X --> hreal_lt X Y) --> + isacut (%w::hrat. EX X::hreal. P X & hreal.cut X w)" by (import hreal HREAL_SUP_ISACUT) -lemma HREAL_SUP: "ALL P. - Ex P & (EX Y. ALL X. P X --> hreal_lt X Y) --> - (ALL Y. (EX X. P X & hreal_lt Y X) = hreal_lt Y (hreal_sup P))" +lemma HREAL_SUP: "ALL P::hreal => bool. + Ex P & (EX Y::hreal. ALL X::hreal. P X --> hreal_lt X Y) --> + (ALL Y::hreal. + (EX X::hreal. P X & hreal_lt Y X) = hreal_lt Y (hreal_sup P))" by (import hreal HREAL_SUP) ;end_setup @@ -1996,22 +2364,22 @@ ;setup_theory numeral lemma numeral_suc: "Suc ALT_ZERO = NUMERAL_BIT1 ALT_ZERO & -(ALL x. Suc (NUMERAL_BIT1 x) = NUMERAL_BIT2 x) & -(ALL x. Suc (NUMERAL_BIT2 x) = NUMERAL_BIT1 (Suc x))" +(ALL x::nat. Suc (NUMERAL_BIT1 x) = NUMERAL_BIT2 x) & +(ALL x::nat. Suc (NUMERAL_BIT2 x) = NUMERAL_BIT1 (Suc x))" by (import numeral numeral_suc) constdefs iZ :: "nat => nat" - "iZ == %x. x" - -lemma iZ: "ALL x. iZ x = x" + "iZ == %x::nat. x" + +lemma iZ: "ALL x::nat. iZ x = x" by (import numeral iZ) constdefs iiSUC :: "nat => nat" - "iiSUC == %n. Suc (Suc n)" - -lemma iiSUC: "ALL n. iiSUC n = Suc (Suc n)" + "iiSUC == %n::nat. Suc (Suc n)" + +lemma iiSUC: "ALL n::nat. iiSUC n = Suc (Suc n)" by (import numeral iiSUC) lemma numeral_distrib: "(ALL x::nat. (0::nat) + x = x) & @@ -2051,11 +2419,11 @@ by (import numeral numeral_distrib) lemma numeral_iisuc: "iiSUC ALT_ZERO = NUMERAL_BIT2 ALT_ZERO & -iiSUC (NUMERAL_BIT1 n) = NUMERAL_BIT1 (Suc n) & +iiSUC (NUMERAL_BIT1 (n::nat)) = NUMERAL_BIT1 (Suc n) & iiSUC (NUMERAL_BIT2 n) = NUMERAL_BIT2 (Suc n)" by (import numeral numeral_iisuc) -lemma numeral_add: "ALL x xa. +lemma numeral_add: "ALL (x::nat) xa::nat. iZ (ALT_ZERO + x) = x & iZ (x + ALT_ZERO) = x & iZ (NUMERAL_BIT1 x + NUMERAL_BIT1 xa) = NUMERAL_BIT2 (iZ (x + xa)) & @@ -2078,7 +2446,7 @@ iiSUC (NUMERAL_BIT2 x + NUMERAL_BIT2 xa) = NUMERAL_BIT2 (iiSUC (x + xa))" by (import numeral numeral_add) -lemma numeral_eq: "ALL x xa. +lemma numeral_eq: "ALL (x::nat) xa::nat. (ALT_ZERO = NUMERAL_BIT1 x) = False & (NUMERAL_BIT1 x = ALT_ZERO) = False & (ALT_ZERO = NUMERAL_BIT2 x) = False & @@ -2089,7 +2457,7 @@ (NUMERAL_BIT2 x = NUMERAL_BIT2 xa) = (x = xa)" by (import numeral numeral_eq) -lemma numeral_lt: "ALL x xa. +lemma numeral_lt: "ALL (x::nat) xa::nat. (ALT_ZERO < NUMERAL_BIT1 x) = True & (ALT_ZERO < NUMERAL_BIT2 x) = True & (x < ALT_ZERO) = False & @@ -2099,7 +2467,7 @@ (NUMERAL_BIT2 x < NUMERAL_BIT1 xa) = (x < xa)" by (import numeral numeral_lt) -lemma numeral_lte: "ALL x xa. +lemma numeral_lte: "ALL (x::nat) xa::nat. (ALT_ZERO <= x) = True & (NUMERAL_BIT1 x <= ALT_ZERO) = False & (NUMERAL_BIT2 x <= ALT_ZERO) = False & @@ -2111,67 +2479,70 @@ lemma numeral_pre: "PRE ALT_ZERO = ALT_ZERO & PRE (NUMERAL_BIT1 ALT_ZERO) = ALT_ZERO & -(ALL x. +(ALL x::nat. PRE (NUMERAL_BIT1 (NUMERAL_BIT1 x)) = NUMERAL_BIT2 (PRE (NUMERAL_BIT1 x))) & -(ALL x. +(ALL x::nat. PRE (NUMERAL_BIT1 (NUMERAL_BIT2 x)) = NUMERAL_BIT2 (NUMERAL_BIT1 x)) & -(ALL x. PRE (NUMERAL_BIT2 x) = NUMERAL_BIT1 x)" +(ALL x::nat. PRE (NUMERAL_BIT2 x) = NUMERAL_BIT1 x)" by (import numeral numeral_pre) -lemma bit_initiality: "ALL zf b1f b2f. - EX x. x ALT_ZERO = zf & - (ALL n. x (NUMERAL_BIT1 n) = b1f n (x n)) & - (ALL n. x (NUMERAL_BIT2 n) = b2f n (x n))" +lemma bit_initiality: "ALL (zf::'a::type) (b1f::nat => 'a::type => 'a::type) + b2f::nat => 'a::type => 'a::type. + EX x::nat => 'a::type. + x ALT_ZERO = zf & + (ALL n::nat. x (NUMERAL_BIT1 n) = b1f n (x n)) & + (ALL n::nat. x (NUMERAL_BIT2 n) = b2f n (x n))" by (import numeral bit_initiality) consts - iBIT_cases :: "nat => 'a => (nat => 'a) => (nat => 'a) => 'a" - -specification (iBIT_cases) iBIT_cases: "(ALL (zf::'a) (bf1::nat => 'a) bf2::nat => 'a. + iBIT_cases :: "nat => 'a::type => (nat => 'a::type) => (nat => 'a::type) => 'a::type" + +specification (iBIT_cases) iBIT_cases: "(ALL (zf::'a::type) (bf1::nat => 'a::type) bf2::nat => 'a::type. iBIT_cases ALT_ZERO zf bf1 bf2 = zf) & -(ALL (n::nat) (zf::'a) (bf1::nat => 'a) bf2::nat => 'a. +(ALL (n::nat) (zf::'a::type) (bf1::nat => 'a::type) bf2::nat => 'a::type. iBIT_cases (NUMERAL_BIT1 n) zf bf1 bf2 = bf1 n) & -(ALL (n::nat) (zf::'a) (bf1::nat => 'a) bf2::nat => 'a. +(ALL (n::nat) (zf::'a::type) (bf1::nat => 'a::type) bf2::nat => 'a::type. iBIT_cases (NUMERAL_BIT2 n) zf bf1 bf2 = bf2 n)" by (import numeral iBIT_cases) constdefs iDUB :: "nat => nat" - "iDUB == %x. x + x" - -lemma iDUB: "ALL x. iDUB x = x + x" + "iDUB == %x::nat. x + x" + +lemma iDUB: "ALL x::nat. iDUB x = x + x" by (import numeral iDUB) consts iSUB :: "bool => nat => nat => nat" -specification (iSUB) iSUB_DEF: "(ALL b x. iSUB b ALT_ZERO x = ALT_ZERO) & -(ALL b n x. +specification (iSUB) iSUB_DEF: "(ALL (b::bool) x::nat. iSUB b ALT_ZERO x = ALT_ZERO) & +(ALL (b::bool) (n::nat) x::nat. iSUB b (NUMERAL_BIT1 n) x = (if b - then iBIT_cases x (NUMERAL_BIT1 n) (%m. iDUB (iSUB True n m)) - (%m. NUMERAL_BIT1 (iSUB False n m)) - else iBIT_cases x (iDUB n) (%m. NUMERAL_BIT1 (iSUB False n m)) - (%m. iDUB (iSUB False n m)))) & -(ALL b n x. + then iBIT_cases x (NUMERAL_BIT1 n) (%m::nat. iDUB (iSUB True n m)) + (%m::nat. NUMERAL_BIT1 (iSUB False n m)) + else iBIT_cases x (iDUB n) (%m::nat. NUMERAL_BIT1 (iSUB False n m)) + (%m::nat. iDUB (iSUB False n m)))) & +(ALL (b::bool) (n::nat) x::nat. iSUB b (NUMERAL_BIT2 n) x = (if b - then iBIT_cases x (NUMERAL_BIT2 n) (%m. NUMERAL_BIT1 (iSUB True n m)) - (%m. iDUB (iSUB True n m)) - else iBIT_cases x (NUMERAL_BIT1 n) (%m. iDUB (iSUB True n m)) - (%m. NUMERAL_BIT1 (iSUB False n m))))" + then iBIT_cases x (NUMERAL_BIT2 n) + (%m::nat. NUMERAL_BIT1 (iSUB True n m)) + (%m::nat. iDUB (iSUB True n m)) + else iBIT_cases x (NUMERAL_BIT1 n) (%m::nat. iDUB (iSUB True n m)) + (%m::nat. NUMERAL_BIT1 (iSUB False n m))))" by (import numeral iSUB_DEF) -lemma bit_induction: "ALL P. +lemma bit_induction: "ALL P::nat => bool. P ALT_ZERO & - (ALL n. P n --> P (NUMERAL_BIT1 n)) & - (ALL n. P n --> P (NUMERAL_BIT2 n)) --> + (ALL n::nat. P n --> P (NUMERAL_BIT1 n)) & + (ALL n::nat. P n --> P (NUMERAL_BIT2 n)) --> All P" by (import numeral bit_induction) -lemma iSUB_THM: "ALL xa xb xc. - iSUB xa ALT_ZERO x = ALT_ZERO & +lemma iSUB_THM: "ALL (xa::bool) (xb::nat) xc::nat. + iSUB xa ALT_ZERO (x::nat) = ALT_ZERO & iSUB True xb ALT_ZERO = xb & iSUB False (NUMERAL_BIT1 xb) ALT_ZERO = iDUB xb & iSUB True (NUMERAL_BIT1 xb) (NUMERAL_BIT1 xc) = iDUB (iSUB True xb xc) & @@ -2190,17 +2561,26 @@ NUMERAL_BIT1 (iSUB False xb xc)" by (import numeral iSUB_THM) -lemma numeral_sub: "ALL x xa. - NUMERAL (x - xa) = (if xa < x then NUMERAL (iSUB True x xa) else 0)" +lemma numeral_sub: "(All::(nat => bool) => bool) + (%x::nat. + (All::(nat => bool) => bool) + (%xa::nat. + (op =::nat => nat => bool) + ((NUMERAL::nat => nat) ((op -::nat => nat => nat) x xa)) + ((If::bool => nat => nat => nat) + ((op <::nat => nat => bool) xa x) + ((NUMERAL::nat => nat) + ((iSUB::bool => nat => nat => nat) (True::bool) x xa)) + (0::nat))))" by (import numeral numeral_sub) -lemma iDUB_removal: "ALL x. +lemma iDUB_removal: "ALL x::nat. iDUB (NUMERAL_BIT1 x) = NUMERAL_BIT2 (iDUB x) & iDUB (NUMERAL_BIT2 x) = NUMERAL_BIT2 (NUMERAL_BIT1 x) & iDUB ALT_ZERO = ALT_ZERO" by (import numeral iDUB_removal) -lemma numeral_mult: "ALL x xa. +lemma numeral_mult: "ALL (x::nat) xa::nat. ALT_ZERO * x = ALT_ZERO & x * ALT_ZERO = ALT_ZERO & NUMERAL_BIT1 x * xa = iZ (iDUB (x * xa) + xa) & @@ -2209,265 +2589,316 @@ constdefs iSQR :: "nat => nat" - "iSQR == %x. x * x" - -lemma iSQR: "ALL x. iSQR x = x * x" + "iSQR == %x::nat. x * x" + +lemma iSQR: "ALL x::nat. iSQR x = x * x" by (import numeral iSQR) -lemma numeral_exp: "(ALL x. x ^ ALT_ZERO = NUMERAL_BIT1 ALT_ZERO) & -(ALL x xa. x ^ NUMERAL_BIT1 xa = x * iSQR (x ^ xa)) & -(ALL x xa. x ^ NUMERAL_BIT2 xa = iSQR x * iSQR (x ^ xa))" +lemma numeral_exp: "(ALL x::nat. x ^ ALT_ZERO = NUMERAL_BIT1 ALT_ZERO) & +(ALL (x::nat) xa::nat. x ^ NUMERAL_BIT1 xa = x * iSQR (x ^ xa)) & +(ALL (x::nat) xa::nat. x ^ NUMERAL_BIT2 xa = iSQR x * iSQR (x ^ xa))" by (import numeral numeral_exp) -lemma numeral_evenodd: "ALL x. +lemma numeral_evenodd: "ALL x::nat. EVEN ALT_ZERO & EVEN (NUMERAL_BIT2 x) & ~ EVEN (NUMERAL_BIT1 x) & ~ ODD ALT_ZERO & ~ ODD (NUMERAL_BIT2 x) & ODD (NUMERAL_BIT1 x)" by (import numeral numeral_evenodd) -lemma numeral_fact: "ALL n. FACT n = (if n = 0 then 1 else n * FACT (PRE n))" +lemma numeral_fact: "ALL n::nat. FACT n = (if n = (0::nat) then 1::nat else n * FACT (PRE n))" by (import numeral numeral_fact) -lemma numeral_funpow: "ALL n. (f ^ n) x = (if n = 0 then x else (f ^ (n - 1)) (f x))" +lemma numeral_funpow: "ALL n::nat. + ((f::'a::type => 'a::type) ^ n) (x::'a::type) = + (if n = (0::nat) then x else (f ^ (n - (1::nat))) (f x))" by (import numeral numeral_funpow) ;end_setup ;setup_theory ind_type -lemma INJ_INVERSE2: "ALL P::'A => 'B => 'C. - (ALL (x1::'A) (y1::'B) (x2::'A) y2::'B. +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 => 'A) Y::'C => 'B. - ALL (xa::'A) y::'B. x (P xa y) = xa & Y (P xa y) = y)" + (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)" by (import ind_type INJ_INVERSE2) constdefs NUMPAIR :: "nat => nat => nat" - "NUMPAIR == %x y. 2 ^ x * (2 * y + 1)" - -lemma NUMPAIR: "ALL x y. NUMPAIR x y = 2 ^ x * (2 * y + 1)" + "NUMPAIR == %(x::nat) y::nat. (2::nat) ^ x * ((2::nat) * y + (1::nat))" + +lemma NUMPAIR: "ALL (x::nat) y::nat. NUMPAIR x y = (2::nat) ^ x * ((2::nat) * y + (1::nat))" by (import ind_type NUMPAIR) -lemma NUMPAIR_INJ_LEMMA: "ALL x xa xb xc. NUMPAIR x xa = NUMPAIR xb xc --> x = xb" +lemma NUMPAIR_INJ_LEMMA: "ALL (x::nat) (xa::nat) (xb::nat) xc::nat. + NUMPAIR x xa = NUMPAIR xb xc --> x = xb" by (import ind_type NUMPAIR_INJ_LEMMA) -lemma NUMPAIR_INJ: "ALL x1 y1 x2 y2. (NUMPAIR x1 y1 = NUMPAIR x2 y2) = (x1 = x2 & y1 = y2)" +lemma NUMPAIR_INJ: "ALL (x1::nat) (y1::nat) (x2::nat) y2::nat. + (NUMPAIR x1 y1 = NUMPAIR x2 y2) = (x1 = x2 & y1 = y2)" by (import ind_type NUMPAIR_INJ) consts NUMSND :: "nat => nat" NUMFST :: "nat => nat" -specification (NUMFST NUMSND) NUMPAIR_DEST: "ALL x y. NUMFST (NUMPAIR x y) = x & NUMSND (NUMPAIR x y) = y" +specification (NUMFST NUMSND) NUMPAIR_DEST: "ALL (x::nat) y::nat. NUMFST (NUMPAIR x y) = x & NUMSND (NUMPAIR x y) = y" by (import ind_type NUMPAIR_DEST) constdefs NUMSUM :: "bool => nat => nat" - "NUMSUM == %b x. if b then Suc (2 * x) else 2 * x" - -lemma NUMSUM: "ALL b x. NUMSUM b x = (if b then Suc (2 * x) else 2 * x)" + "NUMSUM == %(b::bool) x::nat. if b then Suc ((2::nat) * x) else (2::nat) * x" + +lemma NUMSUM: "ALL (b::bool) x::nat. + NUMSUM b x = (if b then Suc ((2::nat) * x) else (2::nat) * x)" by (import ind_type NUMSUM) -lemma NUMSUM_INJ: "ALL b1 x1 b2 x2. (NUMSUM b1 x1 = NUMSUM b2 x2) = (b1 = b2 & x1 = x2)" +lemma NUMSUM_INJ: "ALL (b1::bool) (x1::nat) (b2::bool) x2::nat. + (NUMSUM b1 x1 = NUMSUM b2 x2) = (b1 = b2 & x1 = x2)" by (import ind_type NUMSUM_INJ) consts NUMRIGHT :: "nat => nat" NUMLEFT :: "nat => bool" -specification (NUMLEFT NUMRIGHT) NUMSUM_DEST: "ALL x y. NUMLEFT (NUMSUM x y) = x & NUMRIGHT (NUMSUM x y) = y" +specification (NUMLEFT NUMRIGHT) NUMSUM_DEST: "ALL (x::bool) y::nat. NUMLEFT (NUMSUM x y) = x & NUMRIGHT (NUMSUM x y) = y" by (import ind_type NUMSUM_DEST) constdefs - INJN :: "nat => nat => 'a => bool" - "INJN == %m n a. n = m" - -lemma INJN: "ALL m. INJN m = (%n a. n = m)" + INJN :: "nat => nat => 'a::type => bool" + "INJN == %(m::nat) (n::nat) a::'a::type. n = m" + +lemma INJN: "ALL m::nat. INJN m = (%(n::nat) a::'a::type. n = m)" by (import ind_type INJN) -lemma INJN_INJ: "ALL n1 n2. (INJN n1 = INJN n2) = (n1 = n2)" +lemma INJN_INJ: "ALL (n1::nat) n2::nat. (INJN n1 = INJN n2) = (n1 = n2)" by (import ind_type INJN_INJ) constdefs - INJA :: "'a => nat => 'a => bool" - "INJA == %a n b. b = a" - -lemma INJA: "ALL a. INJA a = (%n b. b = a)" + INJA :: "'a::type => nat => 'a::type => bool" + "INJA == %(a::'a::type) (n::nat) b::'a::type. b = a" + +lemma INJA: "ALL a::'a::type. INJA a = (%(n::nat) b::'a::type. b = a)" by (import ind_type INJA) -lemma INJA_INJ: "ALL a1 a2. (INJA a1 = INJA a2) = (a1 = a2)" +lemma INJA_INJ: "ALL (a1::'a::type) a2::'a::type. (INJA a1 = INJA a2) = (a1 = a2)" by (import ind_type INJA_INJ) constdefs - INJF :: "(nat => nat => 'a => bool) => nat => 'a => bool" - "INJF == %f n. f (NUMFST n) (NUMSND n)" - -lemma INJF: "ALL f. INJF f = (%n. f (NUMFST n) (NUMSND n))" + INJF :: "(nat => nat => 'a::type => bool) => nat => 'a::type => bool" + "INJF == %(f::nat => nat => 'a::type => bool) n::nat. f (NUMFST n) (NUMSND n)" + +lemma INJF: "ALL f::nat => nat => 'a::type => bool. + INJF f = (%n::nat. f (NUMFST n) (NUMSND n))" by (import ind_type INJF) -lemma INJF_INJ: "ALL f1 f2. (INJF f1 = INJF f2) = (f1 = f2)" +lemma INJF_INJ: "ALL (f1::nat => nat => 'a::type => bool) f2::nat => nat => 'a::type => bool. + (INJF f1 = INJF f2) = (f1 = f2)" by (import ind_type INJF_INJ) constdefs - INJP :: "(nat => 'a => bool) => (nat => 'a => bool) => nat => 'a => bool" + INJP :: "(nat => 'a::type => bool) +=> (nat => 'a::type => bool) => nat => 'a::type => bool" "INJP == -%f1 f2 n a. if NUMLEFT n then f1 (NUMRIGHT n) a else f2 (NUMRIGHT n) a" - -lemma INJP: "ALL f1 f2. +%(f1::nat => 'a::type => bool) (f2::nat => 'a::type => bool) (n::nat) + a::'a::type. if NUMLEFT n then f1 (NUMRIGHT n) a else f2 (NUMRIGHT n) a" + +lemma INJP: "ALL (f1::nat => 'a::type => bool) f2::nat => 'a::type => bool. INJP f1 f2 = - (%n a. if NUMLEFT n then f1 (NUMRIGHT n) a else f2 (NUMRIGHT n) a)" + (%(n::nat) a::'a::type. + if NUMLEFT n then f1 (NUMRIGHT n) a else f2 (NUMRIGHT n) a)" by (import ind_type INJP) -lemma INJP_INJ: "ALL f1 f1' f2 f2'. (INJP f1 f2 = INJP f1' f2') = (f1 = f1' & f2 = f2')" +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')" by (import ind_type INJP_INJ) constdefs - ZCONSTR :: "nat => 'a => (nat => nat => 'a => bool) => nat => 'a => bool" - "ZCONSTR == %c i r. INJP (INJN (Suc c)) (INJP (INJA i) (INJF r))" - -lemma ZCONSTR: "ALL c i r. ZCONSTR c i r = INJP (INJN (Suc c)) (INJP (INJA i) (INJF r))" + ZCONSTR :: "nat +=> 'a::type => (nat => nat => 'a::type => bool) => nat => 'a::type => bool" + "ZCONSTR == +%(c::nat) (i::'a::type) r::nat => nat => 'a::type => bool. + INJP (INJN (Suc c)) (INJP (INJA i) (INJF r))" + +lemma ZCONSTR: "ALL (c::nat) (i::'a::type) r::nat => nat => 'a::type => bool. + ZCONSTR c i r = INJP (INJN (Suc c)) (INJP (INJA i) (INJF r))" by (import ind_type ZCONSTR) constdefs - ZBOT :: "nat => 'a => bool" - "ZBOT == INJP (INJN 0) (SOME z. True)" - -lemma ZBOT: "ZBOT = INJP (INJN 0) (SOME z. True)" + ZBOT :: "nat => 'a::type => bool" + "ZBOT == INJP (INJN (0::nat)) (SOME z::nat => 'a::type => bool. True)" + +lemma ZBOT: "ZBOT = INJP (INJN (0::nat)) (SOME z::nat => 'a::type => bool. True)" by (import ind_type ZBOT) -lemma ZCONSTR_ZBOT: "ALL x xa xb. ZCONSTR x xa xb ~= ZBOT" +lemma ZCONSTR_ZBOT: "ALL (x::nat) (xa::'a::type) xb::nat => nat => 'a::type => bool. + ZCONSTR x xa xb ~= ZBOT" by (import ind_type ZCONSTR_ZBOT) constdefs - ZRECSPACE :: "(nat => 'a => bool) => bool" + ZRECSPACE :: "(nat => 'a::type => bool) => bool" "ZRECSPACE == -%a0. ALL ZRECSPACE'. - (ALL a0. - a0 = ZBOT | - (EX c i r. a0 = ZCONSTR c i r & (ALL n. ZRECSPACE' (r n))) --> - ZRECSPACE' a0) --> - ZRECSPACE' a0" +%a0::nat => 'a::type => bool. + ALL ZRECSPACE'::(nat => 'a::type => bool) => bool. + (ALL a0::nat => 'a::type => bool. + a0 = ZBOT | + (EX (c::nat) (i::'a::type) r::nat => nat => 'a::type => bool. + a0 = ZCONSTR c i r & (ALL n::nat. ZRECSPACE' (r n))) --> + ZRECSPACE' a0) --> + ZRECSPACE' a0" lemma ZRECSPACE: "ZRECSPACE = -(%a0. ALL ZRECSPACE'. - (ALL a0. - a0 = ZBOT | - (EX c i r. a0 = ZCONSTR c i r & (ALL n. ZRECSPACE' (r n))) --> - ZRECSPACE' a0) --> - ZRECSPACE' a0)" +(%a0::nat => 'a::type => bool. + ALL ZRECSPACE'::(nat => 'a::type => bool) => bool. + (ALL a0::nat => 'a::type => bool. + a0 = ZBOT | + (EX (c::nat) (i::'a::type) r::nat => nat => 'a::type => bool. + a0 = ZCONSTR c i r & (ALL n::nat. ZRECSPACE' (r n))) --> + ZRECSPACE' a0) --> + ZRECSPACE' a0)" by (import ind_type ZRECSPACE) lemma ZRECSPACE_rules: "(op &::bool => bool => bool) - ((ZRECSPACE::(nat => 'a => bool) => bool) (ZBOT::nat => 'a => bool)) + ((ZRECSPACE::(nat => 'a::type => bool) => bool) + (ZBOT::nat => 'a::type => bool)) ((All::(nat => bool) => bool) (%c::nat. - (All::('a => bool) => bool) - (%i::'a. - (All::((nat => nat => 'a => bool) => bool) => bool) - (%r::nat => nat => 'a => bool. + (All::('a::type => bool) => bool) + (%i::'a::type. + (All::((nat => nat => 'a::type => bool) => bool) => bool) + (%r::nat => nat => 'a::type => bool. (op -->::bool => bool => bool) ((All::(nat => bool) => bool) (%n::nat. - (ZRECSPACE::(nat => 'a => bool) => bool) (r n))) - ((ZRECSPACE::(nat => 'a => bool) => bool) + (ZRECSPACE::(nat => 'a::type => bool) => bool) + (r n))) + ((ZRECSPACE::(nat => 'a::type => bool) => bool) ((ZCONSTR::nat - => 'a => (nat => nat => 'a => bool) - => nat => 'a => bool) + => 'a::type + => (nat => nat => 'a::type => bool) + => nat => 'a::type => bool) c i r))))))" by (import ind_type ZRECSPACE_rules) -lemma ZRECSPACE_ind: "ALL x. - x ZBOT & (ALL c i r. (ALL n. x (r n)) --> x (ZCONSTR c i r)) --> - (ALL a0. ZRECSPACE a0 --> x a0)" +lemma ZRECSPACE_ind: "ALL x::(nat => 'a::type => bool) => bool. + x ZBOT & + (ALL (c::nat) (i::'a::type) r::nat => nat => 'a::type => bool. + (ALL n::nat. x (r n)) --> x (ZCONSTR c i r)) --> + (ALL a0::nat => 'a::type => bool. ZRECSPACE a0 --> x a0)" by (import ind_type ZRECSPACE_ind) -lemma ZRECSPACE_cases: "ALL a0. +lemma ZRECSPACE_cases: "ALL a0::nat => 'a::type => bool. ZRECSPACE a0 = - (a0 = ZBOT | (EX c i r. a0 = ZCONSTR c i r & (ALL n. ZRECSPACE (r n))))" + (a0 = ZBOT | + (EX (c::nat) (i::'a::type) r::nat => nat => 'a::type => bool. + a0 = ZCONSTR c i r & (ALL n::nat. ZRECSPACE (r n))))" by (import ind_type ZRECSPACE_cases) -typedef (open) ('a) recspace = "(Collect::((nat => 'a => bool) => bool) => (nat => 'a => bool) set) - (ZRECSPACE::(nat => 'a => bool) => bool)" +typedef (open) ('a) recspace = "(Collect::((nat => 'a::type => bool) => bool) + => (nat => 'a::type => bool) set) + (ZRECSPACE::(nat => 'a::type => bool) => bool)" by (rule typedef_helper,import ind_type recspace_TY_DEF) lemmas recspace_TY_DEF = typedef_hol2hol4 [OF type_definition_recspace] consts - mk_rec :: "(nat => 'a => bool) => 'a recspace" - dest_rec :: "'a recspace => nat => 'a => bool" - -specification (dest_rec mk_rec) recspace_repfns: "(ALL a::'a recspace. mk_rec (dest_rec a) = a) & -(ALL r::nat => 'a => bool. ZRECSPACE r = (dest_rec (mk_rec r) = r))" + mk_rec :: "(nat => 'a::type => bool) => 'a::type recspace" + dest_rec :: "'a::type recspace => nat => 'a::type => bool" + +specification (dest_rec mk_rec) recspace_repfns: "(ALL a::'a::type recspace. mk_rec (dest_rec a) = a) & +(ALL r::nat => 'a::type => bool. ZRECSPACE r = (dest_rec (mk_rec r) = r))" by (import ind_type recspace_repfns) constdefs - BOTTOM :: "'a recspace" + BOTTOM :: "'a::type recspace" "BOTTOM == mk_rec ZBOT" lemma BOTTOM: "BOTTOM = mk_rec ZBOT" by (import ind_type BOTTOM) constdefs - CONSTR :: "nat => 'a => (nat => 'a recspace) => 'a recspace" - "CONSTR == %c i r. mk_rec (ZCONSTR c i (%n. dest_rec (r n)))" - -lemma CONSTR: "ALL c i r. CONSTR c i r = mk_rec (ZCONSTR c i (%n. dest_rec (r n)))" + CONSTR :: "nat => 'a::type => (nat => 'a::type recspace) => 'a::type recspace" + "CONSTR == +%(c::nat) (i::'a::type) r::nat => 'a::type recspace. + mk_rec (ZCONSTR c i (%n::nat. dest_rec (r n)))" + +lemma CONSTR: "ALL (c::nat) (i::'a::type) r::nat => 'a::type recspace. + CONSTR c i r = mk_rec (ZCONSTR c i (%n::nat. dest_rec (r n)))" by (import ind_type CONSTR) -lemma MK_REC_INJ: "ALL x y. mk_rec x = mk_rec y --> ZRECSPACE x & ZRECSPACE y --> x = y" +lemma MK_REC_INJ: "ALL (x::nat => 'a::type => bool) y::nat => 'a::type => bool. + mk_rec x = mk_rec y --> ZRECSPACE x & ZRECSPACE y --> x = y" by (import ind_type MK_REC_INJ) -lemma DEST_REC_INJ: "ALL x y. (dest_rec x = dest_rec y) = (x = y)" +lemma DEST_REC_INJ: "ALL (x::'a::type recspace) y::'a::type recspace. + (dest_rec x = dest_rec y) = (x = y)" by (import ind_type DEST_REC_INJ) -lemma CONSTR_BOT: "ALL c i r. CONSTR c i r ~= BOTTOM" +lemma CONSTR_BOT: "ALL (c::nat) (i::'a::type) r::nat => 'a::type recspace. + CONSTR c i r ~= BOTTOM" by (import ind_type CONSTR_BOT) -lemma CONSTR_INJ: "ALL c1 i1 r1 c2 i2 r2. +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)" by (import ind_type CONSTR_INJ) -lemma CONSTR_IND: "ALL P. - P BOTTOM & (ALL c i r. (ALL n. P (r n)) --> P (CONSTR c i r)) --> All P" +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" by (import ind_type CONSTR_IND) -lemma CONSTR_REC: "ALL Fn. EX f. ALL c i r. f (CONSTR c i r) = Fn c i r (%n. f (r n))" +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))" by (import ind_type CONSTR_REC) consts - FCONS :: "'a => (nat => 'a) => nat => 'a" - -specification (FCONS) FCONS: "(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)" + FCONS :: "'a::type => (nat => 'a::type) => nat => 'a::type" + +specification (FCONS) FCONS: "(ALL (a::'a::type) f::nat => 'a::type. FCONS a f (0::nat) = a) & +(ALL (a::'a::type) (f::nat => 'a::type) n::nat. FCONS a f (Suc n) = f n)" by (import ind_type FCONS) constdefs - FNIL :: "nat => 'a" - "FNIL == %n. SOME x. True" - -lemma FNIL: "ALL n. FNIL n = (SOME x. True)" + FNIL :: "nat => 'a::type" + "FNIL == %n::nat. SOME x::'a::type. True" + +lemma FNIL: "ALL n::nat. FNIL n = (SOME x::'a::type. True)" by (import ind_type FNIL) constdefs - ISO :: "('a => 'b) => ('b => 'a) => bool" - "ISO == %f g. (ALL x. f (g x) = x) & (ALL y. g (f y) = y)" - -lemma ISO: "ALL f g. ISO f g = ((ALL x. f (g x) = x) & (ALL y. g (f y) = y))" + ISO :: "('a::type => 'b::type) => ('b::type => 'a::type) => bool" + "ISO == +%(f::'a::type => 'b::type) g::'b::type => 'a::type. + (ALL x::'b::type. f (g x) = x) & (ALL y::'a::type. g (f y) = y)" + +lemma ISO: "ALL (f::'a::type => 'b::type) g::'b::type => 'a::type. + ISO f g = + ((ALL x::'b::type. f (g x) = x) & (ALL y::'a::type. g (f y) = y))" by (import ind_type ISO) -lemma ISO_REFL: "ISO (%x. x) (%x. x)" +lemma ISO_REFL: "ISO (%x::'a::type. x) (%x::'a::type. x)" by (import ind_type ISO_REFL) -lemma ISO_FUN: "ISO (f::'a => 'c) (f'::'c => 'a) & ISO (g::'b => 'd) (g'::'d => 'b) --> -ISO (%(h::'a => 'b) a'::'c. g (h (f' a'))) - (%(h::'c => 'd) a::'a. g' (h (f a)))" +lemma ISO_FUN: "ISO (f::'a::type => 'c::type) (f'::'c::type => 'a::type) & +ISO (g::'b::type => 'd::type) (g'::'d::type => 'b::type) --> +ISO (%(h::'a::type => 'b::type) a'::'c::type. g (h (f' a'))) + (%(h::'c::type => 'd::type) a::'a::type. g' (h (f a)))" by (import ind_type ISO_FUN) -lemma ISO_USAGE: "ISO f g --> -(ALL P. All P = (ALL x. P (g x))) & -(ALL P. Ex P = (EX x. P (g x))) & (ALL a b. (a = g b) = (f a = b))" +lemma ISO_USAGE: "ISO (f::'a::type => 'b::type) (g::'b::type => 'a::type) --> +(ALL P::'a::type => bool. All P = (ALL x::'b::type. P (g x))) & +(ALL P::'a::type => bool. Ex P = (EX x::'b::type. P (g x))) & +(ALL (a::'a::type) b::'b::type. (a = g b) = (f a = b))" by (import ind_type ISO_USAGE) ;end_setup @@ -2480,7 +2911,7 @@ lemma DIVIDES_ADD_2: "ALL (a::nat) (b::nat) c::nat. a dvd b & a dvd b + c --> a dvd c" by (import divides DIVIDES_ADD_2) -lemma DIVIDES_FACT: "ALL b>0. b dvd FACT b" +lemma DIVIDES_FACT: "ALL b>0::nat. b dvd FACT b" by (import divides DIVIDES_FACT) lemma DIVIDES_MULT_LEFT: "ALL (x::nat) xa::nat. (x * xa dvd xa) = (xa = (0::nat) | x = (1::nat))" @@ -2494,15 +2925,18 @@ prime :: "nat => bool" defs - prime_primdef: "prime.prime == %a. a ~= 1 & (ALL b. b dvd a --> b = a | b = 1)" - -lemma prime_def: "ALL a. prime.prime a = (a ~= 1 & (ALL b. b dvd a --> b = a | b = 1))" + prime_primdef: "prime.prime == +%a::nat. a ~= (1::nat) & (ALL b::nat. b dvd a --> b = a | b = (1::nat))" + +lemma prime_def: "ALL a::nat. + prime.prime a = + (a ~= (1::nat) & (ALL b::nat. b dvd a --> b = a | b = (1::nat)))" by (import prime prime_def) -lemma NOT_PRIME_0: "~ prime.prime 0" +lemma NOT_PRIME_0: "~ prime.prime (0::nat)" by (import prime NOT_PRIME_0) -lemma NOT_PRIME_1: "~ prime.prime 1" +lemma NOT_PRIME_1: "~ prime.prime (1::nat)" by (import prime NOT_PRIME_1) ;end_setup @@ -2510,224 +2944,286 @@ ;setup_theory list consts - EL :: "nat => 'a list => 'a" - -specification (EL) EL: "(ALL l::'a list. EL (0::nat) l = hd l) & -(ALL (l::'a list) n::nat. EL (Suc n) l = EL n (tl l))" + EL :: "nat => 'a::type list => 'a::type" + +specification (EL) EL: "(ALL l::'a::type list. EL (0::nat) l = hd l) & +(ALL (l::'a::type list) n::nat. EL (Suc n) l = EL n (tl l))" by (import list EL) -lemma NULL: "(op &::bool => bool => bool) ((null::'a list => bool) ([]::'a list)) - ((All::('a => bool) => bool) - (%x::'a. - (All::('a list => bool) => bool) - (%xa::'a list. +lemma NULL: "(op &::bool => bool => bool) + ((null::'a::type list => bool) ([]::'a::type list)) + ((All::('a::type => bool) => bool) + (%x::'a::type. + (All::('a::type list => bool) => bool) + (%xa::'a::type list. (Not::bool => bool) - ((null::'a list => bool) - ((op #::'a => 'a list => 'a list) x xa)))))" + ((null::'a::type list => bool) + ((op #::'a::type => 'a::type list => 'a::type list) x xa)))))" by (import list NULL) -lemma list_case_compute: "ALL l. list_case b f l = (if null l then b else f (hd l) (tl l))" +lemma list_case_compute: "ALL l::'a::type list. + list_case (b::'b::type) (f::'a::type => 'a::type list => 'b::type) l = + (if null l then b else f (hd l) (tl l))" by (import list list_case_compute) -lemma LIST_NOT_EQ: "ALL l1 l2. l1 ~= l2 --> (ALL x xa. x # l1 ~= xa # l2)" +lemma LIST_NOT_EQ: "ALL (l1::'a::type list) l2::'a::type list. + l1 ~= l2 --> (ALL (x::'a::type) xa::'a::type. x # l1 ~= xa # l2)" by (import list LIST_NOT_EQ) -lemma NOT_EQ_LIST: "ALL h1 h2. h1 ~= h2 --> (ALL x xa. h1 # x ~= h2 # xa)" +lemma NOT_EQ_LIST: "ALL (h1::'a::type) h2::'a::type. + h1 ~= h2 --> + (ALL (x::'a::type list) xa::'a::type list. h1 # x ~= h2 # xa)" by (import list NOT_EQ_LIST) -lemma EQ_LIST: "ALL h1 h2. h1 = h2 --> (ALL l1 l2. l1 = l2 --> h1 # l1 = h2 # l2)" +lemma EQ_LIST: "ALL (h1::'a::type) h2::'a::type. + h1 = h2 --> + (ALL (l1::'a::type list) l2::'a::type list. + l1 = l2 --> h1 # l1 = h2 # l2)" by (import list EQ_LIST) -lemma CONS: "ALL l. ~ null l --> hd l # tl l = l" +lemma CONS: "ALL l::'a::type list. ~ null l --> hd l # tl l = l" by (import list CONS) -lemma MAP_EQ_NIL: "ALL l f. (map f l = []) = (l = []) & ([] = map f l) = (l = [])" +lemma MAP_EQ_NIL: "ALL (l::'a::type list) f::'a::type => 'b::type. + (map f l = []) = (l = []) & ([] = map f l) = (l = [])" by (import list MAP_EQ_NIL) -lemma EVERY_EL: "(All::('a list => bool) => bool) - (%l::'a list. - (All::(('a => bool) => bool) => bool) - (%P::'a => bool. +lemma EVERY_EL: "(All::('a::type list => bool) => bool) + (%l::'a::type list. + (All::(('a::type => bool) => bool) => bool) + (%P::'a::type => bool. (op =::bool => bool => bool) - ((list_all::('a => bool) => 'a list => bool) P l) + ((list_all::('a::type => bool) => 'a::type list => bool) P l) ((All::(nat => bool) => bool) (%n::nat. (op -->::bool => bool => bool) - ((op <::nat => nat => bool) n ((size::'a list => nat) l)) - (P ((EL::nat => 'a list => 'a) n l))))))" + ((op <::nat => nat => bool) n + ((size::'a::type list => nat) l)) + (P ((EL::nat => 'a::type list => 'a::type) n l))))))" by (import list EVERY_EL) -lemma EVERY_CONJ: "ALL l. list_all (%x. P x & Q x) l = (list_all P l & list_all Q l)" +lemma EVERY_CONJ: "ALL l::'a::type list. + list_all + (%x::'a::type. (P::'a::type => bool) x & (Q::'a::type => bool) x) l = + (list_all P l & list_all Q l)" by (import list EVERY_CONJ) -lemma EVERY_MEM: "ALL P l. list_all P l = (ALL e. e mem l --> P e)" +lemma EVERY_MEM: "ALL (P::'a::type => bool) l::'a::type list. + list_all P l = (ALL e::'a::type. e mem l --> P e)" by (import list EVERY_MEM) -lemma EXISTS_MEM: "ALL P l. list_exists P l = (EX e. e mem l & P e)" +lemma EXISTS_MEM: "ALL (P::'a::type => bool) l::'a::type list. + list_exists P l = (EX e::'a::type. e mem l & P e)" by (import list EXISTS_MEM) -lemma MEM_APPEND: "ALL e l1 l2. e mem l1 @ l2 = (e mem l1 | e mem l2)" +lemma MEM_APPEND: "ALL (e::'a::type) (l1::'a::type list) l2::'a::type list. + e mem l1 @ l2 = (e mem l1 | e mem l2)" by (import list MEM_APPEND) -lemma EXISTS_APPEND: "ALL P l1 l2. list_exists P (l1 @ l2) = (list_exists P l1 | list_exists P l2)" +lemma EXISTS_APPEND: "ALL (P::'a::type => bool) (l1::'a::type list) l2::'a::type list. + list_exists P (l1 @ l2) = (list_exists P l1 | list_exists P l2)" by (import list EXISTS_APPEND) -lemma NOT_EVERY: "ALL P l. (~ list_all P l) = list_exists (Not o P) l" +lemma NOT_EVERY: "ALL (P::'a::type => bool) l::'a::type list. + (~ list_all P l) = list_exists (Not o P) l" by (import list NOT_EVERY) -lemma NOT_EXISTS: "ALL P l. (~ list_exists P l) = list_all (Not o P) l" +lemma NOT_EXISTS: "ALL (P::'a::type => bool) l::'a::type list. + (~ list_exists P l) = list_all (Not o P) l" by (import list NOT_EXISTS) -lemma MEM_MAP: "ALL l f x. x mem map f l = (EX y. x = f y & y mem l)" +lemma MEM_MAP: "ALL (l::'a::type list) (f::'a::type => 'b::type) x::'b::type. + x mem map f l = (EX y::'a::type. x = f y & y mem l)" by (import list MEM_MAP) -lemma LENGTH_CONS: "ALL l n. (length l = Suc n) = (EX h l'. length l' = n & l = h # l')" +lemma LENGTH_CONS: "ALL (l::'a::type list) n::nat. + (length l = Suc n) = + (EX (h::'a::type) l'::'a::type list. length l' = n & l = h # l')" by (import list LENGTH_CONS) -lemma LENGTH_EQ_CONS: "ALL P n. - (ALL l. length l = Suc n --> P l) = - (ALL l. length l = n --> (ALL x. P (x # l)))" +lemma LENGTH_EQ_CONS: "ALL (P::'a::type list => bool) n::nat. + (ALL l::'a::type list. length l = Suc n --> P l) = + (ALL l::'a::type list. length l = n --> (ALL x::'a::type. P (x # l)))" by (import list LENGTH_EQ_CONS) -lemma LENGTH_EQ_NIL: "ALL P. (ALL l. length l = 0 --> P l) = P []" +lemma LENGTH_EQ_NIL: "ALL P::'a::type list => bool. + (ALL l::'a::type list. length l = (0::nat) --> P l) = P []" by (import list LENGTH_EQ_NIL) -lemma CONS_ACYCLIC: "ALL l x. l ~= x # l & x # l ~= l" +lemma CONS_ACYCLIC: "ALL (l::'a::type list) x::'a::type. l ~= x # l & x # l ~= l" by (import list CONS_ACYCLIC) -lemma APPEND_eq_NIL: "(ALL (l1::'a list) l2::'a list. ([] = l1 @ l2) = (l1 = [] & l2 = [])) & -(ALL (l1::'a list) l2::'a list. (l1 @ l2 = []) = (l1 = [] & l2 = []))" +lemma APPEND_eq_NIL: "(ALL (l1::'a::type list) l2::'a::type list. + ([] = l1 @ l2) = (l1 = [] & l2 = [])) & +(ALL (l1::'a::type list) l2::'a::type list. + (l1 @ l2 = []) = (l1 = [] & l2 = []))" by (import list APPEND_eq_NIL) -lemma APPEND_11: "(ALL (l1::'a list) (l2::'a list) l3::'a list. +lemma APPEND_11: "(ALL (l1::'a::type list) (l2::'a::type list) l3::'a::type list. (l1 @ l2 = l1 @ l3) = (l2 = l3)) & -(ALL (l1::'a list) (l2::'a list) l3::'a list. +(ALL (l1::'a::type list) (l2::'a::type list) l3::'a::type list. (l2 @ l1 = l3 @ l1) = (l2 = l3))" by (import list APPEND_11) -lemma EL_compute: "ALL n. EL n l = (if n = 0 then hd l else EL (PRE n) (tl l))" +lemma EL_compute: "ALL n::nat. + EL n (l::'a::type list) = + (if n = (0::nat) then hd l else EL (PRE n) (tl l))" by (import list EL_compute) -lemma WF_LIST_PRED: "WF (%L1 L2. EX h. L2 = h # L1)" +lemma WF_LIST_PRED: "WF (%(L1::'a::type list) L2::'a::type list. EX h::'a::type. L2 = h # L1)" by (import list WF_LIST_PRED) -lemma list_size_cong: "ALL M N f f'. - M = N & (ALL x. x mem N --> f x = f' x) --> +lemma list_size_cong: "ALL (M::'a::type list) (N::'a::type list) (f::'a::type => nat) + f'::'a::type => nat. + M = N & (ALL x::'a::type. x mem N --> f x = f' x) --> list_size f M = list_size f' N" by (import list list_size_cong) -lemma FOLDR_CONG: "ALL l l' b b' f f'. - l = l' & b = b' & (ALL x a. x mem l' --> f x a = f' x a) --> +lemma FOLDR_CONG: "ALL (l::'a::type list) (l'::'a::type list) (b::'b::type) (b'::'b::type) + (f::'a::type => 'b::type => 'b::type) + f'::'a::type => 'b::type => 'b::type. + l = l' & + b = b' & (ALL (x::'a::type) a::'b::type. x mem l' --> f x a = f' x a) --> foldr f l b = foldr f' l' b'" by (import list FOLDR_CONG) -lemma FOLDL_CONG: "ALL l l' b b' f f'. - l = l' & b = b' & (ALL x a. x mem l' --> f a x = f' a x) --> +lemma FOLDL_CONG: "ALL (l::'a::type list) (l'::'a::type list) (b::'b::type) (b'::'b::type) + (f::'b::type => 'a::type => 'b::type) + f'::'b::type => 'a::type => 'b::type. + l = l' & + b = b' & (ALL (x::'a::type) a::'b::type. x mem l' --> f a x = f' a x) --> foldl f b l = foldl f' b' l'" by (import list FOLDL_CONG) -lemma MAP_CONG: "ALL l1 l2 f f'. - l1 = l2 & (ALL x. x mem l2 --> f x = f' x) --> map f l1 = map f' l2" +lemma MAP_CONG: "ALL (l1::'a::type list) (l2::'a::type list) (f::'a::type => 'b::type) + f'::'a::type => 'b::type. + l1 = l2 & (ALL x::'a::type. x mem l2 --> f x = f' x) --> + map f l1 = map f' l2" by (import list MAP_CONG) -lemma EXISTS_CONG: "ALL l1 l2 P P'. - l1 = l2 & (ALL x. x mem l2 --> P x = P' x) --> +lemma EXISTS_CONG: "ALL (l1::'a::type list) (l2::'a::type list) (P::'a::type => bool) + P'::'a::type => bool. + l1 = l2 & (ALL x::'a::type. x mem l2 --> P x = P' x) --> list_exists P l1 = list_exists P' l2" by (import list EXISTS_CONG) -lemma EVERY_CONG: "ALL l1 l2 P P'. - l1 = l2 & (ALL x. x mem l2 --> P x = P' x) --> +lemma EVERY_CONG: "ALL (l1::'a::type list) (l2::'a::type list) (P::'a::type => bool) + P'::'a::type => bool. + l1 = l2 & (ALL x::'a::type. x mem l2 --> P x = P' x) --> list_all P l1 = list_all P' l2" by (import list EVERY_CONG) -lemma EVERY_MONOTONIC: "ALL P Q. (ALL x. P x --> Q x) --> (ALL l. list_all P l --> list_all Q l)" +lemma EVERY_MONOTONIC: "ALL (P::'a::type => bool) Q::'a::type => bool. + (ALL x::'a::type. P x --> Q x) --> + (ALL l::'a::type list. list_all P l --> list_all Q l)" by (import list EVERY_MONOTONIC) -lemma LENGTH_ZIP: "ALL l1 l2. +lemma LENGTH_ZIP: "ALL (l1::'a::type list) l2::'b::type list. length l1 = length l2 --> length (zip l1 l2) = length l1 & length (zip l1 l2) = length l2" by (import list LENGTH_ZIP) -lemma LENGTH_UNZIP: "ALL pl. +lemma LENGTH_UNZIP: "ALL pl::('a::type * 'b::type) list. length (fst (unzip pl)) = length pl & length (snd (unzip pl)) = length pl" by (import list LENGTH_UNZIP) -lemma ZIP_UNZIP: "ALL l. ZIP (unzip l) = l" +lemma ZIP_UNZIP: "ALL l::('a::type * 'b::type) list. ZIP (unzip l) = l" by (import list ZIP_UNZIP) -lemma UNZIP_ZIP: "ALL l1 l2. length l1 = length l2 --> unzip (zip l1 l2) = (l1, l2)" +lemma UNZIP_ZIP: "ALL (l1::'a::type list) l2::'b::type list. + length l1 = length l2 --> unzip (zip l1 l2) = (l1, l2)" by (import list UNZIP_ZIP) -lemma ZIP_MAP: "ALL l1 l2 f1 f2. +lemma ZIP_MAP: "ALL (l1::'a::type list) (l2::'b::type list) (f1::'a::type => 'c::type) + f2::'b::type => 'd::type. length l1 = length l2 --> - zip (map f1 l1) l2 = map (%p. (f1 (fst p), snd p)) (zip l1 l2) & - zip l1 (map f2 l2) = map (%p. (fst p, f2 (snd p))) (zip l1 l2)" + zip (map f1 l1) l2 = + map (%p::'a::type * 'b::type. (f1 (fst p), snd p)) (zip l1 l2) & + zip l1 (map f2 l2) = + map (%p::'a::type * 'b::type. (fst p, f2 (snd p))) (zip l1 l2)" by (import list ZIP_MAP) -lemma MEM_ZIP: "(All::('a list => bool) => bool) - (%l1::'a list. - (All::('b list => bool) => bool) - (%l2::'b list. - (All::('a * 'b => bool) => bool) - (%p::'a * 'b. +lemma MEM_ZIP: "(All::('a::type list => bool) => bool) + (%l1::'a::type list. + (All::('b::type list => bool) => bool) + (%l2::'b::type list. + (All::('a::type * 'b::type => bool) => bool) + (%p::'a::type * 'b::type. (op -->::bool => bool => bool) - ((op =::nat => nat => bool) ((size::'a list => nat) l1) - ((size::'b list => nat) l2)) + ((op =::nat => nat => bool) + ((size::'a::type list => nat) l1) + ((size::'b::type list => nat) l2)) ((op =::bool => bool => bool) - ((op mem::'a * 'b => ('a * 'b) list => bool) p - ((zip::'a list => 'b list => ('a * 'b) list) l1 l2)) + ((op mem::'a::type * 'b::type + => ('a::type * 'b::type) list => bool) + p ((zip::'a::type list + => 'b::type list => ('a::type * 'b::type) list) + l1 l2)) ((Ex::(nat => bool) => bool) (%n::nat. (op &::bool => bool => bool) ((op <::nat => nat => bool) n - ((size::'a list => nat) l1)) - ((op =::'a * 'b => 'a * 'b => bool) p - ((Pair::'a => 'b => 'a * 'b) - ((EL::nat => 'a list => 'a) n l1) - ((EL::nat => 'b list => 'b) n l2)))))))))" + ((size::'a::type list => nat) l1)) + ((op =::'a::type * 'b::type + => 'a::type * 'b::type => bool) + p ((Pair::'a::type + => 'b::type => 'a::type * 'b::type) + ((EL::nat => 'a::type list => 'a::type) n l1) + ((EL::nat => 'b::type list => 'b::type) n + l2)))))))))" by (import list MEM_ZIP) -lemma EL_ZIP: "ALL l1 l2 n. +lemma EL_ZIP: "ALL (l1::'a::type list) (l2::'b::type list) n::nat. length l1 = length l2 & n < length l1 --> EL n (zip l1 l2) = (EL n l1, EL n l2)" by (import list EL_ZIP) -lemma MAP2_ZIP: "(All::('a list => bool) => bool) - (%l1::'a list. - (All::('b list => bool) => bool) - (%l2::'b list. +lemma MAP2_ZIP: "(All::('a::type list => bool) => bool) + (%l1::'a::type list. + (All::('b::type list => bool) => bool) + (%l2::'b::type list. (op -->::bool => bool => bool) - ((op =::nat => nat => bool) ((size::'a list => nat) l1) - ((size::'b list => nat) l2)) - ((All::(('a => 'b => 'c) => bool) => bool) - (%f::'a => 'b => 'c. - (op =::'c list => 'c list => bool) - ((map2::('a => 'b => 'c) => 'a list => 'b list => 'c list) + ((op =::nat => nat => bool) ((size::'a::type list => nat) l1) + ((size::'b::type list => nat) l2)) + ((All::(('a::type => 'b::type => 'c::type) => bool) => bool) + (%f::'a::type => 'b::type => 'c::type. + (op =::'c::type list => 'c::type list => bool) + ((map2::('a::type => 'b::type => 'c::type) + => 'a::type list + => 'b::type list => 'c::type list) f l1 l2) - ((map::('a * 'b => 'c) => ('a * 'b) list => 'c list) - ((split::('a => 'b => 'c) => 'a * 'b => 'c) f) - ((zip::'a list => 'b list => ('a * 'b) list) l1 l2))))))" + ((map::('a::type * 'b::type => 'c::type) + => ('a::type * 'b::type) list => 'c::type list) + ((split::('a::type => 'b::type => 'c::type) + => 'a::type * 'b::type => 'c::type) + f) + ((zip::'a::type list + => 'b::type list => ('a::type * 'b::type) list) + l1 l2))))))" by (import list MAP2_ZIP) -lemma MEM_EL: "(All::('a list => bool) => bool) - (%l::'a list. - (All::('a => bool) => bool) - (%x::'a. - (op =::bool => bool => bool) ((op mem::'a => 'a list => bool) x l) +lemma MEM_EL: "(All::('a::type list => bool) => bool) + (%l::'a::type list. + (All::('a::type => bool) => bool) + (%x::'a::type. + (op =::bool => bool => bool) + ((op mem::'a::type => 'a::type list => bool) x l) ((Ex::(nat => bool) => bool) (%n::nat. (op &::bool => bool => bool) - ((op <::nat => nat => bool) n ((size::'a list => nat) l)) - ((op =::'a => 'a => bool) x - ((EL::nat => 'a list => 'a) n l))))))" + ((op <::nat => nat => bool) n + ((size::'a::type list => nat) l)) + ((op =::'a::type => 'a::type => bool) x + ((EL::nat => 'a::type list => 'a::type) n l))))))" by (import list MEM_EL) -lemma LAST_CONS: "(ALL x::'a. last [x] = x) & -(ALL (x::'a) (xa::'a) xb::'a list. last (x # xa # xb) = last (xa # xb))" +lemma LAST_CONS: "(ALL x::'a::type. last [x] = x) & +(ALL (x::'a::type) (xa::'a::type) xb::'a::type list. + last (x # xa # xb) = last (xa # xb))" by (import list LAST_CONS) -lemma FRONT_CONS: "(ALL x::'a. butlast [x] = []) & -(ALL (x::'a) (xa::'a) xb::'a list. +lemma FRONT_CONS: "(ALL x::'a::type. butlast [x] = []) & +(ALL (x::'a::type) (xa::'a::type) xb::'a::type list. butlast (x # xa # xb) = x # butlast (xa # xb))" by (import list FRONT_CONS) @@ -2735,10 +3231,12 @@ ;setup_theory pred_set -lemma EXTENSION: "ALL s t. (s = t) = (ALL x. IN x s = IN x t)" +lemma EXTENSION: "ALL (s::'a::type => bool) t::'a::type => bool. + (s = t) = (ALL x::'a::type. IN x s = IN x t)" by (import pred_set EXTENSION) -lemma NOT_EQUAL_SETS: "ALL x xa. (x ~= xa) = (EX xb. IN xb xa = (~ IN xb x))" +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))" by (import pred_set NOT_EQUAL_SETS) lemma NUM_SET_WOP: "ALL s::nat => bool. @@ -2747,39 +3245,40 @@ by (import pred_set NUM_SET_WOP) consts - GSPEC :: "('b => 'a * bool) => 'a => bool" - -specification (GSPEC) GSPECIFICATION: "ALL (f::'b => 'a * bool) v::'a. IN v (GSPEC f) = (EX x::'b. (v, True) = f x)" + GSPEC :: "('b::type => 'a::type * bool) => 'a::type => bool" + +specification (GSPEC) GSPECIFICATION: "ALL (f::'b::type => 'a::type * bool) v::'a::type. + IN v (GSPEC f) = (EX x::'b::type. (v, True) = f x)" by (import pred_set GSPECIFICATION) -lemma SET_MINIMUM: "ALL (s::'a => bool) M::'a => nat. - (EX x::'a. IN x s) = - (EX x::'a. IN x s & (ALL y::'a. IN y s --> M x <= M y))" +lemma SET_MINIMUM: "ALL (s::'a::type => bool) M::'a::type => nat. + (EX x::'a::type. IN x s) = + (EX x::'a::type. IN x s & (ALL y::'a::type. IN y s --> M x <= M y))" by (import pred_set SET_MINIMUM) constdefs - EMPTY :: "'a => bool" - "EMPTY == %x. False" - -lemma EMPTY_DEF: "EMPTY = (%x. False)" + EMPTY :: "'a::type => bool" + "EMPTY == %x::'a::type. False" + +lemma EMPTY_DEF: "EMPTY = (%x::'a::type. False)" by (import pred_set EMPTY_DEF) -lemma NOT_IN_EMPTY: "ALL x. ~ IN x EMPTY" +lemma NOT_IN_EMPTY: "ALL x::'a::type. ~ IN x EMPTY" by (import pred_set NOT_IN_EMPTY) -lemma MEMBER_NOT_EMPTY: "ALL x. (EX xa. IN xa x) = (x ~= EMPTY)" +lemma MEMBER_NOT_EMPTY: "ALL x::'a::type => bool. (EX xa::'a::type. IN xa x) = (x ~= EMPTY)" by (import pred_set MEMBER_NOT_EMPTY) consts - UNIV :: "'a => bool" + UNIV :: "'a::type => bool" defs - UNIV_def: "pred_set.UNIV == %x. True" - -lemma UNIV_DEF: "pred_set.UNIV = (%x. True)" + UNIV_def: "pred_set.UNIV == %x::'a::type. True" + +lemma UNIV_DEF: "pred_set.UNIV = (%x::'a::type. True)" by (import pred_set UNIV_DEF) -lemma IN_UNIV: "ALL x. IN x pred_set.UNIV" +lemma IN_UNIV: "ALL x::'a::type. IN x pred_set.UNIV" by (import pred_set IN_UNIV) lemma UNIV_NOT_EMPTY: "pred_set.UNIV ~= EMPTY" @@ -2788,868 +3287,1025 @@ lemma EMPTY_NOT_UNIV: "EMPTY ~= pred_set.UNIV" by (import pred_set EMPTY_NOT_UNIV) -lemma EQ_UNIV: "(ALL x. IN x s) = (s = pred_set.UNIV)" +lemma EQ_UNIV: "(ALL x::'a::type. IN x (s::'a::type => bool)) = (s = pred_set.UNIV)" by (import pred_set EQ_UNIV) constdefs - SUBSET :: "('a => bool) => ('a => bool) => bool" - "SUBSET == %s t. ALL x. IN x s --> IN x t" - -lemma SUBSET_DEF: "ALL s t. SUBSET s t = (ALL x. IN x s --> IN x t)" + SUBSET :: "('a::type => bool) => ('a::type => bool) => bool" + "SUBSET == +%(s::'a::type => bool) t::'a::type => bool. + ALL x::'a::type. IN x s --> IN x t" + +lemma SUBSET_DEF: "ALL (s::'a::type => bool) t::'a::type => bool. + SUBSET s t = (ALL x::'a::type. IN x s --> IN x t)" by (import pred_set SUBSET_DEF) -lemma SUBSET_TRANS: "ALL x xa xb. SUBSET x xa & SUBSET xa xb --> SUBSET x xb" +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 pred_set SUBSET_TRANS) -lemma SUBSET_REFL: "ALL x. SUBSET x x" +lemma SUBSET_REFL: "ALL x::'a::type => bool. SUBSET x x" by (import pred_set SUBSET_REFL) -lemma SUBSET_ANTISYM: "ALL x xa. SUBSET x xa & SUBSET xa x --> x = xa" +lemma SUBSET_ANTISYM: "ALL (x::'a::type => bool) xa::'a::type => bool. + SUBSET x xa & SUBSET xa x --> x = xa" by (import pred_set SUBSET_ANTISYM) lemma EMPTY_SUBSET: "All (SUBSET EMPTY)" by (import pred_set EMPTY_SUBSET) -lemma SUBSET_EMPTY: "ALL x. SUBSET x EMPTY = (x = EMPTY)" +lemma SUBSET_EMPTY: "ALL x::'a::type => bool. SUBSET x EMPTY = (x = EMPTY)" by (import pred_set SUBSET_EMPTY) -lemma SUBSET_UNIV: "ALL x. SUBSET x pred_set.UNIV" +lemma SUBSET_UNIV: "ALL x::'a::type => bool. SUBSET x pred_set.UNIV" by (import pred_set SUBSET_UNIV) -lemma UNIV_SUBSET: "ALL x. SUBSET pred_set.UNIV x = (x = pred_set.UNIV)" +lemma UNIV_SUBSET: "ALL x::'a::type => bool. SUBSET pred_set.UNIV x = (x = pred_set.UNIV)" by (import pred_set UNIV_SUBSET) constdefs - PSUBSET :: "('a => bool) => ('a => bool) => bool" - "PSUBSET == %s t. SUBSET s t & s ~= t" - -lemma PSUBSET_DEF: "ALL s t. PSUBSET s t = (SUBSET s t & s ~= t)" + PSUBSET :: "('a::type => bool) => ('a::type => bool) => bool" + "PSUBSET == %(s::'a::type => bool) t::'a::type => bool. SUBSET s t & s ~= t" + +lemma PSUBSET_DEF: "ALL (s::'a::type => bool) t::'a::type => bool. + PSUBSET s t = (SUBSET s t & s ~= t)" by (import pred_set PSUBSET_DEF) -lemma PSUBSET_TRANS: "ALL x xa xb. PSUBSET x xa & PSUBSET xa xb --> PSUBSET x xb" +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 pred_set PSUBSET_TRANS) -lemma PSUBSET_IRREFL: "ALL x. ~ PSUBSET x x" +lemma PSUBSET_IRREFL: "ALL x::'a::type => bool. ~ PSUBSET x x" by (import pred_set PSUBSET_IRREFL) -lemma NOT_PSUBSET_EMPTY: "ALL x. ~ PSUBSET x EMPTY" +lemma NOT_PSUBSET_EMPTY: "ALL x::'a::type => bool. ~ PSUBSET x EMPTY" by (import pred_set NOT_PSUBSET_EMPTY) -lemma NOT_UNIV_PSUBSET: "ALL x. ~ PSUBSET pred_set.UNIV x" +lemma NOT_UNIV_PSUBSET: "ALL x::'a::type => bool. ~ PSUBSET pred_set.UNIV x" by (import pred_set NOT_UNIV_PSUBSET) -lemma PSUBSET_UNIV: "ALL x. PSUBSET x pred_set.UNIV = (EX xa. ~ IN xa x)" +lemma PSUBSET_UNIV: "ALL x::'a::type => bool. + PSUBSET x pred_set.UNIV = (EX xa::'a::type. ~ IN xa x)" by (import pred_set PSUBSET_UNIV) consts - UNION :: "('a => bool) => ('a => bool) => 'a => bool" + UNION :: "('a::type => bool) => ('a::type => bool) => 'a::type => bool" defs - UNION_def: "pred_set.UNION == %s t. GSPEC (%x. (x, IN x s | IN x t))" - -lemma UNION_DEF: "ALL s t. pred_set.UNION s t = GSPEC (%x. (x, IN x s | IN x t))" + UNION_def: "pred_set.UNION == +%(s::'a::type => bool) t::'a::type => bool. + GSPEC (%x::'a::type. (x, IN x s | IN x t))" + +lemma UNION_DEF: "ALL (s::'a::type => bool) t::'a::type => bool. + pred_set.UNION s t = GSPEC (%x::'a::type. (x, IN x s | IN x t))" by (import pred_set UNION_DEF) -lemma IN_UNION: "ALL x xa xb. IN xb (pred_set.UNION x xa) = (IN xb x | IN xb xa)" +lemma IN_UNION: "ALL (x::'a::type => bool) (xa::'a::type => bool) xb::'a::type. + IN xb (pred_set.UNION x xa) = (IN xb x | IN xb xa)" by (import pred_set IN_UNION) -lemma UNION_ASSOC: "ALL x xa xb. +lemma UNION_ASSOC: "ALL (x::'a::type => bool) (xa::'a::type => bool) xb::'a::type => bool. pred_set.UNION x (pred_set.UNION xa xb) = pred_set.UNION (pred_set.UNION x xa) xb" by (import pred_set UNION_ASSOC) -lemma UNION_IDEMPOT: "ALL x. pred_set.UNION x x = x" +lemma UNION_IDEMPOT: "ALL x::'a::type => bool. pred_set.UNION x x = x" by (import pred_set UNION_IDEMPOT) -lemma UNION_COMM: "ALL x xa. pred_set.UNION x xa = pred_set.UNION xa x" +lemma UNION_COMM: "ALL (x::'a::type => bool) xa::'a::type => bool. + pred_set.UNION x xa = pred_set.UNION xa x" by (import pred_set UNION_COMM) -lemma SUBSET_UNION: "(ALL (x::'a => bool) xa::'a => bool. SUBSET x (pred_set.UNION x xa)) & -(ALL (x::'a => bool) xa::'a => bool. SUBSET x (pred_set.UNION xa x))" +lemma SUBSET_UNION: "(ALL (x::'a::type => bool) xa::'a::type => bool. + SUBSET x (pred_set.UNION x xa)) & +(ALL (x::'a::type => bool) xa::'a::type => bool. + SUBSET x (pred_set.UNION xa x))" by (import pred_set SUBSET_UNION) -lemma UNION_SUBSET: "ALL s t u. SUBSET (pred_set.UNION s t) u = (SUBSET s u & SUBSET t u)" +lemma UNION_SUBSET: "ALL (s::'a::type => bool) (t::'a::type => bool) u::'a::type => bool. + SUBSET (pred_set.UNION s t) u = (SUBSET s u & SUBSET t u)" by (import pred_set UNION_SUBSET) -lemma SUBSET_UNION_ABSORPTION: "ALL x xa. SUBSET x xa = (pred_set.UNION x xa = xa)" +lemma SUBSET_UNION_ABSORPTION: "ALL (x::'a::type => bool) xa::'a::type => bool. + SUBSET x xa = (pred_set.UNION x xa = xa)" by (import pred_set SUBSET_UNION_ABSORPTION) -lemma UNION_EMPTY: "(ALL x::'a => bool. pred_set.UNION EMPTY x = x) & -(ALL x::'a => bool. pred_set.UNION x EMPTY = x)" +lemma UNION_EMPTY: "(ALL x::'a::type => bool. pred_set.UNION EMPTY x = x) & +(ALL x::'a::type => bool. pred_set.UNION x EMPTY = x)" by (import pred_set UNION_EMPTY) -lemma UNION_UNIV: "(ALL x::'a => bool. pred_set.UNION pred_set.UNIV x = pred_set.UNIV) & -(ALL x::'a => bool. pred_set.UNION x pred_set.UNIV = pred_set.UNIV)" +lemma UNION_UNIV: "(ALL x::'a::type => bool. pred_set.UNION pred_set.UNIV x = pred_set.UNIV) & +(ALL x::'a::type => bool. pred_set.UNION x pred_set.UNIV = pred_set.UNIV)" by (import pred_set UNION_UNIV) -lemma EMPTY_UNION: "ALL x xa. (pred_set.UNION x xa = EMPTY) = (x = EMPTY & xa = EMPTY)" +lemma EMPTY_UNION: "ALL (x::'a::type => bool) xa::'a::type => bool. + (pred_set.UNION x xa = EMPTY) = (x = EMPTY & xa = EMPTY)" by (import pred_set EMPTY_UNION) consts - INTER :: "('a => bool) => ('a => bool) => 'a => bool" + INTER :: "('a::type => bool) => ('a::type => bool) => 'a::type => bool" defs - INTER_def: "pred_set.INTER == %s t. GSPEC (%x. (x, IN x s & IN x t))" - -lemma INTER_DEF: "ALL s t. pred_set.INTER s t = GSPEC (%x. (x, IN x s & IN x t))" + INTER_def: "pred_set.INTER == +%(s::'a::type => bool) t::'a::type => bool. + GSPEC (%x::'a::type. (x, IN x s & IN x t))" + +lemma INTER_DEF: "ALL (s::'a::type => bool) t::'a::type => bool. + pred_set.INTER s t = GSPEC (%x::'a::type. (x, IN x s & IN x t))" by (import pred_set INTER_DEF) -lemma IN_INTER: "ALL x xa xb. IN xb (pred_set.INTER x xa) = (IN xb x & IN xb xa)" +lemma IN_INTER: "ALL (x::'a::type => bool) (xa::'a::type => bool) xb::'a::type. + IN xb (pred_set.INTER x xa) = (IN xb x & IN xb xa)" by (import pred_set IN_INTER) -lemma INTER_ASSOC: "ALL x xa xb. +lemma INTER_ASSOC: "ALL (x::'a::type => bool) (xa::'a::type => bool) xb::'a::type => bool. pred_set.INTER x (pred_set.INTER xa xb) = pred_set.INTER (pred_set.INTER x xa) xb" by (import pred_set INTER_ASSOC) -lemma INTER_IDEMPOT: "ALL x. pred_set.INTER x x = x" +lemma INTER_IDEMPOT: "ALL x::'a::type => bool. pred_set.INTER x x = x" by (import pred_set INTER_IDEMPOT) -lemma INTER_COMM: "ALL x xa. pred_set.INTER x xa = pred_set.INTER xa x" +lemma INTER_COMM: "ALL (x::'a::type => bool) xa::'a::type => bool. + pred_set.INTER x xa = pred_set.INTER xa x" by (import pred_set INTER_COMM) -lemma INTER_SUBSET: "(ALL (x::'a => bool) xa::'a => bool. SUBSET (pred_set.INTER x xa) x) & -(ALL (x::'a => bool) xa::'a => bool. SUBSET (pred_set.INTER xa x) x)" +lemma INTER_SUBSET: "(ALL (x::'a::type => bool) xa::'a::type => bool. + SUBSET (pred_set.INTER x xa) x) & +(ALL (x::'a::type => bool) xa::'a::type => bool. + SUBSET (pred_set.INTER xa x) x)" by (import pred_set INTER_SUBSET) -lemma SUBSET_INTER: "ALL s t u. SUBSET s (pred_set.INTER t u) = (SUBSET s t & SUBSET s u)" +lemma SUBSET_INTER: "ALL (s::'a::type => bool) (t::'a::type => bool) u::'a::type => bool. + SUBSET s (pred_set.INTER t u) = (SUBSET s t & SUBSET s u)" by (import pred_set SUBSET_INTER) -lemma SUBSET_INTER_ABSORPTION: "ALL x xa. SUBSET x xa = (pred_set.INTER x xa = x)" +lemma SUBSET_INTER_ABSORPTION: "ALL (x::'a::type => bool) xa::'a::type => bool. + SUBSET x xa = (pred_set.INTER x xa = x)" by (import pred_set SUBSET_INTER_ABSORPTION) -lemma INTER_EMPTY: "(ALL x::'a => bool. pred_set.INTER EMPTY x = EMPTY) & -(ALL x::'a => bool. pred_set.INTER x EMPTY = EMPTY)" +lemma INTER_EMPTY: "(ALL x::'a::type => bool. pred_set.INTER EMPTY x = EMPTY) & +(ALL x::'a::type => bool. pred_set.INTER x EMPTY = EMPTY)" by (import pred_set INTER_EMPTY) -lemma INTER_UNIV: "(ALL x::'a => bool. pred_set.INTER pred_set.UNIV x = x) & -(ALL x::'a => bool. pred_set.INTER x pred_set.UNIV = x)" +lemma INTER_UNIV: "(ALL x::'a::type => bool. pred_set.INTER pred_set.UNIV x = x) & +(ALL x::'a::type => bool. pred_set.INTER x pred_set.UNIV = x)" by (import pred_set INTER_UNIV) -lemma UNION_OVER_INTER: "ALL x xa xb. +lemma UNION_OVER_INTER: "ALL (x::'a::type => bool) (xa::'a::type => bool) xb::'a::type => bool. pred_set.INTER x (pred_set.UNION xa xb) = pred_set.UNION (pred_set.INTER x xa) (pred_set.INTER x xb)" by (import pred_set UNION_OVER_INTER) -lemma INTER_OVER_UNION: "ALL x xa xb. +lemma INTER_OVER_UNION: "ALL (x::'a::type => bool) (xa::'a::type => bool) xb::'a::type => bool. pred_set.UNION x (pred_set.INTER xa xb) = pred_set.INTER (pred_set.UNION x xa) (pred_set.UNION x xb)" by (import pred_set INTER_OVER_UNION) constdefs - DISJOINT :: "('a => bool) => ('a => bool) => bool" - "DISJOINT == %s t. pred_set.INTER s t = EMPTY" - -lemma DISJOINT_DEF: "ALL s t. DISJOINT s t = (pred_set.INTER s t = EMPTY)" + DISJOINT :: "('a::type => bool) => ('a::type => bool) => bool" + "DISJOINT == +%(s::'a::type => bool) t::'a::type => bool. pred_set.INTER s t = EMPTY" + +lemma DISJOINT_DEF: "ALL (s::'a::type => bool) t::'a::type => bool. + DISJOINT s t = (pred_set.INTER s t = EMPTY)" by (import pred_set DISJOINT_DEF) -lemma IN_DISJOINT: "ALL x xa. DISJOINT x xa = (~ (EX xb. IN xb x & IN xb xa))" +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))" by (import pred_set IN_DISJOINT) -lemma DISJOINT_SYM: "ALL x xa. DISJOINT x xa = DISJOINT xa x" +lemma DISJOINT_SYM: "ALL (x::'a::type => bool) xa::'a::type => bool. + DISJOINT x xa = DISJOINT xa x" by (import pred_set DISJOINT_SYM) -lemma DISJOINT_EMPTY: "ALL x. DISJOINT EMPTY x & DISJOINT x EMPTY" +lemma DISJOINT_EMPTY: "ALL x::'a::type => bool. DISJOINT EMPTY x & DISJOINT x EMPTY" by (import pred_set DISJOINT_EMPTY) -lemma DISJOINT_EMPTY_REFL: "ALL x. (x = EMPTY) = DISJOINT x x" +lemma DISJOINT_EMPTY_REFL: "ALL x::'a::type => bool. (x = EMPTY) = DISJOINT x x" by (import pred_set DISJOINT_EMPTY_REFL) -lemma DISJOINT_UNION: "ALL x xa xb. +lemma DISJOINT_UNION: "ALL (x::'a::type => bool) (xa::'a::type => bool) xb::'a::type => bool. DISJOINT (pred_set.UNION x xa) xb = (DISJOINT x xb & DISJOINT xa xb)" by (import pred_set DISJOINT_UNION) -lemma DISJOINT_UNION_BOTH: "ALL s t u. +lemma DISJOINT_UNION_BOTH: "ALL (s::'a::type => bool) (t::'a::type => bool) u::'a::type => bool. DISJOINT (pred_set.UNION s t) u = (DISJOINT s u & DISJOINT t u) & DISJOINT u (pred_set.UNION s t) = (DISJOINT s u & DISJOINT t u)" by (import pred_set DISJOINT_UNION_BOTH) constdefs - DIFF :: "('a => bool) => ('a => bool) => 'a => bool" - "DIFF == %s t. GSPEC (%x. (x, IN x s & ~ IN x t))" - -lemma DIFF_DEF: "ALL s t. DIFF s t = GSPEC (%x. (x, IN x s & ~ IN x t))" + DIFF :: "('a::type => bool) => ('a::type => bool) => 'a::type => bool" + "DIFF == +%(s::'a::type => bool) t::'a::type => bool. + GSPEC (%x::'a::type. (x, IN x s & ~ IN x t))" + +lemma DIFF_DEF: "ALL (s::'a::type => bool) t::'a::type => bool. + DIFF s t = GSPEC (%x::'a::type. (x, IN x s & ~ IN x t))" by (import pred_set DIFF_DEF) -lemma IN_DIFF: "ALL s t x. IN x (DIFF s t) = (IN x s & ~ IN x t)" +lemma IN_DIFF: "ALL (s::'a::type => bool) (t::'a::type => bool) x::'a::type. + IN x (DIFF s t) = (IN x s & ~ IN x t)" by (import pred_set IN_DIFF) -lemma DIFF_EMPTY: "ALL s. DIFF s EMPTY = s" +lemma DIFF_EMPTY: "ALL s::'a::type => bool. DIFF s EMPTY = s" by (import pred_set DIFF_EMPTY) -lemma EMPTY_DIFF: "ALL s. DIFF EMPTY s = EMPTY" +lemma EMPTY_DIFF: "ALL s::'a::type => bool. DIFF EMPTY s = EMPTY" by (import pred_set EMPTY_DIFF) -lemma DIFF_UNIV: "ALL s. DIFF s pred_set.UNIV = EMPTY" +lemma DIFF_UNIV: "ALL s::'a::type => bool. DIFF s pred_set.UNIV = EMPTY" by (import pred_set DIFF_UNIV) -lemma DIFF_DIFF: "ALL x xa. DIFF (DIFF x xa) xa = DIFF x xa" +lemma DIFF_DIFF: "ALL (x::'a::type => bool) xa::'a::type => bool. + DIFF (DIFF x xa) xa = DIFF x xa" by (import pred_set DIFF_DIFF) -lemma DIFF_EQ_EMPTY: "ALL x. DIFF x x = EMPTY" +lemma DIFF_EQ_EMPTY: "ALL x::'a::type => bool. DIFF x x = EMPTY" by (import pred_set DIFF_EQ_EMPTY) constdefs - INSERT :: "'a => ('a => bool) => 'a => bool" - "INSERT == %x s. GSPEC (%y. (y, y = x | IN y s))" - -lemma INSERT_DEF: "ALL x s. INSERT x s = GSPEC (%y. (y, y = x | IN y s))" + INSERT :: "'a::type => ('a::type => bool) => 'a::type => bool" + "INSERT == +%(x::'a::type) s::'a::type => bool. + GSPEC (%y::'a::type. (y, y = x | IN y s))" + +lemma INSERT_DEF: "ALL (x::'a::type) s::'a::type => bool. + INSERT x s = GSPEC (%y::'a::type. (y, y = x | IN y s))" by (import pred_set INSERT_DEF) -lemma IN_INSERT: "ALL x xa xb. IN x (INSERT xa xb) = (x = xa | IN x xb)" +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 pred_set IN_INSERT) -lemma COMPONENT: "ALL x xa. IN x (INSERT x xa)" +lemma COMPONENT: "ALL (x::'a::type) xa::'a::type => bool. IN x (INSERT x xa)" by (import pred_set COMPONENT) -lemma SET_CASES: "ALL x. x = EMPTY | (EX xa xb. x = INSERT xa xb & ~ IN xa xb)" +lemma SET_CASES: "ALL x::'a::type => bool. + x = EMPTY | + (EX (xa::'a::type) xb::'a::type => bool. x = INSERT xa xb & ~ IN xa xb)" by (import pred_set SET_CASES) -lemma DECOMPOSITION: "ALL s x. IN x s = (EX t. s = INSERT x t & ~ IN x t)" +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)" by (import pred_set DECOMPOSITION) -lemma ABSORPTION: "ALL x xa. IN x xa = (INSERT x xa = xa)" +lemma ABSORPTION: "ALL (x::'a::type) xa::'a::type => bool. IN x xa = (INSERT x xa = xa)" by (import pred_set ABSORPTION) -lemma INSERT_INSERT: "ALL x xa. INSERT x (INSERT x xa) = INSERT x xa" +lemma INSERT_INSERT: "ALL (x::'a::type) xa::'a::type => bool. INSERT x (INSERT x xa) = INSERT x xa" by (import pred_set INSERT_INSERT) -lemma INSERT_COMM: "ALL x xa xb. INSERT x (INSERT xa xb) = INSERT xa (INSERT x xb)" +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 pred_set INSERT_COMM) -lemma INSERT_UNIV: "ALL x. INSERT x pred_set.UNIV = pred_set.UNIV" +lemma INSERT_UNIV: "ALL x::'a::type. INSERT x pred_set.UNIV = pred_set.UNIV" by (import pred_set INSERT_UNIV) -lemma NOT_INSERT_EMPTY: "ALL x xa. INSERT x xa ~= EMPTY" +lemma NOT_INSERT_EMPTY: "ALL (x::'a::type) xa::'a::type => bool. INSERT x xa ~= EMPTY" by (import pred_set NOT_INSERT_EMPTY) -lemma NOT_EMPTY_INSERT: "ALL x xa. EMPTY ~= INSERT x xa" +lemma NOT_EMPTY_INSERT: "ALL (x::'a::type) xa::'a::type => bool. EMPTY ~= INSERT x xa" by (import pred_set NOT_EMPTY_INSERT) -lemma INSERT_UNION: "ALL x s t. +lemma INSERT_UNION: "ALL (x::'a::type) (s::'a::type => bool) t::'a::type => bool. pred_set.UNION (INSERT x s) t = (if IN x t then pred_set.UNION s t else INSERT x (pred_set.UNION s t))" by (import pred_set INSERT_UNION) -lemma INSERT_UNION_EQ: "ALL x s t. pred_set.UNION (INSERT x s) t = INSERT x (pred_set.UNION s t)" +lemma INSERT_UNION_EQ: "ALL (x::'a::type) (s::'a::type => bool) t::'a::type => bool. + pred_set.UNION (INSERT x s) t = INSERT x (pred_set.UNION s t)" by (import pred_set INSERT_UNION_EQ) -lemma INSERT_INTER: "ALL x s t. +lemma INSERT_INTER: "ALL (x::'a::type) (s::'a::type => bool) t::'a::type => bool. pred_set.INTER (INSERT x s) t = (if IN x t then INSERT x (pred_set.INTER s t) else pred_set.INTER s t)" by (import pred_set INSERT_INTER) -lemma DISJOINT_INSERT: "ALL x xa xb. DISJOINT (INSERT x xa) xb = (DISJOINT xa xb & ~ IN x xb)" +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)" by (import pred_set DISJOINT_INSERT) -lemma INSERT_SUBSET: "ALL x xa xb. SUBSET (INSERT x xa) xb = (IN x xb & SUBSET xa xb)" +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 pred_set INSERT_SUBSET) -lemma SUBSET_INSERT: "ALL x xa. ~ IN x xa --> (ALL xb. SUBSET xa (INSERT x xb) = SUBSET xa xb)" +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 pred_set SUBSET_INSERT) -lemma INSERT_DIFF: "ALL s t x. +lemma INSERT_DIFF: "ALL (s::'a::type => bool) (t::'a::type => bool) x::'a::type. DIFF (INSERT x s) t = (if IN x t then DIFF s t else INSERT x (DIFF s t))" by (import pred_set INSERT_DIFF) constdefs - DELETE :: "('a => bool) => 'a => 'a => bool" - "DELETE == %s x. DIFF s (INSERT x EMPTY)" - -lemma DELETE_DEF: "ALL s x. DELETE s x = DIFF s (INSERT x EMPTY)" + DELETE :: "('a::type => bool) => 'a::type => 'a::type => bool" + "DELETE == %(s::'a::type => bool) x::'a::type. DIFF s (INSERT x EMPTY)" + +lemma DELETE_DEF: "ALL (s::'a::type => bool) x::'a::type. DELETE s x = DIFF s (INSERT x EMPTY)" by (import pred_set DELETE_DEF) -lemma IN_DELETE: "ALL x xa xb. IN xa (DELETE x xb) = (IN xa x & xa ~= xb)" +lemma IN_DELETE: "ALL (x::'a::type => bool) (xa::'a::type) xb::'a::type. + IN xa (DELETE x xb) = (IN xa x & xa ~= xb)" by (import pred_set IN_DELETE) -lemma DELETE_NON_ELEMENT: "ALL x xa. (~ IN x xa) = (DELETE xa x = xa)" +lemma DELETE_NON_ELEMENT: "ALL (x::'a::type) xa::'a::type => bool. (~ IN x xa) = (DELETE xa x = xa)" by (import pred_set DELETE_NON_ELEMENT) -lemma IN_DELETE_EQ: "ALL s x x'. (IN x s = IN x' s) = (IN x (DELETE s x') = IN x' (DELETE s x))" +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))" by (import pred_set IN_DELETE_EQ) -lemma EMPTY_DELETE: "ALL x. DELETE EMPTY x = EMPTY" +lemma EMPTY_DELETE: "ALL x::'a::type. DELETE EMPTY x = EMPTY" by (import pred_set EMPTY_DELETE) -lemma DELETE_DELETE: "ALL x xa. DELETE (DELETE xa x) x = DELETE xa x" +lemma DELETE_DELETE: "ALL (x::'a::type) xa::'a::type => bool. DELETE (DELETE xa x) x = DELETE xa x" by (import pred_set DELETE_DELETE) -lemma DELETE_COMM: "ALL x xa xb. DELETE (DELETE xb x) xa = DELETE (DELETE xb xa) x" +lemma DELETE_COMM: "ALL (x::'a::type) (xa::'a::type) xb::'a::type => bool. + DELETE (DELETE xb x) xa = DELETE (DELETE xb xa) x" by (import pred_set DELETE_COMM) -lemma DELETE_SUBSET: "ALL x xa. SUBSET (DELETE xa x) xa" +lemma DELETE_SUBSET: "ALL (x::'a::type) xa::'a::type => bool. SUBSET (DELETE xa x) xa" by (import pred_set DELETE_SUBSET) -lemma SUBSET_DELETE: "ALL x xa xb. SUBSET xa (DELETE xb x) = (~ IN x xa & SUBSET xa xb)" +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)" by (import pred_set SUBSET_DELETE) -lemma SUBSET_INSERT_DELETE: "ALL x s t. SUBSET s (INSERT x t) = SUBSET (DELETE s x) t" +lemma SUBSET_INSERT_DELETE: "ALL (x::'a::type) (s::'a::type => bool) t::'a::type => bool. + SUBSET s (INSERT x t) = SUBSET (DELETE s x) t" by (import pred_set SUBSET_INSERT_DELETE) -lemma DIFF_INSERT: "ALL x xa xb. DIFF x (INSERT xb xa) = DIFF (DELETE x xb) xa" +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 pred_set DIFF_INSERT) -lemma PSUBSET_INSERT_SUBSET: "ALL x xa. PSUBSET x xa = (EX xb. ~ IN xb x & SUBSET (INSERT xb x) xa)" +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)" by (import pred_set PSUBSET_INSERT_SUBSET) -lemma PSUBSET_MEMBER: "ALL s t. PSUBSET s t = (SUBSET s t & (EX y. IN y t & ~ IN y s))" +lemma PSUBSET_MEMBER: "ALL (s::'a::type => bool) t::'a::type => bool. + PSUBSET s t = (SUBSET s t & (EX y::'a::type. IN y t & ~ IN y s))" by (import pred_set PSUBSET_MEMBER) -lemma DELETE_INSERT: "ALL x xa xb. +lemma DELETE_INSERT: "ALL (x::'a::type) (xa::'a::type) xb::'a::type => bool. DELETE (INSERT x xb) xa = (if x = xa then DELETE xb xa else INSERT x (DELETE xb xa))" by (import pred_set DELETE_INSERT) -lemma INSERT_DELETE: "ALL x xa. IN x xa --> INSERT x (DELETE xa x) = xa" +lemma INSERT_DELETE: "ALL (x::'a::type) xa::'a::type => bool. + IN x xa --> INSERT x (DELETE xa x) = xa" by (import pred_set INSERT_DELETE) -lemma DELETE_INTER: "ALL x xa xb. +lemma DELETE_INTER: "ALL (x::'a::type => bool) (xa::'a::type => bool) xb::'a::type. pred_set.INTER (DELETE x xb) xa = DELETE (pred_set.INTER x xa) xb" by (import pred_set DELETE_INTER) -lemma DISJOINT_DELETE_SYM: "ALL x xa xb. DISJOINT (DELETE x xb) xa = DISJOINT (DELETE xa xb) x" +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" by (import pred_set DISJOINT_DELETE_SYM) consts - CHOICE :: "('a => bool) => 'a" - -specification (CHOICE) CHOICE_DEF: "ALL x. x ~= EMPTY --> IN (CHOICE x) x" + CHOICE :: "('a::type => bool) => 'a::type" + +specification (CHOICE) CHOICE_DEF: "ALL x::'a::type => bool. x ~= EMPTY --> IN (CHOICE x) x" by (import pred_set CHOICE_DEF) constdefs - REST :: "('a => bool) => 'a => bool" - "REST == %s. DELETE s (CHOICE s)" - -lemma REST_DEF: "ALL s. REST s = DELETE s (CHOICE s)" + REST :: "('a::type => bool) => 'a::type => bool" + "REST == %s::'a::type => bool. DELETE s (CHOICE s)" + +lemma REST_DEF: "ALL s::'a::type => bool. REST s = DELETE s (CHOICE s)" by (import pred_set REST_DEF) -lemma CHOICE_NOT_IN_REST: "ALL x. ~ IN (CHOICE x) (REST x)" +lemma CHOICE_NOT_IN_REST: "ALL x::'a::type => bool. ~ IN (CHOICE x) (REST x)" by (import pred_set CHOICE_NOT_IN_REST) -lemma CHOICE_INSERT_REST: "ALL s. s ~= EMPTY --> INSERT (CHOICE s) (REST s) = s" +lemma CHOICE_INSERT_REST: "ALL s::'a::type => bool. s ~= EMPTY --> INSERT (CHOICE s) (REST s) = s" by (import pred_set CHOICE_INSERT_REST) -lemma REST_SUBSET: "ALL x. SUBSET (REST x) x" +lemma REST_SUBSET: "ALL x::'a::type => bool. SUBSET (REST x) x" by (import pred_set REST_SUBSET) -lemma REST_PSUBSET: "ALL x. x ~= EMPTY --> PSUBSET (REST x) x" +lemma REST_PSUBSET: "ALL x::'a::type => bool. x ~= EMPTY --> PSUBSET (REST x) x" by (import pred_set REST_PSUBSET) constdefs - SING :: "('a => bool) => bool" - "SING == %s. EX x. s = INSERT x EMPTY" - -lemma SING_DEF: "ALL s. SING s = (EX x. s = INSERT x EMPTY)" + SING :: "('a::type => bool) => bool" + "SING == %s::'a::type => bool. EX x::'a::type. s = INSERT x EMPTY" + +lemma SING_DEF: "ALL s::'a::type => bool. SING s = (EX x::'a::type. s = INSERT x EMPTY)" by (import pred_set SING_DEF) -lemma SING: "ALL x. SING (INSERT x EMPTY)" +lemma SING: "ALL x::'a::type. SING (INSERT x EMPTY)" by (import pred_set SING) -lemma IN_SING: "ALL x xa. IN x (INSERT xa EMPTY) = (x = xa)" +lemma IN_SING: "ALL (x::'a::type) xa::'a::type. IN x (INSERT xa EMPTY) = (x = xa)" by (import pred_set IN_SING) -lemma NOT_SING_EMPTY: "ALL x. INSERT x EMPTY ~= EMPTY" +lemma NOT_SING_EMPTY: "ALL x::'a::type. INSERT x EMPTY ~= EMPTY" by (import pred_set NOT_SING_EMPTY) -lemma NOT_EMPTY_SING: "ALL x. EMPTY ~= INSERT x EMPTY" +lemma NOT_EMPTY_SING: "ALL x::'a::type. EMPTY ~= INSERT x EMPTY" by (import pred_set NOT_EMPTY_SING) -lemma EQUAL_SING: "ALL x xa. (INSERT x EMPTY = INSERT xa EMPTY) = (x = xa)" +lemma EQUAL_SING: "ALL (x::'a::type) xa::'a::type. + (INSERT x EMPTY = INSERT xa EMPTY) = (x = xa)" by (import pred_set EQUAL_SING) -lemma DISJOINT_SING_EMPTY: "ALL x. DISJOINT (INSERT x EMPTY) EMPTY" +lemma DISJOINT_SING_EMPTY: "ALL x::'a::type. DISJOINT (INSERT x EMPTY) EMPTY" by (import pred_set DISJOINT_SING_EMPTY) -lemma INSERT_SING_UNION: "ALL x xa. INSERT xa x = pred_set.UNION (INSERT xa EMPTY) x" +lemma INSERT_SING_UNION: "ALL (x::'a::type => bool) xa::'a::type. + INSERT xa x = pred_set.UNION (INSERT xa EMPTY) x" by (import pred_set INSERT_SING_UNION) -lemma SING_DELETE: "ALL x. DELETE (INSERT x EMPTY) x = EMPTY" +lemma SING_DELETE: "ALL x::'a::type. DELETE (INSERT x EMPTY) x = EMPTY" by (import pred_set SING_DELETE) -lemma DELETE_EQ_SING: "ALL x xa. IN xa x --> (DELETE x xa = EMPTY) = (x = INSERT xa EMPTY)" +lemma DELETE_EQ_SING: "ALL (x::'a::type => bool) xa::'a::type. + IN xa x --> (DELETE x xa = EMPTY) = (x = INSERT xa EMPTY)" by (import pred_set DELETE_EQ_SING) -lemma CHOICE_SING: "ALL x. CHOICE (INSERT x EMPTY) = x" +lemma CHOICE_SING: "ALL x::'a::type. CHOICE (INSERT x EMPTY) = x" by (import pred_set CHOICE_SING) -lemma REST_SING: "ALL x. REST (INSERT x EMPTY) = EMPTY" +lemma REST_SING: "ALL x::'a::type. REST (INSERT x EMPTY) = EMPTY" by (import pred_set REST_SING) -lemma SING_IFF_EMPTY_REST: "ALL x. SING x = (x ~= EMPTY & REST x = EMPTY)" +lemma SING_IFF_EMPTY_REST: "ALL x::'a::type => bool. SING x = (x ~= EMPTY & REST x = EMPTY)" by (import pred_set SING_IFF_EMPTY_REST) constdefs - IMAGE :: "('a => 'b) => ('a => bool) => 'b => bool" - "IMAGE == %f s. GSPEC (%x. (f x, IN x s))" - -lemma IMAGE_DEF: "ALL f s. IMAGE f s = GSPEC (%x. (f x, IN x s))" + IMAGE :: "('a::type => 'b::type) => ('a::type => bool) => 'b::type => bool" + "IMAGE == +%(f::'a::type => 'b::type) s::'a::type => bool. + GSPEC (%x::'a::type. (f x, IN x s))" + +lemma IMAGE_DEF: "ALL (f::'a::type => 'b::type) s::'a::type => bool. + IMAGE f s = GSPEC (%x::'a::type. (f x, IN x s))" by (import pred_set IMAGE_DEF) -lemma IN_IMAGE: "ALL (x::'b) (xa::'a => bool) xb::'a => 'b. - IN x (IMAGE xb xa) = (EX xc::'a. x = xb xc & IN xc xa)" +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)" by (import pred_set IN_IMAGE) -lemma IMAGE_IN: "ALL x xa. IN x xa --> (ALL xb. IN (xb x) (IMAGE xb xa))" +lemma IMAGE_IN: "ALL (x::'a::type) xa::'a::type => bool. + IN x xa --> (ALL xb::'a::type => 'b::type. IN (xb x) (IMAGE xb xa))" by (import pred_set IMAGE_IN) -lemma IMAGE_EMPTY: "ALL x. IMAGE x EMPTY = EMPTY" +lemma IMAGE_EMPTY: "ALL x::'a::type => 'b::type. IMAGE x EMPTY = EMPTY" by (import pred_set IMAGE_EMPTY) -lemma IMAGE_ID: "ALL x. IMAGE (%x. x) x = x" +lemma IMAGE_ID: "ALL x::'a::type => bool. IMAGE (%x::'a::type. x) x = x" by (import pred_set IMAGE_ID) -lemma IMAGE_COMPOSE: "ALL (x::'b => 'c) (xa::'a => 'b) xb::'a => bool. - IMAGE (x o xa) xb = IMAGE x (IMAGE xa xb)" +lemma IMAGE_COMPOSE: "ALL (x::'b::type => 'c::type) (xa::'a::type => 'b::type) + xb::'a::type => bool. IMAGE (x o xa) xb = IMAGE x (IMAGE xa xb)" by (import pred_set IMAGE_COMPOSE) -lemma IMAGE_INSERT: "ALL x xa xb. IMAGE x (INSERT xa xb) = INSERT (x xa) (IMAGE x xb)" +lemma IMAGE_INSERT: "ALL (x::'a::type => 'b::type) (xa::'a::type) xb::'a::type => bool. + IMAGE x (INSERT xa xb) = INSERT (x xa) (IMAGE x xb)" by (import pred_set IMAGE_INSERT) -lemma IMAGE_EQ_EMPTY: "ALL s x. (IMAGE x s = EMPTY) = (s = EMPTY)" +lemma IMAGE_EQ_EMPTY: "ALL (s::'a::type => bool) x::'a::type => 'b::type. + (IMAGE x s = EMPTY) = (s = EMPTY)" by (import pred_set IMAGE_EQ_EMPTY) -lemma IMAGE_DELETE: "ALL f x s. ~ IN x s --> IMAGE f (DELETE s x) = IMAGE f s" +lemma IMAGE_DELETE: "ALL (f::'a::type => 'b::type) (x::'a::type) s::'a::type => bool. + ~ IN x s --> IMAGE f (DELETE s x) = IMAGE f s" by (import pred_set IMAGE_DELETE) -lemma IMAGE_UNION: "ALL x xa xb. +lemma IMAGE_UNION: "ALL (x::'a::type => 'b::type) (xa::'a::type => bool) xb::'a::type => bool. IMAGE x (pred_set.UNION xa xb) = pred_set.UNION (IMAGE x xa) (IMAGE x xb)" by (import pred_set IMAGE_UNION) -lemma IMAGE_SUBSET: "ALL x xa. SUBSET x xa --> (ALL xb. SUBSET (IMAGE xb x) (IMAGE xb xa))" +lemma IMAGE_SUBSET: "ALL (x::'a::type => bool) xa::'a::type => bool. + SUBSET x xa --> + (ALL xb::'a::type => 'b::type. SUBSET (IMAGE xb x) (IMAGE xb xa))" by (import pred_set IMAGE_SUBSET) -lemma IMAGE_INTER: "ALL f s t. +lemma IMAGE_INTER: "ALL (f::'a::type => 'b::type) (s::'a::type => bool) t::'a::type => bool. SUBSET (IMAGE f (pred_set.INTER s t)) (pred_set.INTER (IMAGE f s) (IMAGE f t))" by (import pred_set IMAGE_INTER) constdefs - INJ :: "('a => 'b) => ('a => bool) => ('b => bool) => bool" + INJ :: "('a::type => 'b::type) => ('a::type => bool) => ('b::type => bool) => bool" "INJ == -%f s t. - (ALL x. IN x s --> IN (f x) t) & - (ALL x y. IN x s & IN y s --> f x = f y --> x = y)" - -lemma INJ_DEF: "ALL f s t. +%(f::'a::type => 'b::type) (s::'a::type => bool) t::'b::type => bool. + (ALL x::'a::type. IN x s --> IN (f x) t) & + (ALL (x::'a::type) y::'a::type. IN x s & IN y s --> f x = f y --> x = y)" + +lemma INJ_DEF: "ALL (f::'a::type => 'b::type) (s::'a::type => bool) t::'b::type => bool. INJ f s t = - ((ALL x. IN x s --> IN (f x) t) & - (ALL x y. IN x s & IN y s --> f x = f y --> x = y))" + ((ALL x::'a::type. IN x s --> IN (f x) t) & + (ALL (x::'a::type) y::'a::type. + IN x s & IN y s --> f x = f y --> x = y))" by (import pred_set INJ_DEF) -lemma INJ_ID: "ALL x. INJ (%x. x) x x" +lemma INJ_ID: "ALL x::'a::type => bool. INJ (%x::'a::type. x) x x" by (import pred_set INJ_ID) -lemma INJ_COMPOSE: "ALL x xa xb xc xd. INJ x xb xc & INJ xa xc xd --> INJ (xa o x) xb xd" +lemma INJ_COMPOSE: "ALL (x::'a::type => 'b::type) (xa::'b::type => 'c::type) + (xb::'a::type => bool) (xc::'b::type => bool) xd::'c::type => bool. + INJ x xb xc & INJ xa xc xd --> INJ (xa o x) xb xd" by (import pred_set INJ_COMPOSE) -lemma INJ_EMPTY: "ALL x. All (INJ x EMPTY) & (ALL xa. INJ x xa EMPTY = (xa = EMPTY))" +lemma INJ_EMPTY: "ALL x::'a::type => 'b::type. + All (INJ x EMPTY) & + (ALL xa::'a::type => bool. INJ x xa EMPTY = (xa = EMPTY))" by (import pred_set INJ_EMPTY) constdefs - SURJ :: "('a => 'b) => ('a => bool) => ('b => bool) => bool" + SURJ :: "('a::type => 'b::type) => ('a::type => bool) => ('b::type => bool) => bool" "SURJ == -%f s t. - (ALL x. IN x s --> IN (f x) t) & - (ALL x. IN x t --> (EX y. IN y s & f y = x))" - -lemma SURJ_DEF: "ALL f s t. +%(f::'a::type => 'b::type) (s::'a::type => bool) t::'b::type => bool. + (ALL x::'a::type. IN x s --> IN (f x) t) & + (ALL x::'b::type. IN x t --> (EX y::'a::type. IN y s & f y = x))" + +lemma SURJ_DEF: "ALL (f::'a::type => 'b::type) (s::'a::type => bool) t::'b::type => bool. SURJ f s t = - ((ALL x. IN x s --> IN (f x) t) & - (ALL x. IN x t --> (EX y. IN y s & f y = x)))" + ((ALL x::'a::type. IN x s --> IN (f x) t) & + (ALL x::'b::type. IN x t --> (EX y::'a::type. IN y s & f y = x)))" by (import pred_set SURJ_DEF) -lemma SURJ_ID: "ALL x. SURJ (%x. x) x x" +lemma SURJ_ID: "ALL x::'a::type => bool. SURJ (%x::'a::type. x) x x" by (import pred_set SURJ_ID) -lemma SURJ_COMPOSE: "ALL x xa xb xc xd. SURJ x xb xc & SURJ xa xc xd --> SURJ (xa o x) xb xd" +lemma SURJ_COMPOSE: "ALL (x::'a::type => 'b::type) (xa::'b::type => 'c::type) + (xb::'a::type => bool) (xc::'b::type => bool) xd::'c::type => bool. + SURJ x xb xc & SURJ xa xc xd --> SURJ (xa o x) xb xd" by (import pred_set SURJ_COMPOSE) -lemma SURJ_EMPTY: "ALL x. - (ALL xa. SURJ x EMPTY xa = (xa = EMPTY)) & - (ALL xa. SURJ x xa EMPTY = (xa = EMPTY))" +lemma SURJ_EMPTY: "ALL x::'a::type => 'b::type. + (ALL xa::'b::type => bool. SURJ x EMPTY xa = (xa = EMPTY)) & + (ALL xa::'a::type => bool. SURJ x xa EMPTY = (xa = EMPTY))" by (import pred_set SURJ_EMPTY) -lemma IMAGE_SURJ: "ALL x xa xb. SURJ x xa xb = (IMAGE x xa = xb)" +lemma IMAGE_SURJ: "ALL (x::'a::type => 'b::type) (xa::'a::type => bool) xb::'b::type => bool. + SURJ x xa xb = (IMAGE x xa = xb)" by (import pred_set IMAGE_SURJ) constdefs - BIJ :: "('a => 'b) => ('a => bool) => ('b => bool) => bool" - "BIJ == %f s t. INJ f s t & SURJ f s t" - -lemma BIJ_DEF: "ALL f s t. BIJ f s t = (INJ f s t & SURJ f s t)" + BIJ :: "('a::type => 'b::type) => ('a::type => bool) => ('b::type => bool) => bool" + "BIJ == +%(f::'a::type => 'b::type) (s::'a::type => bool) t::'b::type => bool. + INJ f s t & SURJ f s t" + +lemma BIJ_DEF: "ALL (f::'a::type => 'b::type) (s::'a::type => bool) t::'b::type => bool. + BIJ f s t = (INJ f s t & SURJ f s t)" by (import pred_set BIJ_DEF) -lemma BIJ_ID: "ALL x. BIJ (%x. x) x x" +lemma BIJ_ID: "ALL x::'a::type => bool. BIJ (%x::'a::type. x) x x" by (import pred_set BIJ_ID) -lemma BIJ_EMPTY: "ALL x. - (ALL xa. BIJ x EMPTY xa = (xa = EMPTY)) & - (ALL xa. BIJ x xa EMPTY = (xa = EMPTY))" +lemma BIJ_EMPTY: "ALL x::'a::type => 'b::type. + (ALL xa::'b::type => bool. BIJ x EMPTY xa = (xa = EMPTY)) & + (ALL xa::'a::type => bool. BIJ x xa EMPTY = (xa = EMPTY))" by (import pred_set BIJ_EMPTY) -lemma BIJ_COMPOSE: "ALL x xa xb xc xd. BIJ x xb xc & BIJ xa xc xd --> BIJ (xa o x) xb xd" +lemma BIJ_COMPOSE: "ALL (x::'a::type => 'b::type) (xa::'b::type => 'c::type) + (xb::'a::type => bool) (xc::'b::type => bool) xd::'c::type => bool. + BIJ x xb xc & BIJ xa xc xd --> BIJ (xa o x) xb xd" by (import pred_set BIJ_COMPOSE) consts - LINV :: "('a => 'b) => ('a => bool) => 'b => 'a" - -specification (LINV) LINV_DEF: "ALL f s t. INJ f s t --> (ALL x. IN x s --> LINV f s (f x) = x)" + LINV :: "('a::type => 'b::type) => ('a::type => bool) => 'b::type => 'a::type" + +specification (LINV) LINV_DEF: "ALL (f::'a::type => 'b::type) (s::'a::type => bool) t::'b::type => bool. + INJ f s t --> (ALL x::'a::type. IN x s --> LINV f s (f x) = x)" by (import pred_set LINV_DEF) consts - RINV :: "('a => 'b) => ('a => bool) => 'b => 'a" - -specification (RINV) RINV_DEF: "ALL f s t. SURJ f s t --> (ALL x. IN x t --> f (RINV f s x) = x)" + RINV :: "('a::type => 'b::type) => ('a::type => bool) => 'b::type => 'a::type" + +specification (RINV) RINV_DEF: "ALL (f::'a::type => 'b::type) (s::'a::type => bool) t::'b::type => bool. + SURJ f s t --> (ALL x::'b::type. IN x t --> f (RINV f s x) = x)" by (import pred_set RINV_DEF) constdefs - FINITE :: "('a => bool) => bool" + FINITE :: "('a::type => bool) => bool" "FINITE == -%s. ALL P. P EMPTY & (ALL s. P s --> (ALL e. P (INSERT e s))) --> P s" - -lemma FINITE_DEF: "ALL s. +%s::'a::type => bool. + ALL P::('a::type => bool) => bool. + P EMPTY & + (ALL s::'a::type => bool. + P s --> (ALL e::'a::type. P (INSERT e s))) --> + P s" + +lemma FINITE_DEF: "ALL s::'a::type => bool. FINITE s = - (ALL P. P EMPTY & (ALL s. P s --> (ALL e. P (INSERT e s))) --> P s)" + (ALL P::('a::type => bool) => bool. + P EMPTY & + (ALL s::'a::type => bool. + P s --> (ALL e::'a::type. P (INSERT e s))) --> + P s)" by (import pred_set FINITE_DEF) lemma FINITE_EMPTY: "FINITE EMPTY" by (import pred_set FINITE_EMPTY) -lemma FINITE_INDUCT: "ALL P. +lemma FINITE_INDUCT: "ALL P::('a::type => bool) => bool. P EMPTY & - (ALL s. FINITE s & P s --> (ALL e. ~ IN e s --> P (INSERT e s))) --> - (ALL s. FINITE s --> P s)" + (ALL s::'a::type => bool. + FINITE s & P s --> + (ALL e::'a::type. ~ IN e s --> P (INSERT e s))) --> + (ALL s::'a::type => bool. FINITE s --> P s)" by (import pred_set FINITE_INDUCT) -lemma FINITE_INSERT: "ALL x s. FINITE (INSERT x s) = FINITE s" +lemma FINITE_INSERT: "ALL (x::'a::type) s::'a::type => bool. FINITE (INSERT x s) = FINITE s" by (import pred_set FINITE_INSERT) -lemma FINITE_DELETE: "ALL x s. FINITE (DELETE s x) = FINITE s" +lemma FINITE_DELETE: "ALL (x::'a::type) s::'a::type => bool. FINITE (DELETE s x) = FINITE s" by (import pred_set FINITE_DELETE) -lemma FINITE_UNION: "ALL s t. FINITE (pred_set.UNION s t) = (FINITE s & FINITE t)" +lemma FINITE_UNION: "ALL (s::'a::type => bool) t::'a::type => bool. + FINITE (pred_set.UNION s t) = (FINITE s & FINITE t)" by (import pred_set FINITE_UNION) -lemma INTER_FINITE: "ALL s. FINITE s --> (ALL t. FINITE (pred_set.INTER s t))" +lemma INTER_FINITE: "ALL s::'a::type => bool. + FINITE s --> (ALL t::'a::type => bool. FINITE (pred_set.INTER s t))" by (import pred_set INTER_FINITE) -lemma SUBSET_FINITE: "ALL s. FINITE s --> (ALL t. SUBSET t s --> FINITE t)" +lemma SUBSET_FINITE: "ALL s::'a::type => bool. + FINITE s --> (ALL t::'a::type => bool. SUBSET t s --> FINITE t)" by (import pred_set SUBSET_FINITE) -lemma PSUBSET_FINITE: "ALL x. FINITE x --> (ALL xa. PSUBSET xa x --> FINITE xa)" +lemma PSUBSET_FINITE: "ALL x::'a::type => bool. + FINITE x --> (ALL xa::'a::type => bool. PSUBSET xa x --> FINITE xa)" by (import pred_set PSUBSET_FINITE) -lemma FINITE_DIFF: "ALL s. FINITE s --> (ALL t. FINITE (DIFF s t))" +lemma FINITE_DIFF: "ALL s::'a::type => bool. + FINITE s --> (ALL t::'a::type => bool. FINITE (DIFF s t))" by (import pred_set FINITE_DIFF) -lemma FINITE_SING: "ALL x. FINITE (INSERT x EMPTY)" +lemma FINITE_SING: "ALL x::'a::type. FINITE (INSERT x EMPTY)" by (import pred_set FINITE_SING) -lemma SING_FINITE: "ALL x. SING x --> FINITE x" +lemma SING_FINITE: "ALL x::'a::type => bool. SING x --> FINITE x" by (import pred_set SING_FINITE) -lemma IMAGE_FINITE: "ALL s. FINITE s --> (ALL f. FINITE (IMAGE f s))" +lemma IMAGE_FINITE: "ALL s::'a::type => bool. + FINITE s --> (ALL f::'a::type => 'b::type. FINITE (IMAGE f s))" by (import pred_set IMAGE_FINITE) consts - CARD :: "('a => bool) => nat" + CARD :: "('a::type => bool) => nat" specification (CARD) CARD_DEF: "(op &::bool => bool => bool) ((op =::nat => nat => bool) - ((CARD::('a => bool) => nat) (EMPTY::'a => bool)) (0::nat)) - ((All::(('a => bool) => bool) => bool) - (%s::'a => bool. - (op -->::bool => bool => bool) ((FINITE::('a => bool) => bool) s) - ((All::('a => bool) => bool) - (%x::'a. + ((CARD::('a::type => bool) => nat) (EMPTY::'a::type => bool)) (0::nat)) + ((All::(('a::type => bool) => bool) => bool) + (%s::'a::type => bool. + (op -->::bool => bool => bool) + ((FINITE::('a::type => bool) => bool) s) + ((All::('a::type => bool) => bool) + (%x::'a::type. (op =::nat => nat => bool) - ((CARD::('a => bool) => nat) - ((INSERT::'a => ('a => bool) => 'a => bool) x s)) + ((CARD::('a::type => bool) => nat) + ((INSERT::'a::type + => ('a::type => bool) => 'a::type => bool) + x s)) ((If::bool => nat => nat => nat) - ((IN::'a => ('a => bool) => bool) x s) - ((CARD::('a => bool) => nat) s) - ((Suc::nat => nat) ((CARD::('a => bool) => nat) s)))))))" + ((IN::'a::type => ('a::type => bool) => bool) x s) + ((CARD::('a::type => bool) => nat) s) + ((Suc::nat => nat) + ((CARD::('a::type => bool) => nat) s)))))))" by (import pred_set CARD_DEF) -lemma CARD_EMPTY: "CARD EMPTY = 0" +lemma CARD_EMPTY: "CARD EMPTY = (0::nat)" by (import pred_set CARD_EMPTY) -lemma CARD_INSERT: "ALL s. +lemma CARD_INSERT: "ALL s::'a::type => bool. FINITE s --> - (ALL x. CARD (INSERT x s) = (if IN x s then CARD s else Suc (CARD s)))" + (ALL x::'a::type. + CARD (INSERT x s) = (if IN x s then CARD s else Suc (CARD s)))" by (import pred_set CARD_INSERT) -lemma CARD_EQ_0: "ALL s. FINITE s --> (CARD s = 0) = (s = EMPTY)" +lemma CARD_EQ_0: "ALL s::'a::type => bool. FINITE s --> (CARD s = (0::nat)) = (s = EMPTY)" by (import pred_set CARD_EQ_0) -lemma CARD_DELETE: "ALL s. +lemma CARD_DELETE: "ALL s::'a::type => bool. FINITE s --> - (ALL x. CARD (DELETE s x) = (if IN x s then CARD s - 1 else CARD s))" + (ALL x::'a::type. + CARD (DELETE s x) = (if IN x s then CARD s - (1::nat) else CARD s))" by (import pred_set CARD_DELETE) -lemma CARD_INTER_LESS_EQ: "ALL s. FINITE s --> (ALL t. CARD (pred_set.INTER s t) <= CARD s)" +lemma CARD_INTER_LESS_EQ: "ALL s::'a::type => bool. + FINITE s --> + (ALL t::'a::type => bool. CARD (pred_set.INTER s t) <= CARD s)" by (import pred_set CARD_INTER_LESS_EQ) -lemma CARD_UNION: "ALL s. +lemma CARD_UNION: "ALL s::'a::type => bool. FINITE s --> - (ALL t. + (ALL t::'a::type => bool. FINITE t --> CARD (pred_set.UNION s t) + CARD (pred_set.INTER s t) = CARD s + CARD t)" by (import pred_set CARD_UNION) -lemma CARD_SUBSET: "ALL s. FINITE s --> (ALL t. SUBSET t s --> CARD t <= CARD s)" +lemma CARD_SUBSET: "ALL s::'a::type => bool. + FINITE s --> (ALL t::'a::type => bool. SUBSET t s --> CARD t <= CARD s)" by (import pred_set CARD_SUBSET) -lemma CARD_PSUBSET: "ALL s. FINITE s --> (ALL t. PSUBSET t s --> CARD t < CARD s)" +lemma CARD_PSUBSET: "ALL s::'a::type => bool. + FINITE s --> (ALL t::'a::type => bool. PSUBSET t s --> CARD t < CARD s)" by (import pred_set CARD_PSUBSET) -lemma CARD_SING: "ALL x. CARD (INSERT x EMPTY) = 1" +lemma CARD_SING: "ALL x::'a::type. CARD (INSERT x EMPTY) = (1::nat)" by (import pred_set CARD_SING) -lemma SING_IFF_CARD1: "ALL x. SING x = (CARD x = 1 & FINITE x)" +lemma SING_IFF_CARD1: "ALL x::'a::type => bool. SING x = (CARD x = (1::nat) & FINITE x)" by (import pred_set SING_IFF_CARD1) -lemma CARD_DIFF: "ALL t. +lemma CARD_DIFF: "ALL t::'a::type => bool. FINITE t --> - (ALL s. + (ALL s::'a::type => bool. FINITE s --> CARD (DIFF s t) = CARD s - CARD (pred_set.INTER s t))" by (import pred_set CARD_DIFF) -lemma LESS_CARD_DIFF: "ALL t. +lemma LESS_CARD_DIFF: "ALL t::'a::type => bool. FINITE t --> - (ALL s. FINITE s --> CARD t < CARD s --> 0 < CARD (DIFF s t))" + (ALL s::'a::type => bool. + FINITE s --> CARD t < CARD s --> (0::nat) < CARD (DIFF s t))" by (import pred_set LESS_CARD_DIFF) -lemma FINITE_COMPLETE_INDUCTION: "ALL P. - (ALL x. (ALL y. PSUBSET y x --> P y) --> FINITE x --> P x) --> - (ALL x. FINITE x --> P x)" +lemma FINITE_COMPLETE_INDUCTION: "ALL P::('a::type => bool) => bool. + (ALL x::'a::type => bool. + (ALL y::'a::type => bool. PSUBSET y x --> P y) --> + FINITE x --> P x) --> + (ALL x::'a::type => bool. FINITE x --> P x)" by (import pred_set FINITE_COMPLETE_INDUCTION) constdefs - INFINITE :: "('a => bool) => bool" - "INFINITE == %s. ~ FINITE s" - -lemma INFINITE_DEF: "ALL s. INFINITE s = (~ FINITE s)" + INFINITE :: "('a::type => bool) => bool" + "INFINITE == %s::'a::type => bool. ~ FINITE s" + +lemma INFINITE_DEF: "ALL s::'a::type => bool. INFINITE s = (~ FINITE s)" by (import pred_set INFINITE_DEF) lemma NOT_IN_FINITE: "(op =::bool => bool => bool) - ((INFINITE::('a => bool) => bool) (pred_set.UNIV::'a => bool)) - ((All::(('a => bool) => bool) => bool) - (%s::'a => bool. - (op -->::bool => bool => bool) ((FINITE::('a => bool) => bool) s) - ((Ex::('a => bool) => bool) - (%x::'a. - (Not::bool => bool) ((IN::'a => ('a => bool) => bool) x s)))))" + ((INFINITE::('a::type => bool) => bool) (pred_set.UNIV::'a::type => bool)) + ((All::(('a::type => bool) => bool) => bool) + (%s::'a::type => bool. + (op -->::bool => bool => bool) + ((FINITE::('a::type => bool) => bool) s) + ((Ex::('a::type => bool) => bool) + (%x::'a::type. + (Not::bool => bool) + ((IN::'a::type => ('a::type => bool) => bool) x s)))))" by (import pred_set NOT_IN_FINITE) -lemma INFINITE_INHAB: "ALL x. INFINITE x --> (EX xa. IN xa x)" +lemma INFINITE_INHAB: "ALL x::'a::type => bool. INFINITE x --> (EX xa::'a::type. IN xa x)" by (import pred_set INFINITE_INHAB) -lemma IMAGE_11_INFINITE: "ALL f. - (ALL x y. f x = f y --> x = y) --> - (ALL s. INFINITE s --> INFINITE (IMAGE f s))" +lemma IMAGE_11_INFINITE: "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))" by (import pred_set IMAGE_11_INFINITE) -lemma INFINITE_SUBSET: "ALL x. INFINITE x --> (ALL xa. SUBSET x xa --> INFINITE xa)" +lemma INFINITE_SUBSET: "ALL x::'a::type => bool. + INFINITE x --> (ALL xa::'a::type => bool. SUBSET x xa --> INFINITE xa)" by (import pred_set INFINITE_SUBSET) -lemma IN_INFINITE_NOT_FINITE: "ALL x xa. INFINITE x & FINITE xa --> (EX xb. IN xb x & ~ IN xb xa)" +lemma IN_INFINITE_NOT_FINITE: "ALL (x::'a::type => bool) xa::'a::type => bool. + INFINITE x & FINITE xa --> (EX xb::'a::type. IN xb x & ~ IN xb xa)" by (import pred_set IN_INFINITE_NOT_FINITE) lemma INFINITE_UNIV: "(op =::bool => bool => bool) - ((INFINITE::('a => bool) => bool) (pred_set.UNIV::'a => bool)) - ((Ex::(('a => 'a) => bool) => bool) - (%f::'a => 'a. + ((INFINITE::('a::type => bool) => bool) (pred_set.UNIV::'a::type => bool)) + ((Ex::(('a::type => 'a::type) => bool) => bool) + (%f::'a::type => 'a::type. (op &::bool => bool => bool) - ((All::('a => bool) => bool) - (%x::'a. - (All::('a => bool) => bool) - (%y::'a. + ((All::('a::type => bool) => bool) + (%x::'a::type. + (All::('a::type => bool) => bool) + (%y::'a::type. (op -->::bool => bool => bool) - ((op =::'a => 'a => bool) (f x) (f y)) - ((op =::'a => 'a => bool) x y)))) - ((Ex::('a => bool) => bool) - (%y::'a. - (All::('a => bool) => bool) - (%x::'a. + ((op =::'a::type => 'a::type => bool) (f x) (f y)) + ((op =::'a::type => 'a::type => bool) x y)))) + ((Ex::('a::type => bool) => bool) + (%y::'a::type. + (All::('a::type => bool) => bool) + (%x::'a::type. (Not::bool => bool) - ((op =::'a => 'a => bool) (f x) y))))))" + ((op =::'a::type => 'a::type => bool) (f x) y))))))" by (import pred_set INFINITE_UNIV) -lemma FINITE_PSUBSET_INFINITE: "ALL x. INFINITE x = (ALL xa. FINITE xa --> SUBSET xa x --> PSUBSET xa x)" +lemma FINITE_PSUBSET_INFINITE: "ALL x::'a::type => bool. + INFINITE x = + (ALL xa::'a::type => bool. FINITE xa --> SUBSET xa x --> PSUBSET xa x)" by (import pred_set FINITE_PSUBSET_INFINITE) lemma FINITE_PSUBSET_UNIV: "(op =::bool => bool => bool) - ((INFINITE::('a => bool) => bool) (pred_set.UNIV::'a => bool)) - ((All::(('a => bool) => bool) => bool) - (%s::'a => bool. - (op -->::bool => bool => bool) ((FINITE::('a => bool) => bool) s) - ((PSUBSET::('a => bool) => ('a => bool) => bool) s - (pred_set.UNIV::'a => bool))))" + ((INFINITE::('a::type => bool) => bool) (pred_set.UNIV::'a::type => bool)) + ((All::(('a::type => bool) => bool) => bool) + (%s::'a::type => bool. + (op -->::bool => bool => bool) + ((FINITE::('a::type => bool) => bool) s) + ((PSUBSET::('a::type => bool) => ('a::type => bool) => bool) s + (pred_set.UNIV::'a::type => bool))))" by (import pred_set FINITE_PSUBSET_UNIV) -lemma INFINITE_DIFF_FINITE: "ALL s t. INFINITE s & FINITE t --> DIFF s t ~= EMPTY" +lemma INFINITE_DIFF_FINITE: "ALL (s::'a::type => bool) t::'a::type => bool. + INFINITE s & FINITE t --> DIFF s t ~= EMPTY" by (import pred_set INFINITE_DIFF_FINITE) -lemma FINITE_ISO_NUM: "ALL s. +lemma FINITE_ISO_NUM: "ALL s::'a::type => bool. FINITE s --> - (EX f. (ALL n m. n < CARD s & m < CARD s --> f n = f m --> n = m) & - s = GSPEC (%n. (f n, n < CARD s)))" + (EX f::nat => 'a::type. + (ALL (n::nat) m::nat. + n < CARD s & m < CARD s --> f n = f m --> n = m) & + s = GSPEC (%n::nat. (f n, n < CARD s)))" by (import pred_set FINITE_ISO_NUM) -lemma FINITE_WEAK_ENUMERATE: "(All::(('a => bool) => bool) => bool) - (%x::'a => bool. - (op =::bool => bool => bool) ((FINITE::('a => bool) => bool) x) - ((Ex::((nat => 'a) => bool) => bool) - (%f::nat => 'a. +lemma FINITE_WEAK_ENUMERATE: "(All::(('a::type => bool) => bool) => bool) + (%x::'a::type => bool. + (op =::bool => bool => bool) ((FINITE::('a::type => bool) => bool) x) + ((Ex::((nat => 'a::type) => bool) => bool) + (%f::nat => 'a::type. (Ex::(nat => bool) => bool) (%b::nat. - (All::('a => bool) => bool) - (%e::'a. + (All::('a::type => bool) => bool) + (%e::'a::type. (op =::bool => bool => bool) - ((IN::'a => ('a => bool) => bool) e x) + ((IN::'a::type => ('a::type => bool) => bool) e x) ((Ex::(nat => bool) => bool) (%n::nat. (op &::bool => bool => bool) ((op <::nat => nat => bool) n b) - ((op =::'a => 'a => bool) e (f n)))))))))" + ((op =::'a::type => 'a::type => bool) e + (f n)))))))))" by (import pred_set FINITE_WEAK_ENUMERATE) constdefs - BIGUNION :: "(('a => bool) => bool) => 'a => bool" - "BIGUNION == %P. GSPEC (%x. (x, EX p. IN p P & IN x p))" - -lemma BIGUNION: "ALL P. BIGUNION P = GSPEC (%x. (x, EX p. IN p P & IN x p))" + BIGUNION :: "(('a::type => bool) => bool) => 'a::type => bool" + "BIGUNION == +%P::('a::type => bool) => bool. + GSPEC (%x::'a::type. (x, EX p::'a::type => bool. IN p P & IN x p))" + +lemma BIGUNION: "ALL P::('a::type => bool) => bool. + BIGUNION P = + GSPEC (%x::'a::type. (x, EX p::'a::type => bool. IN p P & IN x p))" by (import pred_set BIGUNION) -lemma IN_BIGUNION: "ALL x xa. IN x (BIGUNION xa) = (EX s. IN x s & IN s xa)" +lemma IN_BIGUNION: "ALL (x::'a::type) xa::('a::type => bool) => bool. + IN x (BIGUNION xa) = (EX s::'a::type => bool. IN x s & IN s xa)" by (import pred_set IN_BIGUNION) lemma BIGUNION_EMPTY: "BIGUNION EMPTY = EMPTY" by (import pred_set BIGUNION_EMPTY) -lemma BIGUNION_SING: "ALL x. BIGUNION (INSERT x EMPTY) = x" +lemma BIGUNION_SING: "ALL x::'a::type => bool. BIGUNION (INSERT x EMPTY) = x" by (import pred_set BIGUNION_SING) -lemma BIGUNION_UNION: "ALL x xa. +lemma BIGUNION_UNION: "ALL (x::('a::type => bool) => bool) xa::('a::type => bool) => bool. BIGUNION (pred_set.UNION x xa) = pred_set.UNION (BIGUNION x) (BIGUNION xa)" by (import pred_set BIGUNION_UNION) -lemma DISJOINT_BIGUNION: "(ALL (s::('a => bool) => bool) t::'a => bool. +lemma DISJOINT_BIGUNION: "(ALL (s::('a::type => bool) => bool) t::'a::type => bool. DISJOINT (BIGUNION s) t = - (ALL s'::'a => bool. IN s' s --> DISJOINT s' t)) & -(ALL (x::('a => bool) => bool) xa::'a => bool. + (ALL s'::'a::type => bool. IN s' s --> DISJOINT s' t)) & +(ALL (x::('a::type => bool) => bool) xa::'a::type => bool. DISJOINT xa (BIGUNION x) = - (ALL xb::'a => bool. IN xb x --> DISJOINT xa xb))" + (ALL xb::'a::type => bool. IN xb x --> DISJOINT xa xb))" by (import pred_set DISJOINT_BIGUNION) -lemma BIGUNION_INSERT: "ALL x xa. BIGUNION (INSERT x xa) = pred_set.UNION x (BIGUNION xa)" +lemma BIGUNION_INSERT: "ALL (x::'a::type => bool) xa::('a::type => bool) => bool. + BIGUNION (INSERT x xa) = pred_set.UNION x (BIGUNION xa)" by (import pred_set BIGUNION_INSERT) -lemma BIGUNION_SUBSET: "ALL X P. SUBSET (BIGUNION P) X = (ALL Y. IN Y P --> SUBSET Y X)" +lemma BIGUNION_SUBSET: "ALL (X::'a::type => bool) P::('a::type => bool) => bool. + SUBSET (BIGUNION P) X = (ALL Y::'a::type => bool. IN Y P --> SUBSET Y X)" by (import pred_set BIGUNION_SUBSET) -lemma FINITE_BIGUNION: "ALL x. FINITE x & (ALL s. IN s x --> FINITE s) --> FINITE (BIGUNION x)" +lemma FINITE_BIGUNION: "ALL x::('a::type => bool) => bool. + FINITE x & (ALL s::'a::type => bool. IN s x --> FINITE s) --> + FINITE (BIGUNION x)" by (import pred_set FINITE_BIGUNION) constdefs - BIGINTER :: "(('a => bool) => bool) => 'a => bool" - "BIGINTER == %B. GSPEC (%x. (x, ALL P. IN P B --> IN x P))" - -lemma BIGINTER: "ALL B. BIGINTER B = GSPEC (%x. (x, ALL P. IN P B --> IN x P))" + BIGINTER :: "(('a::type => bool) => bool) => 'a::type => bool" + "BIGINTER == +%B::('a::type => bool) => bool. + GSPEC (%x::'a::type. (x, ALL P::'a::type => bool. IN P B --> IN x P))" + +lemma BIGINTER: "ALL B::('a::type => bool) => bool. + BIGINTER B = + GSPEC (%x::'a::type. (x, ALL P::'a::type => bool. IN P B --> IN x P))" by (import pred_set BIGINTER) -lemma IN_BIGINTER: "IN x (BIGINTER B) = (ALL P. IN P B --> IN x P)" +lemma IN_BIGINTER: "IN (x::'a::type) (BIGINTER (B::('a::type => bool) => bool)) = +(ALL P::'a::type => bool. IN P B --> IN x P)" by (import pred_set IN_BIGINTER) -lemma BIGINTER_INSERT: "ALL P B. BIGINTER (INSERT P B) = pred_set.INTER P (BIGINTER B)" +lemma BIGINTER_INSERT: "ALL (P::'a::type => bool) B::('a::type => bool) => bool. + BIGINTER (INSERT P B) = pred_set.INTER P (BIGINTER B)" by (import pred_set BIGINTER_INSERT) lemma BIGINTER_EMPTY: "BIGINTER EMPTY = pred_set.UNIV" by (import pred_set BIGINTER_EMPTY) -lemma BIGINTER_INTER: "ALL x xa. BIGINTER (INSERT x (INSERT xa EMPTY)) = pred_set.INTER x xa" +lemma BIGINTER_INTER: "ALL (x::'a::type => bool) xa::'a::type => bool. + BIGINTER (INSERT x (INSERT xa EMPTY)) = pred_set.INTER x xa" by (import pred_set BIGINTER_INTER) -lemma BIGINTER_SING: "ALL x. BIGINTER (INSERT x EMPTY) = x" +lemma BIGINTER_SING: "ALL x::'a::type => bool. BIGINTER (INSERT x EMPTY) = x" by (import pred_set BIGINTER_SING) -lemma SUBSET_BIGINTER: "ALL X P. SUBSET X (BIGINTER P) = (ALL x. IN x P --> SUBSET X x)" +lemma SUBSET_BIGINTER: "ALL (X::'a::type => bool) P::('a::type => bool) => bool. + SUBSET X (BIGINTER P) = (ALL x::'a::type => bool. IN x P --> SUBSET X x)" by (import pred_set SUBSET_BIGINTER) -lemma DISJOINT_BIGINTER: "ALL x xa xb. +lemma DISJOINT_BIGINTER: "ALL (x::'a::type => bool) (xa::'a::type => bool) + xb::('a::type => bool) => bool. IN xa xb & DISJOINT xa x --> DISJOINT x (BIGINTER xb) & DISJOINT (BIGINTER xb) x" by (import pred_set DISJOINT_BIGINTER) constdefs - CROSS :: "('a => bool) => ('b => bool) => 'a * 'b => bool" - "CROSS == %P Q. GSPEC (%p. (p, IN (fst p) P & IN (snd p) Q))" - -lemma CROSS_DEF: "ALL P Q. CROSS P Q = GSPEC (%p. (p, IN (fst p) P & IN (snd p) Q))" + CROSS :: "('a::type => bool) => ('b::type => bool) => 'a::type * 'b::type => bool" + "CROSS == +%(P::'a::type => bool) Q::'b::type => bool. + GSPEC (%p::'a::type * 'b::type. (p, IN (fst p) P & IN (snd p) Q))" + +lemma CROSS_DEF: "ALL (P::'a::type => bool) Q::'b::type => bool. + CROSS P Q = + GSPEC (%p::'a::type * 'b::type. (p, IN (fst p) P & IN (snd p) Q))" by (import pred_set CROSS_DEF) -lemma IN_CROSS: "ALL x xa xb. IN xb (CROSS x xa) = (IN (fst xb) x & IN (snd xb) xa)" +lemma IN_CROSS: "ALL (x::'a::type => bool) (xa::'b::type => bool) xb::'a::type * 'b::type. + IN xb (CROSS x xa) = (IN (fst xb) x & IN (snd xb) xa)" by (import pred_set IN_CROSS) -lemma CROSS_EMPTY: "ALL x. CROSS x EMPTY = EMPTY & CROSS EMPTY x = EMPTY" +lemma CROSS_EMPTY: "ALL x::'a::type => bool. CROSS x EMPTY = EMPTY & CROSS EMPTY x = EMPTY" by (import pred_set CROSS_EMPTY) -lemma CROSS_INSERT_LEFT: "ALL x xa xb. +lemma CROSS_INSERT_LEFT: "ALL (x::'a::type => bool) (xa::'b::type => bool) xb::'a::type. CROSS (INSERT xb x) xa = pred_set.UNION (CROSS (INSERT xb EMPTY) xa) (CROSS x xa)" by (import pred_set CROSS_INSERT_LEFT) -lemma CROSS_INSERT_RIGHT: "ALL x xa xb. +lemma CROSS_INSERT_RIGHT: "ALL (x::'a::type => bool) (xa::'b::type => bool) xb::'b::type. CROSS x (INSERT xb xa) = pred_set.UNION (CROSS x (INSERT xb EMPTY)) (CROSS x xa)" by (import pred_set CROSS_INSERT_RIGHT) -lemma FINITE_CROSS: "ALL x xa. FINITE x & FINITE xa --> FINITE (CROSS x xa)" +lemma FINITE_CROSS: "ALL (x::'a::type => bool) xa::'b::type => bool. + FINITE x & FINITE xa --> FINITE (CROSS x xa)" by (import pred_set FINITE_CROSS) -lemma CROSS_SINGS: "ALL x xa. CROSS (INSERT x EMPTY) (INSERT xa EMPTY) = INSERT (x, xa) EMPTY" +lemma CROSS_SINGS: "ALL (x::'a::type) xa::'b::type. + CROSS (INSERT x EMPTY) (INSERT xa EMPTY) = INSERT (x, xa) EMPTY" by (import pred_set CROSS_SINGS) -lemma CARD_SING_CROSS: "ALL x s. FINITE s --> CARD (CROSS (INSERT x EMPTY) s) = CARD s" +lemma CARD_SING_CROSS: "ALL (x::'a::type) s::'b::type => bool. + FINITE s --> CARD (CROSS (INSERT x EMPTY) s) = CARD s" by (import pred_set CARD_SING_CROSS) -lemma CARD_CROSS: "ALL x xa. FINITE x & FINITE xa --> CARD (CROSS x xa) = CARD x * CARD xa" +lemma CARD_CROSS: "ALL (x::'a::type => bool) xa::'b::type => bool. + FINITE x & FINITE xa --> CARD (CROSS x xa) = CARD x * CARD xa" by (import pred_set CARD_CROSS) -lemma CROSS_SUBSET: "ALL x xa xb xc. +lemma CROSS_SUBSET: "ALL (x::'a::type => bool) (xa::'b::type => bool) (xb::'a::type => bool) + xc::'b::type => bool. SUBSET (CROSS xb xc) (CROSS x xa) = (xb = EMPTY | xc = EMPTY | SUBSET xb x & SUBSET xc xa)" by (import pred_set CROSS_SUBSET) -lemma FINITE_CROSS_EQ: "ALL P Q. FINITE (CROSS P Q) = (P = EMPTY | Q = EMPTY | FINITE P & FINITE Q)" +lemma FINITE_CROSS_EQ: "ALL (P::'a::type => bool) Q::'b::type => bool. + FINITE (CROSS P Q) = (P = EMPTY | Q = EMPTY | FINITE P & FINITE Q)" by (import pred_set FINITE_CROSS_EQ) constdefs - COMPL :: "('a => bool) => 'a => bool" + COMPL :: "('a::type => bool) => 'a::type => bool" "COMPL == DIFF pred_set.UNIV" -lemma COMPL_DEF: "ALL P. COMPL P = DIFF pred_set.UNIV P" +lemma COMPL_DEF: "ALL P::'a::type => bool. COMPL P = DIFF pred_set.UNIV P" by (import pred_set COMPL_DEF) -lemma IN_COMPL: "ALL x xa. IN x (COMPL xa) = (~ IN x xa)" +lemma IN_COMPL: "ALL (x::'a::type) xa::'a::type => bool. IN x (COMPL xa) = (~ IN x xa)" by (import pred_set IN_COMPL) -lemma COMPL_COMPL: "ALL x. COMPL (COMPL x) = x" +lemma COMPL_COMPL: "ALL x::'a::type => bool. COMPL (COMPL x) = x" by (import pred_set COMPL_COMPL) -lemma COMPL_CLAUSES: "ALL x. +lemma COMPL_CLAUSES: "ALL x::'a::type => bool. pred_set.INTER (COMPL x) x = EMPTY & pred_set.UNION (COMPL x) x = pred_set.UNIV" by (import pred_set COMPL_CLAUSES) -lemma COMPL_SPLITS: "ALL x xa. +lemma COMPL_SPLITS: "ALL (x::'a::type => bool) xa::'a::type => bool. pred_set.UNION (pred_set.INTER x xa) (pred_set.INTER (COMPL x) xa) = xa" by (import pred_set COMPL_SPLITS) -lemma INTER_UNION_COMPL: "ALL x xa. pred_set.INTER x xa = COMPL (pred_set.UNION (COMPL x) (COMPL xa))" +lemma INTER_UNION_COMPL: "ALL (x::'a::type => bool) xa::'a::type => bool. + pred_set.INTER x xa = COMPL (pred_set.UNION (COMPL x) (COMPL xa))" by (import pred_set INTER_UNION_COMPL) lemma COMPL_EMPTY: "COMPL EMPTY = pred_set.UNIV" @@ -3659,48 +4315,54 @@ count :: "nat => nat => bool" defs - count_primdef: "count == %n. GSPEC (%m. (m, m < n))" - -lemma count_def: "ALL n. count n = GSPEC (%m. (m, m < n))" + count_primdef: "count == %n::nat. GSPEC (%m::nat. (m, m < n))" + +lemma count_def: "ALL n::nat. count n = GSPEC (%m::nat. (m, m < n))" by (import pred_set count_def) -lemma IN_COUNT: "ALL m n. IN m (count n) = (m < n)" +lemma IN_COUNT: "ALL (m::nat) n::nat. IN m (count n) = (m < n)" by (import pred_set IN_COUNT) -lemma COUNT_ZERO: "count 0 = EMPTY" +lemma COUNT_ZERO: "count (0::nat) = EMPTY" by (import pred_set COUNT_ZERO) -lemma COUNT_SUC: "ALL n. count (Suc n) = INSERT n (count n)" +lemma COUNT_SUC: "ALL n::nat. count (Suc n) = INSERT n (count n)" by (import pred_set COUNT_SUC) -lemma FINITE_COUNT: "ALL n. FINITE (count n)" +lemma FINITE_COUNT: "ALL n::nat. FINITE (count n)" by (import pred_set FINITE_COUNT) -lemma CARD_COUNT: "ALL n. CARD (count n) = n" +lemma CARD_COUNT: "ALL n::nat. CARD (count n) = n" by (import pred_set CARD_COUNT) constdefs - ITSET_tupled :: "('a => 'b => 'b) => ('a => bool) * 'b => 'b" + ITSET_tupled :: "('a::type => 'b::type => 'b::type) +=> ('a::type => bool) * 'b::type => 'b::type" "ITSET_tupled == -%f. WFREC - (SOME R. - WF R & - (ALL b s. - FINITE s & s ~= EMPTY --> R (REST s, f (CHOICE s) b) (s, b))) - (%ITSET_tupled (v, v1). - if FINITE v - then if v = EMPTY then v1 - else ITSET_tupled (REST v, f (CHOICE v) v1) - else ARB)" - -lemma ITSET_tupled_primitive_def: "ALL f. +%f::'a::type => 'b::type => 'b::type. + WFREC + (SOME R::('a::type => bool) * 'b::type + => ('a::type => bool) * 'b::type => bool. + WF R & + (ALL (b::'b::type) s::'a::type => bool. + FINITE s & s ~= EMPTY --> R (REST s, f (CHOICE s) b) (s, b))) + (%(ITSET_tupled::('a::type => bool) * 'b::type => 'b::type) + (v::'a::type => bool, v1::'b::type). + if FINITE v + then if v = EMPTY then v1 + else ITSET_tupled (REST v, f (CHOICE v) v1) + else ARB)" + +lemma ITSET_tupled_primitive_def: "ALL f::'a::type => 'b::type => 'b::type. ITSET_tupled f = WFREC - (SOME R. + (SOME R::('a::type => bool) * 'b::type + => ('a::type => bool) * 'b::type => bool. WF R & - (ALL b s. + (ALL (b::'b::type) s::'a::type => bool. FINITE s & s ~= EMPTY --> R (REST s, f (CHOICE s) b) (s, b))) - (%ITSET_tupled (v, v1). + (%(ITSET_tupled::('a::type => bool) * 'b::type => 'b::type) + (v::'a::type => bool, v1::'b::type). if FINITE v then if v = EMPTY then v1 else ITSET_tupled (REST v, f (CHOICE v) v1) @@ -3708,26 +4370,32 @@ by (import pred_set ITSET_tupled_primitive_def) constdefs - ITSET :: "('a => 'b => 'b) => ('a => bool) => 'b => 'b" - "ITSET == %f x x1. ITSET_tupled f (x, x1)" - -lemma ITSET_curried_def: "ALL f x x1. ITSET f x x1 = ITSET_tupled f (x, x1)" + ITSET :: "('a::type => 'b::type => 'b::type) +=> ('a::type => bool) => 'b::type => 'b::type" + "ITSET == +%(f::'a::type => 'b::type => 'b::type) (x::'a::type => bool) x1::'b::type. + ITSET_tupled f (x, x1)" + +lemma ITSET_curried_def: "ALL (f::'a::type => 'b::type => 'b::type) (x::'a::type => bool) + x1::'b::type. ITSET f x x1 = ITSET_tupled f (x, x1)" by (import pred_set ITSET_curried_def) -lemma ITSET_IND: "ALL P. - (ALL s b. - (FINITE s & s ~= EMPTY --> P (REST s) (f (CHOICE s) b)) --> +lemma ITSET_IND: "ALL P::('a::type => bool) => 'b::type => bool. + (ALL (s::'a::type => bool) b::'b::type. + (FINITE s & s ~= EMPTY --> + P (REST s) ((f::'a::type => 'b::type => 'b::type) (CHOICE s) b)) --> P s b) --> - (ALL v. All (P v))" + (ALL v::'a::type => bool. All (P v))" by (import pred_set ITSET_IND) -lemma ITSET_THM: "ALL s f b. +lemma ITSET_THM: "ALL (s::'a::type => bool) (f::'a::type => 'b::type => 'b::type) b::'b::type. FINITE s --> ITSET f s b = (if s = EMPTY then b else ITSET f (REST s) (f (CHOICE s) b))" by (import pred_set ITSET_THM) -lemma ITSET_EMPTY: "ALL x xa. ITSET x EMPTY xa = xa" +lemma ITSET_EMPTY: "ALL (x::'a::type => 'b::type => 'b::type) xa::'b::type. + ITSET x EMPTY xa = xa" by (import pred_set ITSET_EMPTY) ;end_setup @@ -3735,45 +4403,67 @@ ;setup_theory operator constdefs - ASSOC :: "('a => 'a => 'a) => bool" - "ASSOC == %f. ALL x y z. f x (f y z) = f (f x y) z" - -lemma ASSOC_DEF: "ALL f. ASSOC f = (ALL x y z. f x (f y z) = f (f x y) z)" + ASSOC :: "('a::type => 'a::type => 'a::type) => bool" + "ASSOC == +%f::'a::type => 'a::type => 'a::type. + ALL (x::'a::type) (y::'a::type) z::'a::type. f x (f y z) = f (f x y) z" + +lemma ASSOC_DEF: "ALL f::'a::type => 'a::type => 'a::type. + ASSOC f = + (ALL (x::'a::type) (y::'a::type) z::'a::type. f x (f y z) = f (f x y) z)" by (import operator ASSOC_DEF) constdefs - COMM :: "('a => 'a => 'b) => bool" - "COMM == %f. ALL x y. f x y = f y x" - -lemma COMM_DEF: "ALL f. COMM f = (ALL x y. f x y = f y x)" + COMM :: "('a::type => 'a::type => 'b::type) => bool" + "COMM == +%f::'a::type => 'a::type => 'b::type. + ALL (x::'a::type) y::'a::type. f x y = f y x" + +lemma COMM_DEF: "ALL f::'a::type => 'a::type => 'b::type. + COMM f = (ALL (x::'a::type) y::'a::type. f x y = f y x)" by (import operator COMM_DEF) constdefs - FCOMM :: "('a => 'b => 'a) => ('c => 'a => 'a) => bool" - "FCOMM == %f g. ALL x y z. g x (f y z) = f (g x y) z" - -lemma FCOMM_DEF: "ALL f g. FCOMM f g = (ALL x y z. g x (f y z) = f (g x y) z)" + FCOMM :: "('a::type => 'b::type => 'a::type) +=> ('c::type => 'a::type => 'a::type) => bool" + "FCOMM == +%(f::'a::type => 'b::type => 'a::type) g::'c::type => 'a::type => 'a::type. + ALL (x::'c::type) (y::'a::type) z::'b::type. g x (f y z) = f (g x y) z" + +lemma FCOMM_DEF: "ALL (f::'a::type => 'b::type => 'a::type) + g::'c::type => 'a::type => 'a::type. + FCOMM f g = + (ALL (x::'c::type) (y::'a::type) z::'b::type. g x (f y z) = f (g x y) z)" by (import operator FCOMM_DEF) constdefs - RIGHT_ID :: "('a => 'b => 'a) => 'b => bool" - "RIGHT_ID == %f e. ALL x. f x e = x" - -lemma RIGHT_ID_DEF: "ALL f e. RIGHT_ID f e = (ALL x. f x e = x)" + RIGHT_ID :: "('a::type => 'b::type => 'a::type) => 'b::type => bool" + "RIGHT_ID == +%(f::'a::type => 'b::type => 'a::type) e::'b::type. + ALL x::'a::type. f x e = x" + +lemma RIGHT_ID_DEF: "ALL (f::'a::type => 'b::type => 'a::type) e::'b::type. + RIGHT_ID f e = (ALL x::'a::type. f x e = x)" by (import operator RIGHT_ID_DEF) constdefs - LEFT_ID :: "('a => 'b => 'b) => 'a => bool" - "LEFT_ID == %f e. ALL x. f e x = x" - -lemma LEFT_ID_DEF: "ALL f e. LEFT_ID f e = (ALL x. f e x = x)" + LEFT_ID :: "('a::type => 'b::type => 'b::type) => 'a::type => bool" + "LEFT_ID == +%(f::'a::type => 'b::type => 'b::type) e::'a::type. + ALL x::'b::type. f e x = x" + +lemma LEFT_ID_DEF: "ALL (f::'a::type => 'b::type => 'b::type) e::'a::type. + LEFT_ID f e = (ALL x::'b::type. f e x = x)" by (import operator LEFT_ID_DEF) constdefs - MONOID :: "('a => 'a => 'a) => 'a => bool" - "MONOID == %f e. ASSOC f & RIGHT_ID f e & LEFT_ID f e" - -lemma MONOID_DEF: "ALL f e. MONOID f e = (ASSOC f & RIGHT_ID f e & LEFT_ID f e)" + MONOID :: "('a::type => 'a::type => 'a::type) => 'a::type => bool" + "MONOID == +%(f::'a::type => 'a::type => 'a::type) e::'a::type. + ASSOC f & RIGHT_ID f e & LEFT_ID f e" + +lemma MONOID_DEF: "ALL (f::'a::type => 'a::type => 'a::type) e::'a::type. + MONOID f e = (ASSOC f & RIGHT_ID f e & LEFT_ID f e)" by (import operator MONOID_DEF) lemma ASSOC_CONJ: "ASSOC op &" @@ -3782,7 +4472,7 @@ lemma ASSOC_DISJ: "ASSOC op |" by (import operator ASSOC_DISJ) -lemma FCOMM_ASSOC: "ALL x. FCOMM x x = ASSOC x" +lemma FCOMM_ASSOC: "ALL x::'a::type => 'a::type => 'a::type. FCOMM x x = ASSOC x" by (import operator FCOMM_ASSOC) lemma MONOID_CONJ_T: "MONOID op & True" @@ -3796,29 +4486,35 @@ ;setup_theory rich_list consts - SNOC :: "'a => 'a list => 'a list" - -specification (SNOC) SNOC: "(ALL x::'a. SNOC x [] = [x]) & -(ALL (x::'a) (x'::'a) l::'a list. SNOC x (x' # l) = x' # SNOC x l)" + SNOC :: "'a::type => 'a::type list => 'a::type list" + +specification (SNOC) SNOC: "(ALL x::'a::type. SNOC x [] = [x]) & +(ALL (x::'a::type) (x'::'a::type) l::'a::type list. + SNOC x (x' # l) = x' # SNOC x l)" by (import rich_list SNOC) consts - SCANL :: "('b => 'a => 'b) => 'b => 'a list => 'b list" - -specification (SCANL) SCANL: "(ALL (f::'b => 'a => 'b) e::'b. SCANL f e [] = [e]) & -(ALL (f::'b => 'a => 'b) (e::'b) (x::'a) l::'a list. - SCANL f e (x # l) = e # SCANL f (f e x) l)" + SCANL :: "('b::type => 'a::type => 'b::type) +=> 'b::type => 'a::type list => 'b::type list" + +specification (SCANL) SCANL: "(ALL (f::'b::type => 'a::type => 'b::type) e::'b::type. + SCANL f e [] = [e]) & +(ALL (f::'b::type => 'a::type => 'b::type) (e::'b::type) (x::'a::type) + l::'a::type list. SCANL f e (x # l) = e # SCANL f (f e x) l)" by (import rich_list SCANL) consts - SCANR :: "('a => 'b => 'b) => 'b => 'a list => 'b list" - -specification (SCANR) SCANR: "(ALL (f::'a => 'b => 'b) e::'b. SCANR f e [] = [e]) & -(ALL (f::'a => 'b => 'b) (e::'b) (x::'a) l::'a list. + SCANR :: "('a::type => 'b::type => 'b::type) +=> 'b::type => 'a::type list => 'b::type list" + +specification (SCANR) SCANR: "(ALL (f::'a::type => 'b::type => 'b::type) e::'b::type. + SCANR f e [] = [e]) & +(ALL (f::'a::type => 'b::type => 'b::type) (e::'b::type) (x::'a::type) + l::'a::type list. SCANR f e (x # l) = f x (hd (SCANR f e l)) # SCANR f e l)" by (import rich_list SCANR) -lemma IS_EL_DEF: "ALL x l. x mem l = list_exists (op = x) l" +lemma IS_EL_DEF: "ALL (x::'a::type) l::'a::type list. x mem l = list_exists (op = x) l" by (import rich_list IS_EL_DEF) constdefs @@ -3836,255 +4532,290 @@ by (import rich_list OR_EL_DEF) consts - FIRSTN :: "nat => 'a list => 'a list" - -specification (FIRSTN) FIRSTN: "(ALL l::'a list. FIRSTN (0::nat) l = []) & -(ALL (n::nat) (x::'a) l::'a list. FIRSTN (Suc n) (x # l) = x # FIRSTN n l)" + FIRSTN :: "nat => 'a::type list => 'a::type list" + +specification (FIRSTN) FIRSTN: "(ALL l::'a::type list. FIRSTN (0::nat) l = []) & +(ALL (n::nat) (x::'a::type) l::'a::type list. + FIRSTN (Suc n) (x # l) = x # FIRSTN n l)" by (import rich_list FIRSTN) consts - BUTFIRSTN :: "nat => 'a list => 'a list" - -specification (BUTFIRSTN) BUTFIRSTN: "(ALL l::'a list. BUTFIRSTN (0::nat) l = l) & -(ALL (n::nat) (x::'a) l::'a list. BUTFIRSTN (Suc n) (x # l) = BUTFIRSTN n l)" + BUTFIRSTN :: "nat => 'a::type list => 'a::type list" + +specification (BUTFIRSTN) BUTFIRSTN: "(ALL l::'a::type list. BUTFIRSTN (0::nat) l = l) & +(ALL (n::nat) (x::'a::type) l::'a::type list. + BUTFIRSTN (Suc n) (x # l) = BUTFIRSTN n l)" by (import rich_list BUTFIRSTN) consts - SEG :: "nat => nat => 'a list => 'a list" - -specification (SEG) SEG: "(ALL (k::nat) l::'a list. SEG (0::nat) k l = []) & -(ALL (m::nat) (x::'a) l::'a list. + SEG :: "nat => nat => 'a::type list => 'a::type list" + +specification (SEG) SEG: "(ALL (k::nat) l::'a::type list. SEG (0::nat) k l = []) & +(ALL (m::nat) (x::'a::type) l::'a::type list. SEG (Suc m) (0::nat) (x # l) = x # SEG m (0::nat) l) & -(ALL (m::nat) (k::nat) (x::'a) l::'a list. +(ALL (m::nat) (k::nat) (x::'a::type) l::'a::type list. SEG (Suc m) (Suc k) (x # l) = SEG (Suc m) k l)" by (import rich_list SEG) -lemma LAST: "ALL x l. last (SNOC x l) = x" +lemma LAST: "ALL (x::'a::type) l::'a::type list. last (SNOC x l) = x" by (import rich_list LAST) -lemma BUTLAST: "ALL x l. butlast (SNOC x l) = l" +lemma BUTLAST: "ALL (x::'a::type) l::'a::type list. butlast (SNOC x l) = l" by (import rich_list BUTLAST) consts - LASTN :: "nat => 'a list => 'a list" - -specification (LASTN) LASTN: "(ALL l::'a list. LASTN (0::nat) l = []) & -(ALL (n::nat) (x::'a) l::'a list. + LASTN :: "nat => 'a::type list => 'a::type list" + +specification (LASTN) LASTN: "(ALL l::'a::type list. LASTN (0::nat) l = []) & +(ALL (n::nat) (x::'a::type) l::'a::type list. LASTN (Suc n) (SNOC x l) = SNOC x (LASTN n l))" by (import rich_list LASTN) consts - BUTLASTN :: "nat => 'a list => 'a list" - -specification (BUTLASTN) BUTLASTN: "(ALL l::'a list. BUTLASTN (0::nat) l = l) & -(ALL (n::nat) (x::'a) l::'a list. + BUTLASTN :: "nat => 'a::type list => 'a::type list" + +specification (BUTLASTN) BUTLASTN: "(ALL l::'a::type list. BUTLASTN (0::nat) l = l) & +(ALL (n::nat) (x::'a::type) l::'a::type list. BUTLASTN (Suc n) (SNOC x l) = BUTLASTN n l)" by (import rich_list BUTLASTN) -lemma EL: "(ALL x::'a list. EL (0::nat) x = hd x) & -(ALL (x::nat) xa::'a list. EL (Suc x) xa = EL x (tl xa))" +lemma EL: "(ALL x::'a::type list. EL (0::nat) x = hd x) & +(ALL (x::nat) xa::'a::type list. EL (Suc x) xa = EL x (tl xa))" by (import rich_list EL) consts - ELL :: "nat => 'a list => 'a" - -specification (ELL) ELL: "(ALL l::'a list. ELL (0::nat) l = last l) & -(ALL (n::nat) l::'a list. ELL (Suc n) l = ELL n (butlast l))" + ELL :: "nat => 'a::type list => 'a::type" + +specification (ELL) ELL: "(ALL l::'a::type list. ELL (0::nat) l = last l) & +(ALL (n::nat) l::'a::type list. ELL (Suc n) l = ELL n (butlast l))" by (import rich_list ELL) consts - IS_PREFIX :: "'a list => 'a list => bool" - -specification (IS_PREFIX) IS_PREFIX: "(ALL l::'a list. IS_PREFIX l [] = True) & -(ALL (x::'a) l::'a list. IS_PREFIX [] (x # l) = False) & -(ALL (x1::'a) (l1::'a list) (x2::'a) l2::'a list. + IS_PREFIX :: "'a::type list => 'a::type list => bool" + +specification (IS_PREFIX) IS_PREFIX: "(ALL l::'a::type list. IS_PREFIX l [] = True) & +(ALL (x::'a::type) l::'a::type list. IS_PREFIX [] (x # l) = False) & +(ALL (x1::'a::type) (l1::'a::type list) (x2::'a::type) l2::'a::type list. IS_PREFIX (x1 # l1) (x2 # l2) = (x1 = x2 & IS_PREFIX l1 l2))" by (import rich_list IS_PREFIX) -lemma SNOC_APPEND: "ALL x l. SNOC x l = l @ [x]" +lemma SNOC_APPEND: "ALL (x::'a::type) l::'a::type list. SNOC x l = l @ [x]" by (import rich_list SNOC_APPEND) -lemma REVERSE: "rev [] = [] & (ALL (x::'a) xa::'a list. rev (x # xa) = SNOC x (rev xa))" +lemma REVERSE: "rev [] = [] & +(ALL (x::'a::type) xa::'a::type list. rev (x # xa) = SNOC x (rev xa))" by (import rich_list REVERSE) -lemma REVERSE_SNOC: "ALL x l. rev (SNOC x l) = x # rev l" +lemma REVERSE_SNOC: "ALL (x::'a::type) l::'a::type list. rev (SNOC x l) = x # rev l" by (import rich_list REVERSE_SNOC) -lemma SNOC_Axiom: "ALL (e::'b) f::'a => 'a list => 'b => 'b. - EX x::'a list => 'b. - x [] = e & (ALL (xa::'a) l::'a list. x (SNOC xa l) = f xa l (x l))" +lemma SNOC_Axiom: "ALL (e::'b::type) f::'a::type => 'a::type list => 'b::type => 'b::type. + EX x::'a::type list => 'b::type. + x [] = e & + (ALL (xa::'a::type) l::'a::type list. x (SNOC xa l) = f xa l (x l))" by (import rich_list SNOC_Axiom) consts - IS_SUFFIX :: "'a list => 'a list => bool" - -specification (IS_SUFFIX) IS_SUFFIX: "(ALL l::'a list. IS_SUFFIX l [] = True) & -(ALL (x::'a) l::'a list. IS_SUFFIX [] (SNOC x l) = False) & -(ALL (x1::'a) (l1::'a list) (x2::'a) l2::'a list. + IS_SUFFIX :: "'a::type list => 'a::type list => bool" + +specification (IS_SUFFIX) IS_SUFFIX: "(ALL l::'a::type list. IS_SUFFIX l [] = True) & +(ALL (x::'a::type) l::'a::type list. IS_SUFFIX [] (SNOC x l) = False) & +(ALL (x1::'a::type) (l1::'a::type list) (x2::'a::type) l2::'a::type list. IS_SUFFIX (SNOC x1 l1) (SNOC x2 l2) = (x1 = x2 & IS_SUFFIX l1 l2))" by (import rich_list IS_SUFFIX) consts - IS_SUBLIST :: "'a list => 'a list => bool" - -specification (IS_SUBLIST) IS_SUBLIST: "(ALL l::'a list. IS_SUBLIST l [] = True) & -(ALL (x::'a) l::'a list. IS_SUBLIST [] (x # l) = False) & -(ALL (x1::'a) (l1::'a list) (x2::'a) l2::'a list. + IS_SUBLIST :: "'a::type list => 'a::type list => bool" + +specification (IS_SUBLIST) IS_SUBLIST: "(ALL l::'a::type list. IS_SUBLIST l [] = True) & +(ALL (x::'a::type) l::'a::type list. IS_SUBLIST [] (x # l) = False) & +(ALL (x1::'a::type) (l1::'a::type list) (x2::'a::type) l2::'a::type list. IS_SUBLIST (x1 # l1) (x2 # l2) = (x1 = x2 & IS_PREFIX l1 l2 | IS_SUBLIST l1 (x2 # l2)))" by (import rich_list IS_SUBLIST) consts - SPLITP :: "('a => bool) => 'a list => 'a list * 'a list" - -specification (SPLITP) SPLITP: "(ALL P::'a => bool. SPLITP P [] = ([], [])) & -(ALL (P::'a => bool) (x::'a) l::'a list. + SPLITP :: "('a::type => bool) => 'a::type list => 'a::type list * 'a::type list" + +specification (SPLITP) SPLITP: "(ALL P::'a::type => bool. SPLITP P [] = ([], [])) & +(ALL (P::'a::type => bool) (x::'a::type) l::'a::type list. SPLITP P (x # l) = (if P x then ([], x # l) else (x # fst (SPLITP P l), snd (SPLITP P l))))" by (import rich_list SPLITP) constdefs - PREFIX :: "('a => bool) => 'a list => 'a list" - "PREFIX == %P l. fst (SPLITP (Not o P) l)" - -lemma PREFIX_DEF: "ALL P l. PREFIX P l = fst (SPLITP (Not o P) l)" + PREFIX :: "('a::type => bool) => 'a::type list => 'a::type list" + "PREFIX == %(P::'a::type => bool) l::'a::type list. fst (SPLITP (Not o P) l)" + +lemma PREFIX_DEF: "ALL (P::'a::type => bool) l::'a::type list. + PREFIX P l = fst (SPLITP (Not o P) l)" by (import rich_list PREFIX_DEF) constdefs - SUFFIX :: "('a => bool) => 'a list => 'a list" - "SUFFIX == %P. foldl (%l' x. if P x then SNOC x l' else []) []" - -lemma SUFFIX_DEF: "ALL P l. SUFFIX P l = foldl (%l' x. if P x then SNOC x l' else []) [] l" + SUFFIX :: "('a::type => bool) => 'a::type list => 'a::type list" + "SUFFIX == +%P::'a::type => bool. + foldl (%(l'::'a::type list) x::'a::type. if P x then SNOC x l' else []) + []" + +lemma SUFFIX_DEF: "ALL (P::'a::type => bool) l::'a::type list. + SUFFIX P l = + foldl (%(l'::'a::type list) x::'a::type. if P x then SNOC x l' else []) + [] l" by (import rich_list SUFFIX_DEF) constdefs - UNZIP_FST :: "('a * 'b) list => 'a list" - "UNZIP_FST == %l. fst (unzip l)" - -lemma UNZIP_FST_DEF: "ALL l. UNZIP_FST l = fst (unzip l)" + UNZIP_FST :: "('a::type * 'b::type) list => 'a::type list" + "UNZIP_FST == %l::('a::type * 'b::type) list. fst (unzip l)" + +lemma UNZIP_FST_DEF: "ALL l::('a::type * 'b::type) list. UNZIP_FST l = fst (unzip l)" by (import rich_list UNZIP_FST_DEF) constdefs - UNZIP_SND :: "('a * 'b) list => 'b list" - "UNZIP_SND == %l. snd (unzip l)" - -lemma UNZIP_SND_DEF: "ALL l. UNZIP_SND l = snd (unzip l)" + UNZIP_SND :: "('a::type * 'b::type) list => 'b::type list" + "UNZIP_SND == %l::('a::type * 'b::type) list. snd (unzip l)" + +lemma UNZIP_SND_DEF: "ALL l::('a::type * 'b::type) list. UNZIP_SND l = snd (unzip l)" by (import rich_list UNZIP_SND_DEF) consts - GENLIST :: "(nat => 'a) => nat => 'a list" - -specification (GENLIST) GENLIST: "(ALL f::nat => 'a. GENLIST f (0::nat) = []) & -(ALL (f::nat => 'a) n::nat. GENLIST f (Suc n) = SNOC (f n) (GENLIST f n))" + GENLIST :: "(nat => 'a::type) => nat => 'a::type list" + +specification (GENLIST) GENLIST: "(ALL f::nat => 'a::type. GENLIST f (0::nat) = []) & +(ALL (f::nat => 'a::type) n::nat. + GENLIST f (Suc n) = SNOC (f n) (GENLIST f n))" by (import rich_list GENLIST) consts - REPLICATE :: "nat => 'a => 'a list" - -specification (REPLICATE) REPLICATE: "(ALL x::'a. REPLICATE (0::nat) x = []) & -(ALL (n::nat) x::'a. REPLICATE (Suc n) x = x # REPLICATE n x)" + REPLICATE :: "nat => 'a::type => 'a::type list" + +specification (REPLICATE) REPLICATE: "(ALL x::'a::type. REPLICATE (0::nat) x = []) & +(ALL (n::nat) x::'a::type. REPLICATE (Suc n) x = x # REPLICATE n x)" by (import rich_list REPLICATE) -lemma LENGTH_MAP2: "ALL l1 l2. +lemma LENGTH_MAP2: "ALL (l1::'a::type list) l2::'b::type list. length l1 = length l2 --> - (ALL f. + (ALL f::'a::type => 'b::type => 'c::type. length (map2 f l1 l2) = length l1 & length (map2 f l1 l2) = length l2)" by (import rich_list LENGTH_MAP2) -lemma NULL_EQ_NIL: "ALL l. null l = (l = [])" +lemma NULL_EQ_NIL: "ALL l::'a::type list. null l = (l = [])" by (import rich_list NULL_EQ_NIL) -lemma LENGTH_EQ: "ALL x y. x = y --> length x = length y" +lemma LENGTH_EQ: "ALL (x::'a::type list) y::'a::type list. x = y --> length x = length y" by (import rich_list LENGTH_EQ) -lemma LENGTH_NOT_NULL: "ALL l. (0 < length l) = (~ null l)" +lemma LENGTH_NOT_NULL: "ALL l::'a::type list. ((0::nat) < length l) = (~ null l)" by (import rich_list LENGTH_NOT_NULL) -lemma SNOC_INDUCT: "ALL P. P [] & (ALL l. P l --> (ALL x. P (SNOC x l))) --> All P" +lemma SNOC_INDUCT: "ALL P::'a::type list => bool. + P [] & + (ALL l::'a::type list. P l --> (ALL x::'a::type. P (SNOC x l))) --> + All P" by (import rich_list SNOC_INDUCT) -lemma SNOC_CASES: "ALL x'. x' = [] | (EX x l. x' = SNOC x l)" +lemma SNOC_CASES: "ALL x'::'a::type list. + x' = [] | (EX (x::'a::type) l::'a::type list. x' = SNOC x l)" by (import rich_list SNOC_CASES) -lemma LENGTH_SNOC: "ALL x l. length (SNOC x l) = Suc (length l)" +lemma LENGTH_SNOC: "ALL (x::'a::type) l::'a::type list. length (SNOC x l) = Suc (length l)" by (import rich_list LENGTH_SNOC) -lemma NOT_NIL_SNOC: "ALL x xa. [] ~= SNOC x xa" +lemma NOT_NIL_SNOC: "ALL (x::'a::type) xa::'a::type list. [] ~= SNOC x xa" by (import rich_list NOT_NIL_SNOC) -lemma NOT_SNOC_NIL: "ALL x xa. SNOC x xa ~= []" +lemma NOT_SNOC_NIL: "ALL (x::'a::type) xa::'a::type list. SNOC x xa ~= []" by (import rich_list NOT_SNOC_NIL) -lemma SNOC_11: "ALL x l x' l'. (SNOC x l = SNOC x' l') = (x = x' & l = l')" +lemma SNOC_11: "ALL (x::'a::type) (l::'a::type list) (x'::'a::type) l'::'a::type list. + (SNOC x l = SNOC x' l') = (x = x' & l = l')" by (import rich_list SNOC_11) -lemma SNOC_EQ_LENGTH_EQ: "ALL x1 l1 x2 l2. SNOC x1 l1 = SNOC x2 l2 --> length l1 = length l2" +lemma SNOC_EQ_LENGTH_EQ: "ALL (x1::'a::type) (l1::'a::type list) (x2::'a::type) l2::'a::type list. + SNOC x1 l1 = SNOC x2 l2 --> length l1 = length l2" by (import rich_list SNOC_EQ_LENGTH_EQ) -lemma SNOC_REVERSE_CONS: "ALL x xa. SNOC x xa = rev (x # rev xa)" +lemma SNOC_REVERSE_CONS: "ALL (x::'a::type) xa::'a::type list. SNOC x xa = rev (x # rev xa)" by (import rich_list SNOC_REVERSE_CONS) -lemma MAP_SNOC: "ALL x xa xb. map x (SNOC xa xb) = SNOC (x xa) (map x xb)" +lemma MAP_SNOC: "ALL (x::'a::type => 'b::type) (xa::'a::type) xb::'a::type list. + map x (SNOC xa xb) = SNOC (x xa) (map x xb)" by (import rich_list MAP_SNOC) -lemma FOLDR_SNOC: "ALL f e x l. foldr f (SNOC x l) e = foldr f l (f x e)" +lemma FOLDR_SNOC: "ALL (f::'a::type => 'b::type => 'b::type) (e::'b::type) (x::'a::type) + l::'a::type list. foldr f (SNOC x l) e = foldr f l (f x e)" by (import rich_list FOLDR_SNOC) -lemma FOLDL_SNOC: "ALL (f::'b => 'a => 'b) (e::'b) (x::'a) l::'a list. - foldl f e (SNOC x l) = f (foldl f e l) x" +lemma FOLDL_SNOC: "ALL (f::'b::type => 'a::type => 'b::type) (e::'b::type) (x::'a::type) + l::'a::type list. foldl f e (SNOC x l) = f (foldl f e l) x" by (import rich_list FOLDL_SNOC) -lemma FOLDR_FOLDL: "ALL f e. MONOID f e --> (ALL l. foldr f l e = foldl f e l)" +lemma FOLDR_FOLDL: "ALL (f::'a::type => 'a::type => 'a::type) e::'a::type. + MONOID f e --> (ALL l::'a::type list. foldr f l e = foldl f e l)" by (import rich_list FOLDR_FOLDL) -lemma LENGTH_FOLDR: "ALL l. length l = foldr (%x. Suc) l 0" +lemma LENGTH_FOLDR: "ALL l::'a::type list. length l = foldr (%x::'a::type. Suc) l (0::nat)" by (import rich_list LENGTH_FOLDR) -lemma LENGTH_FOLDL: "ALL l. length l = foldl (%l' x. Suc l') 0 l" +lemma LENGTH_FOLDL: "ALL l::'a::type list. + length l = foldl (%(l'::nat) x::'a::type. Suc l') (0::nat) l" by (import rich_list LENGTH_FOLDL) -lemma MAP_FOLDR: "ALL f l. map f l = foldr (%x. op # (f x)) l []" +lemma MAP_FOLDR: "ALL (f::'a::type => 'b::type) l::'a::type list. + map f l = foldr (%x::'a::type. op # (f x)) l []" by (import rich_list MAP_FOLDR) -lemma MAP_FOLDL: "ALL f l. map f l = foldl (%l' x. SNOC (f x) l') [] l" +lemma MAP_FOLDL: "ALL (f::'a::type => 'b::type) l::'a::type list. + map f l = foldl (%(l'::'b::type list) x::'a::type. SNOC (f x) l') [] l" by (import rich_list MAP_FOLDL) -lemma MAP_o: "ALL (f::'b => 'c) g::'a => 'b. map (f o g) = map f o map g" +lemma MAP_o: "ALL (f::'b::type => 'c::type) g::'a::type => 'b::type. + map (f o g) = map f o map g" by (import rich_list MAP_o) -lemma FILTER_FOLDR: "ALL P l. filter P l = foldr (%x l'. if P x then x # l' else l') l []" +lemma FILTER_FOLDR: "ALL (P::'a::type => bool) l::'a::type list. + filter P l = + foldr (%(x::'a::type) l'::'a::type list. if P x then x # l' else l') l []" by (import rich_list FILTER_FOLDR) -lemma FILTER_SNOC: "ALL P x l. +lemma FILTER_SNOC: "ALL (P::'a::type => bool) (x::'a::type) l::'a::type list. filter P (SNOC x l) = (if P x then SNOC x (filter P l) else filter P l)" by (import rich_list FILTER_SNOC) -lemma FILTER_FOLDL: "ALL P l. filter P l = foldl (%l' x. if P x then SNOC x l' else l') [] l" +lemma FILTER_FOLDL: "ALL (P::'a::type => bool) l::'a::type list. + filter P l = + foldl (%(l'::'a::type list) x::'a::type. if P x then SNOC x l' else l') + [] l" by (import rich_list FILTER_FOLDL) -lemma FILTER_COMM: "ALL f1 f2 l. filter f1 (filter f2 l) = filter f2 (filter f1 l)" +lemma FILTER_COMM: "ALL (f1::'a::type => bool) (f2::'a::type => bool) l::'a::type list. + filter f1 (filter f2 l) = filter f2 (filter f1 l)" by (import rich_list FILTER_COMM) -lemma FILTER_IDEM: "ALL f l. filter f (filter f l) = filter f l" +lemma FILTER_IDEM: "ALL (f::'a::type => bool) l::'a::type list. + filter f (filter f l) = filter f l" by (import rich_list FILTER_IDEM) -lemma LENGTH_SEG: "ALL n k l. n + k <= length l --> length (SEG n k l) = n" +lemma LENGTH_SEG: "ALL (n::nat) (k::nat) l::'a::type list. + n + k <= length l --> length (SEG n k l) = n" by (import rich_list LENGTH_SEG) -lemma APPEND_NIL: "(ALL l::'a list. l @ [] = l) & (ALL x::'a list. [] @ x = x)" +lemma APPEND_NIL: "(ALL l::'a::type list. l @ [] = l) & (ALL x::'a::type list. [] @ x = x)" by (import rich_list APPEND_NIL) -lemma APPEND_SNOC: "ALL l1 x l2. l1 @ SNOC x l2 = SNOC x (l1 @ l2)" +lemma APPEND_SNOC: "ALL (l1::'a::type list) (x::'a::type) l2::'a::type list. + l1 @ SNOC x l2 = SNOC x (l1 @ l2)" by (import rich_list APPEND_SNOC) -lemma APPEND_FOLDR: "ALL l1 l2. l1 @ l2 = foldr op # l1 l2" +lemma APPEND_FOLDR: "ALL (l1::'a::type list) l2::'a::type list. l1 @ l2 = foldr op # l1 l2" by (import rich_list APPEND_FOLDR) -lemma APPEND_FOLDL: "ALL l1 l2. l1 @ l2 = foldl (%l' x. SNOC x l') l1 l2" +lemma APPEND_FOLDL: "ALL (l1::'a::type list) l2::'a::type list. + l1 @ l2 = foldl (%(l'::'a::type list) x::'a::type. SNOC x l') l1 l2" by (import rich_list APPEND_FOLDL) -lemma CONS_APPEND: "ALL x l. x # l = [x] @ l" +lemma CONS_APPEND: "ALL (x::'a::type) l::'a::type list. x # l = [x] @ l" by (import rich_list CONS_APPEND) lemma ASSOC_APPEND: "ASSOC op @" @@ -4093,767 +4824,924 @@ lemma MONOID_APPEND_NIL: "MONOID op @ []" by (import rich_list MONOID_APPEND_NIL) -lemma APPEND_LENGTH_EQ: "ALL l1 l1'. +lemma APPEND_LENGTH_EQ: "ALL (l1::'a::type list) l1'::'a::type list. length l1 = length l1' --> - (ALL l2 l2'. + (ALL (l2::'a::type list) l2'::'a::type list. length l2 = length l2' --> (l1 @ l2 = l1' @ l2') = (l1 = l1' & l2 = l2'))" by (import rich_list APPEND_LENGTH_EQ) -lemma FLAT_SNOC: "ALL x l. concat (SNOC x l) = concat l @ x" +lemma FLAT_SNOC: "ALL (x::'a::type list) l::'a::type list list. + concat (SNOC x l) = concat l @ x" by (import rich_list FLAT_SNOC) -lemma FLAT_FOLDR: "ALL l. concat l = foldr op @ l []" +lemma FLAT_FOLDR: "ALL l::'a::type list list. concat l = foldr op @ l []" by (import rich_list FLAT_FOLDR) -lemma FLAT_FOLDL: "ALL l. concat l = foldl op @ [] l" +lemma FLAT_FOLDL: "ALL l::'a::type list list. concat l = foldl op @ [] l" by (import rich_list FLAT_FOLDL) -lemma LENGTH_FLAT: "ALL l. length (concat l) = sum (map size l)" +lemma LENGTH_FLAT: "ALL l::'a::type list list. length (concat l) = sum (map size l)" by (import rich_list LENGTH_FLAT) -lemma REVERSE_FOLDR: "ALL l. rev l = foldr SNOC l []" +lemma REVERSE_FOLDR: "ALL l::'a::type list. rev l = foldr SNOC l []" by (import rich_list REVERSE_FOLDR) -lemma REVERSE_FOLDL: "ALL l. rev l = foldl (%l' x. x # l') [] l" +lemma REVERSE_FOLDL: "ALL l::'a::type list. + rev l = foldl (%(l'::'a::type list) x::'a::type. x # l') [] l" by (import rich_list REVERSE_FOLDL) -lemma ALL_EL_SNOC: "ALL P x l. list_all P (SNOC x l) = (list_all P l & P x)" +lemma ALL_EL_SNOC: "ALL (P::'a::type => bool) (x::'a::type) l::'a::type list. + list_all P (SNOC x l) = (list_all P l & P x)" by (import rich_list ALL_EL_SNOC) -lemma ALL_EL_MAP: "ALL (P::'b => bool) (f::'a => 'b) l::'a list. +lemma ALL_EL_MAP: "ALL (P::'b::type => bool) (f::'a::type => 'b::type) l::'a::type list. list_all P (map f l) = list_all (P o f) l" by (import rich_list ALL_EL_MAP) -lemma SOME_EL_SNOC: "ALL P x l. list_exists P (SNOC x l) = (P x | list_exists P l)" +lemma SOME_EL_SNOC: "ALL (P::'a::type => bool) (x::'a::type) l::'a::type list. + list_exists P (SNOC x l) = (P x | list_exists P l)" by (import rich_list SOME_EL_SNOC) -lemma IS_EL_SNOC: "ALL y x l. y mem SNOC x l = (y = x | y mem l)" +lemma IS_EL_SNOC: "ALL (y::'a::type) (x::'a::type) l::'a::type list. + y mem SNOC x l = (y = x | y mem l)" by (import rich_list IS_EL_SNOC) -lemma SUM_SNOC: "ALL x l. sum (SNOC x l) = sum l + x" +lemma SUM_SNOC: "ALL (x::nat) l::nat list. sum (SNOC x l) = sum l + x" by (import rich_list SUM_SNOC) -lemma SUM_FOLDL: "ALL l. sum l = foldl op + 0 l" +lemma SUM_FOLDL: "ALL l::nat list. sum l = foldl op + (0::nat) l" by (import rich_list SUM_FOLDL) -lemma IS_PREFIX_APPEND: "ALL l1 l2. IS_PREFIX l1 l2 = (EX l. l1 = l2 @ l)" +lemma IS_PREFIX_APPEND: "ALL (l1::'a::type list) l2::'a::type list. + IS_PREFIX l1 l2 = (EX l::'a::type list. l1 = l2 @ l)" by (import rich_list IS_PREFIX_APPEND) -lemma IS_SUFFIX_APPEND: "ALL l1 l2. IS_SUFFIX l1 l2 = (EX l. l1 = l @ l2)" +lemma IS_SUFFIX_APPEND: "ALL (l1::'a::type list) l2::'a::type list. + IS_SUFFIX l1 l2 = (EX l::'a::type list. l1 = l @ l2)" by (import rich_list IS_SUFFIX_APPEND) -lemma IS_SUBLIST_APPEND: "ALL l1 l2. IS_SUBLIST l1 l2 = (EX l l'. l1 = l @ l2 @ l')" +lemma IS_SUBLIST_APPEND: "ALL (l1::'a::type list) l2::'a::type list. + IS_SUBLIST l1 l2 = + (EX (l::'a::type list) l'::'a::type list. l1 = l @ l2 @ l')" by (import rich_list IS_SUBLIST_APPEND) -lemma IS_PREFIX_IS_SUBLIST: "ALL l1 l2. IS_PREFIX l1 l2 --> IS_SUBLIST l1 l2" +lemma IS_PREFIX_IS_SUBLIST: "ALL (l1::'a::type list) l2::'a::type list. + IS_PREFIX l1 l2 --> IS_SUBLIST l1 l2" by (import rich_list IS_PREFIX_IS_SUBLIST) -lemma IS_SUFFIX_IS_SUBLIST: "ALL l1 l2. IS_SUFFIX l1 l2 --> IS_SUBLIST l1 l2" +lemma IS_SUFFIX_IS_SUBLIST: "ALL (l1::'a::type list) l2::'a::type list. + IS_SUFFIX l1 l2 --> IS_SUBLIST l1 l2" by (import rich_list IS_SUFFIX_IS_SUBLIST) -lemma IS_PREFIX_REVERSE: "ALL l1 l2. IS_PREFIX (rev l1) (rev l2) = IS_SUFFIX l1 l2" +lemma IS_PREFIX_REVERSE: "ALL (l1::'a::type list) l2::'a::type list. + IS_PREFIX (rev l1) (rev l2) = IS_SUFFIX l1 l2" by (import rich_list IS_PREFIX_REVERSE) -lemma IS_SUFFIX_REVERSE: "ALL l2 l1. IS_SUFFIX (rev l1) (rev l2) = IS_PREFIX l1 l2" +lemma IS_SUFFIX_REVERSE: "ALL (l2::'a::type list) l1::'a::type list. + IS_SUFFIX (rev l1) (rev l2) = IS_PREFIX l1 l2" by (import rich_list IS_SUFFIX_REVERSE) -lemma IS_SUBLIST_REVERSE: "ALL l1 l2. IS_SUBLIST (rev l1) (rev l2) = IS_SUBLIST l1 l2" +lemma IS_SUBLIST_REVERSE: "ALL (l1::'a::type list) l2::'a::type list. + IS_SUBLIST (rev l1) (rev l2) = IS_SUBLIST l1 l2" by (import rich_list IS_SUBLIST_REVERSE) -lemma PREFIX_FOLDR: "ALL P x. PREFIX P x = foldr (%x l'. if P x then x # l' else []) x []" +lemma PREFIX_FOLDR: "ALL (P::'a::type => bool) x::'a::type list. + PREFIX P x = + foldr (%(x::'a::type) l'::'a::type list. if P x then x # l' else []) x []" by (import rich_list PREFIX_FOLDR) -lemma PREFIX: "(ALL x::'a => bool. PREFIX x [] = []) & -(ALL (x::'a => bool) (xa::'a) xb::'a list. +lemma PREFIX: "(ALL x::'a::type => bool. PREFIX x [] = []) & +(ALL (x::'a::type => bool) (xa::'a::type) xb::'a::type list. PREFIX x (xa # xb) = (if x xa then xa # PREFIX x xb else []))" by (import rich_list PREFIX) -lemma IS_PREFIX_PREFIX: "ALL P l. IS_PREFIX l (PREFIX P l)" +lemma IS_PREFIX_PREFIX: "ALL (P::'a::type => bool) l::'a::type list. IS_PREFIX l (PREFIX P l)" by (import rich_list IS_PREFIX_PREFIX) -lemma LENGTH_SCANL: "ALL (f::'b => 'a => 'b) (e::'b) l::'a list. +lemma LENGTH_SCANL: "ALL (f::'b::type => 'a::type => 'b::type) (e::'b::type) l::'a::type list. length (SCANL f e l) = Suc (length l)" by (import rich_list LENGTH_SCANL) -lemma LENGTH_SCANR: "ALL f e l. length (SCANR f e l) = Suc (length l)" +lemma LENGTH_SCANR: "ALL (f::'a::type => 'b::type => 'b::type) (e::'b::type) l::'a::type list. + length (SCANR f e l) = Suc (length l)" by (import rich_list LENGTH_SCANR) -lemma COMM_MONOID_FOLDL: "ALL x. +lemma COMM_MONOID_FOLDL: "ALL x::'a::type => 'a::type => 'a::type. COMM x --> - (ALL xa. MONOID x xa --> (ALL e l. foldl x e l = x e (foldl x xa l)))" + (ALL xa::'a::type. + MONOID x xa --> + (ALL (e::'a::type) l::'a::type list. + foldl x e l = x e (foldl x xa l)))" by (import rich_list COMM_MONOID_FOLDL) -lemma COMM_MONOID_FOLDR: "ALL x. +lemma COMM_MONOID_FOLDR: "ALL x::'a::type => 'a::type => 'a::type. COMM x --> - (ALL xa. MONOID x xa --> (ALL e l. foldr x l e = x e (foldr x l xa)))" + (ALL xa::'a::type. + MONOID x xa --> + (ALL (e::'a::type) l::'a::type list. + foldr x l e = x e (foldr x l xa)))" by (import rich_list COMM_MONOID_FOLDR) -lemma FCOMM_FOLDR_APPEND: "ALL x xa. +lemma FCOMM_FOLDR_APPEND: "ALL (x::'a::type => 'a::type => 'a::type) + xa::'b::type => 'a::type => 'a::type. FCOMM x xa --> - (ALL xb. + (ALL xb::'a::type. LEFT_ID x xb --> - (ALL l1 l2. + (ALL (l1::'b::type list) l2::'b::type list. foldr xa (l1 @ l2) xb = x (foldr xa l1 xb) (foldr xa l2 xb)))" by (import rich_list FCOMM_FOLDR_APPEND) -lemma FCOMM_FOLDL_APPEND: "ALL x xa. +lemma FCOMM_FOLDL_APPEND: "ALL (x::'a::type => 'b::type => 'a::type) + xa::'a::type => 'a::type => 'a::type. FCOMM x xa --> - (ALL xb. + (ALL xb::'a::type. RIGHT_ID xa xb --> - (ALL l1 l2. + (ALL (l1::'b::type list) l2::'b::type list. foldl x xb (l1 @ l2) = xa (foldl x xb l1) (foldl x xb l2)))" by (import rich_list FCOMM_FOLDL_APPEND) -lemma FOLDL_SINGLE: "ALL x xa xb. foldl x xa [xb] = x xa xb" +lemma FOLDL_SINGLE: "ALL (x::'a::type => 'b::type => 'a::type) (xa::'a::type) xb::'b::type. + foldl x xa [xb] = x xa xb" by (import rich_list FOLDL_SINGLE) -lemma FOLDR_SINGLE: "ALL x xa xb. foldr x [xb] xa = x xb xa" +lemma FOLDR_SINGLE: "ALL (x::'a::type => 'b::type => 'b::type) (xa::'b::type) xb::'a::type. + foldr x [xb] xa = x xb xa" by (import rich_list FOLDR_SINGLE) -lemma FOLDR_CONS_NIL: "ALL l. foldr op # l [] = l" +lemma FOLDR_CONS_NIL: "ALL l::'a::type list. foldr op # l [] = l" by (import rich_list FOLDR_CONS_NIL) -lemma FOLDL_SNOC_NIL: "ALL l. foldl (%xs x. SNOC x xs) [] l = l" +lemma FOLDL_SNOC_NIL: "ALL l::'a::type list. + foldl (%(xs::'a::type list) x::'a::type. SNOC x xs) [] l = l" by (import rich_list FOLDL_SNOC_NIL) -lemma FOLDR_REVERSE: "ALL x xa xb. foldr x (rev xb) xa = foldl (%xa y. x y xa) xa xb" +lemma FOLDR_REVERSE: "ALL (x::'a::type => 'b::type => 'b::type) (xa::'b::type) xb::'a::type list. + foldr x (rev xb) xa = foldl (%(xa::'b::type) y::'a::type. x y xa) xa xb" by (import rich_list FOLDR_REVERSE) -lemma FOLDL_REVERSE: "ALL x xa xb. foldl x xa (rev xb) = foldr (%xa y. x y xa) xb xa" +lemma FOLDL_REVERSE: "ALL (x::'a::type => 'b::type => 'a::type) (xa::'a::type) xb::'b::type list. + foldl x xa (rev xb) = foldr (%(xa::'b::type) y::'a::type. x y xa) xb xa" by (import rich_list FOLDL_REVERSE) -lemma FOLDR_MAP: "ALL (f::'a => 'a => 'a) (e::'a) (g::'b => 'a) l::'b list. - foldr f (map g l) e = foldr (%x::'b. f (g x)) l e" +lemma FOLDR_MAP: "ALL (f::'a::type => 'a::type => 'a::type) (e::'a::type) + (g::'b::type => 'a::type) l::'b::type list. + foldr f (map g l) e = foldr (%x::'b::type. f (g x)) l e" by (import rich_list FOLDR_MAP) -lemma FOLDL_MAP: "ALL (f::'a => 'a => 'a) (e::'a) (g::'b => 'a) l::'b list. - foldl f e (map g l) = foldl (%(x::'a) y::'b. f x (g y)) e l" +lemma FOLDL_MAP: "ALL (f::'a::type => 'a::type => 'a::type) (e::'a::type) + (g::'b::type => 'a::type) l::'b::type list. + foldl f e (map g l) = foldl (%(x::'a::type) y::'b::type. f x (g y)) e l" by (import rich_list FOLDL_MAP) -lemma ALL_EL_FOLDR: "ALL P l. list_all P l = foldr (%x. op & (P x)) l True" +lemma ALL_EL_FOLDR: "ALL (P::'a::type => bool) l::'a::type list. + list_all P l = foldr (%x::'a::type. op & (P x)) l True" by (import rich_list ALL_EL_FOLDR) -lemma ALL_EL_FOLDL: "ALL P l. list_all P l = foldl (%l' x. l' & P x) True l" +lemma ALL_EL_FOLDL: "ALL (P::'a::type => bool) l::'a::type list. + list_all P l = foldl (%(l'::bool) x::'a::type. l' & P x) True l" by (import rich_list ALL_EL_FOLDL) -lemma SOME_EL_FOLDR: "ALL P l. list_exists P l = foldr (%x. op | (P x)) l False" +lemma SOME_EL_FOLDR: "ALL (P::'a::type => bool) l::'a::type list. + list_exists P l = foldr (%x::'a::type. op | (P x)) l False" by (import rich_list SOME_EL_FOLDR) -lemma SOME_EL_FOLDL: "ALL P l. list_exists P l = foldl (%l' x. l' | P x) False l" +lemma SOME_EL_FOLDL: "ALL (P::'a::type => bool) l::'a::type list. + list_exists P l = foldl (%(l'::bool) x::'a::type. l' | P x) False l" by (import rich_list SOME_EL_FOLDL) -lemma ALL_EL_FOLDR_MAP: "ALL x xa. list_all x xa = foldr op & (map x xa) True" +lemma ALL_EL_FOLDR_MAP: "ALL (x::'a::type => bool) xa::'a::type list. + list_all x xa = foldr op & (map x xa) True" by (import rich_list ALL_EL_FOLDR_MAP) -lemma ALL_EL_FOLDL_MAP: "ALL x xa. list_all x xa = foldl op & True (map x xa)" +lemma ALL_EL_FOLDL_MAP: "ALL (x::'a::type => bool) xa::'a::type list. + list_all x xa = foldl op & True (map x xa)" by (import rich_list ALL_EL_FOLDL_MAP) -lemma SOME_EL_FOLDR_MAP: "ALL x xa. list_exists x xa = foldr op | (map x xa) False" +lemma SOME_EL_FOLDR_MAP: "ALL (x::'a::type => bool) xa::'a::type list. + list_exists x xa = foldr op | (map x xa) False" by (import rich_list SOME_EL_FOLDR_MAP) -lemma SOME_EL_FOLDL_MAP: "ALL x xa. list_exists x xa = foldl op | False (map x xa)" +lemma SOME_EL_FOLDL_MAP: "ALL (x::'a::type => bool) xa::'a::type list. + list_exists x xa = foldl op | False (map x xa)" by (import rich_list SOME_EL_FOLDL_MAP) -lemma FOLDR_FILTER: "ALL (f::'a => 'a => 'a) (e::'a) (P::'a => bool) l::'a list. +lemma FOLDR_FILTER: "ALL (f::'a::type => 'a::type => 'a::type) (e::'a::type) + (P::'a::type => bool) l::'a::type list. foldr f (filter P l) e = - foldr (%(x::'a) y::'a. if P x then f x y else y) l e" + foldr (%(x::'a::type) y::'a::type. if P x then f x y else y) l e" by (import rich_list FOLDR_FILTER) -lemma FOLDL_FILTER: "ALL (f::'a => 'a => 'a) (e::'a) (P::'a => bool) l::'a list. +lemma FOLDL_FILTER: "ALL (f::'a::type => 'a::type => 'a::type) (e::'a::type) + (P::'a::type => bool) l::'a::type list. foldl f e (filter P l) = - foldl (%(x::'a) y::'a. if P y then f x y else x) e l" + foldl (%(x::'a::type) y::'a::type. if P y then f x y else x) e l" by (import rich_list FOLDL_FILTER) -lemma ASSOC_FOLDR_FLAT: "ALL f. +lemma ASSOC_FOLDR_FLAT: "ALL f::'a::type => 'a::type => 'a::type. ASSOC f --> - (ALL e. + (ALL e::'a::type. LEFT_ID f e --> - (ALL l. foldr f (concat l) e = foldr f (map (FOLDR f e) l) e))" + (ALL l::'a::type list list. + foldr f (concat l) e = foldr f (map (FOLDR f e) l) e))" by (import rich_list ASSOC_FOLDR_FLAT) -lemma ASSOC_FOLDL_FLAT: "ALL f. +lemma ASSOC_FOLDL_FLAT: "ALL f::'a::type => 'a::type => 'a::type. ASSOC f --> - (ALL e. + (ALL e::'a::type. RIGHT_ID f e --> - (ALL l. foldl f e (concat l) = foldl f e (map (foldl f e) l)))" + (ALL l::'a::type list list. + foldl f e (concat l) = foldl f e (map (foldl f e) l)))" by (import rich_list ASSOC_FOLDL_FLAT) -lemma SOME_EL_MAP: "ALL (P::'b => bool) (f::'a => 'b) l::'a list. +lemma SOME_EL_MAP: "ALL (P::'b::type => bool) (f::'a::type => 'b::type) l::'a::type list. list_exists P (map f l) = list_exists (P o f) l" by (import rich_list SOME_EL_MAP) -lemma SOME_EL_DISJ: "ALL P Q l. - list_exists (%x. P x | Q x) l = (list_exists P l | list_exists Q l)" +lemma SOME_EL_DISJ: "ALL (P::'a::type => bool) (Q::'a::type => bool) l::'a::type list. + list_exists (%x::'a::type. P x | Q x) l = + (list_exists P l | list_exists Q l)" by (import rich_list SOME_EL_DISJ) -lemma IS_EL_FOLDR: "ALL x xa. x mem xa = foldr (%xa. op | (x = xa)) xa False" +lemma IS_EL_FOLDR: "ALL (x::'a::type) xa::'a::type list. + x mem xa = foldr (%xa::'a::type. op | (x = xa)) xa False" by (import rich_list IS_EL_FOLDR) -lemma IS_EL_FOLDL: "ALL x xa. x mem xa = foldl (%l' xa. l' | x = xa) False xa" +lemma IS_EL_FOLDL: "ALL (x::'a::type) xa::'a::type list. + x mem xa = foldl (%(l'::bool) xa::'a::type. l' | x = xa) False xa" by (import rich_list IS_EL_FOLDL) -lemma NULL_FOLDR: "ALL l. null l = foldr (%x l'. False) l True" +lemma NULL_FOLDR: "ALL l::'a::type list. null l = foldr (%(x::'a::type) l'::bool. False) l True" by (import rich_list NULL_FOLDR) -lemma NULL_FOLDL: "ALL l. null l = foldl (%x l'. False) True l" +lemma NULL_FOLDL: "ALL l::'a::type list. null l = foldl (%(x::bool) l'::'a::type. False) True l" by (import rich_list NULL_FOLDL) -lemma SEG_LENGTH_ID: "ALL l. SEG (length l) 0 l = l" +lemma SEG_LENGTH_ID: "ALL l::'a::type list. SEG (length l) (0::nat) l = l" by (import rich_list SEG_LENGTH_ID) -lemma SEG_SUC_CONS: "ALL m n l x. SEG m (Suc n) (x # l) = SEG m n l" +lemma SEG_SUC_CONS: "ALL (m::nat) (n::nat) (l::'a::type list) x::'a::type. + SEG m (Suc n) (x # l) = SEG m n l" by (import rich_list SEG_SUC_CONS) -lemma SEG_0_SNOC: "ALL m l x. m <= length l --> SEG m 0 (SNOC x l) = SEG m 0 l" +lemma SEG_0_SNOC: "ALL (m::nat) (l::'a::type list) x::'a::type. + m <= length l --> SEG m (0::nat) (SNOC x l) = SEG m (0::nat) l" by (import rich_list SEG_0_SNOC) -lemma BUTLASTN_SEG: "ALL n l. n <= length l --> BUTLASTN n l = SEG (length l - n) 0 l" +lemma BUTLASTN_SEG: "ALL (n::nat) l::'a::type list. + n <= length l --> BUTLASTN n l = SEG (length l - n) (0::nat) l" by (import rich_list BUTLASTN_SEG) -lemma LASTN_CONS: "ALL n l. n <= length l --> (ALL x. LASTN n (x # l) = LASTN n l)" +lemma LASTN_CONS: "ALL (n::nat) l::'a::type list. + n <= length l --> (ALL x::'a::type. LASTN n (x # l) = LASTN n l)" by (import rich_list LASTN_CONS) -lemma LENGTH_LASTN: "ALL n l. n <= length l --> length (LASTN n l) = n" +lemma LENGTH_LASTN: "ALL (n::nat) l::'a::type list. n <= length l --> length (LASTN n l) = n" by (import rich_list LENGTH_LASTN) -lemma LASTN_LENGTH_ID: "ALL l. LASTN (length l) l = l" +lemma LASTN_LENGTH_ID: "ALL l::'a::type list. LASTN (length l) l = l" by (import rich_list LASTN_LENGTH_ID) -lemma LASTN_LASTN: "ALL l n m. m <= length l --> n <= m --> LASTN n (LASTN m l) = LASTN n l" +lemma LASTN_LASTN: "ALL (l::'a::type list) (n::nat) m::nat. + m <= length l --> n <= m --> LASTN n (LASTN m l) = LASTN n l" by (import rich_list LASTN_LASTN) -lemma FIRSTN_LENGTH_ID: "ALL l. FIRSTN (length l) l = l" +lemma FIRSTN_LENGTH_ID: "ALL l::'a::type list. FIRSTN (length l) l = l" by (import rich_list FIRSTN_LENGTH_ID) -lemma FIRSTN_SNOC: "ALL n l. n <= length l --> (ALL x. FIRSTN n (SNOC x l) = FIRSTN n l)" +lemma FIRSTN_SNOC: "ALL (n::nat) l::'a::type list. + n <= length l --> (ALL x::'a::type. FIRSTN n (SNOC x l) = FIRSTN n l)" by (import rich_list FIRSTN_SNOC) -lemma BUTLASTN_LENGTH_NIL: "ALL l. BUTLASTN (length l) l = []" +lemma BUTLASTN_LENGTH_NIL: "ALL l::'a::type list. BUTLASTN (length l) l = []" by (import rich_list BUTLASTN_LENGTH_NIL) -lemma BUTLASTN_SUC_BUTLAST: "ALL n l. n < length l --> BUTLASTN (Suc n) l = BUTLASTN n (butlast l)" +lemma BUTLASTN_SUC_BUTLAST: "ALL (n::nat) l::'a::type list. + n < length l --> BUTLASTN (Suc n) l = BUTLASTN n (butlast l)" by (import rich_list BUTLASTN_SUC_BUTLAST) -lemma BUTLASTN_BUTLAST: "ALL n l. n < length l --> BUTLASTN n (butlast l) = butlast (BUTLASTN n l)" +lemma BUTLASTN_BUTLAST: "ALL (n::nat) l::'a::type list. + n < length l --> BUTLASTN n (butlast l) = butlast (BUTLASTN n l)" by (import rich_list BUTLASTN_BUTLAST) -lemma LENGTH_BUTLASTN: "ALL n l. n <= length l --> length (BUTLASTN n l) = length l - n" +lemma LENGTH_BUTLASTN: "ALL (n::nat) l::'a::type list. + n <= length l --> length (BUTLASTN n l) = length l - n" by (import rich_list LENGTH_BUTLASTN) -lemma BUTLASTN_BUTLASTN: "ALL m n l. +lemma BUTLASTN_BUTLASTN: "ALL (m::nat) (n::nat) l::'a::type list. n + m <= length l --> BUTLASTN n (BUTLASTN m l) = BUTLASTN (n + m) l" by (import rich_list BUTLASTN_BUTLASTN) -lemma APPEND_BUTLASTN_LASTN: "ALL n l. n <= length l --> BUTLASTN n l @ LASTN n l = l" +lemma APPEND_BUTLASTN_LASTN: "ALL (n::nat) l::'a::type list. + n <= length l --> BUTLASTN n l @ LASTN n l = l" by (import rich_list APPEND_BUTLASTN_LASTN) -lemma APPEND_FIRSTN_LASTN: "ALL m n l. m + n = length l --> FIRSTN n l @ LASTN m l = l" +lemma APPEND_FIRSTN_LASTN: "ALL (m::nat) (n::nat) l::'a::type list. + m + n = length l --> FIRSTN n l @ LASTN m l = l" by (import rich_list APPEND_FIRSTN_LASTN) -lemma BUTLASTN_APPEND2: "ALL n l1 l2. n <= length l2 --> BUTLASTN n (l1 @ l2) = l1 @ BUTLASTN n l2" +lemma BUTLASTN_APPEND2: "ALL (n::nat) (l1::'a::type list) l2::'a::type list. + n <= length l2 --> BUTLASTN n (l1 @ l2) = l1 @ BUTLASTN n l2" by (import rich_list BUTLASTN_APPEND2) -lemma BUTLASTN_LENGTH_APPEND: "ALL l2 l1. BUTLASTN (length l2) (l1 @ l2) = l1" +lemma BUTLASTN_LENGTH_APPEND: "ALL (l2::'a::type list) l1::'a::type list. + BUTLASTN (length l2) (l1 @ l2) = l1" by (import rich_list BUTLASTN_LENGTH_APPEND) -lemma LASTN_LENGTH_APPEND: "ALL l2 l1. LASTN (length l2) (l1 @ l2) = l2" +lemma LASTN_LENGTH_APPEND: "ALL (l2::'a::type list) l1::'a::type list. LASTN (length l2) (l1 @ l2) = l2" by (import rich_list LASTN_LENGTH_APPEND) -lemma BUTLASTN_CONS: "ALL n l. n <= length l --> (ALL x. BUTLASTN n (x # l) = x # BUTLASTN n l)" +lemma BUTLASTN_CONS: "ALL (n::nat) l::'a::type list. + n <= length l --> + (ALL x::'a::type. BUTLASTN n (x # l) = x # BUTLASTN n l)" by (import rich_list BUTLASTN_CONS) -lemma BUTLASTN_LENGTH_CONS: "ALL l x. BUTLASTN (length l) (x # l) = [x]" +lemma BUTLASTN_LENGTH_CONS: "ALL (l::'a::type list) x::'a::type. BUTLASTN (length l) (x # l) = [x]" by (import rich_list BUTLASTN_LENGTH_CONS) -lemma LAST_LASTN_LAST: "ALL n l. n <= length l --> 0 < n --> last (LASTN n l) = last l" +lemma LAST_LASTN_LAST: "ALL (n::nat) l::'a::type list. + n <= length l --> (0::nat) < n --> last (LASTN n l) = last l" by (import rich_list LAST_LASTN_LAST) -lemma BUTLASTN_LASTN_NIL: "ALL n l. n <= length l --> BUTLASTN n (LASTN n l) = []" +lemma BUTLASTN_LASTN_NIL: "ALL (n::nat) l::'a::type list. n <= length l --> BUTLASTN n (LASTN n l) = []" by (import rich_list BUTLASTN_LASTN_NIL) -lemma LASTN_BUTLASTN: "ALL n m l. +lemma LASTN_BUTLASTN: "ALL (n::nat) (m::nat) l::'a::type list. n + m <= length l --> LASTN n (BUTLASTN m l) = BUTLASTN m (LASTN (n + m) l)" by (import rich_list LASTN_BUTLASTN) -lemma BUTLASTN_LASTN: "ALL m n l. +lemma BUTLASTN_LASTN: "ALL (m::nat) (n::nat) l::'a::type list. m <= n & n <= length l --> BUTLASTN m (LASTN n l) = LASTN (n - m) (BUTLASTN m l)" by (import rich_list BUTLASTN_LASTN) -lemma LASTN_1: "ALL l. l ~= [] --> LASTN 1 l = [last l]" +lemma LASTN_1: "ALL l::'a::type list. l ~= [] --> LASTN (1::nat) l = [last l]" by (import rich_list LASTN_1) -lemma BUTLASTN_1: "ALL l. l ~= [] --> BUTLASTN 1 l = butlast l" +lemma BUTLASTN_1: "ALL l::'a::type list. l ~= [] --> BUTLASTN (1::nat) l = butlast l" by (import rich_list BUTLASTN_1) -lemma BUTLASTN_APPEND1: "ALL l2 n. +lemma BUTLASTN_APPEND1: "ALL (l2::'a::type list) n::nat. length l2 <= n --> - (ALL l1. BUTLASTN n (l1 @ l2) = BUTLASTN (n - length l2) l1)" + (ALL l1::'a::type list. + BUTLASTN n (l1 @ l2) = BUTLASTN (n - length l2) l1)" by (import rich_list BUTLASTN_APPEND1) -lemma LASTN_APPEND2: "ALL n l2. n <= length l2 --> (ALL l1. LASTN n (l1 @ l2) = LASTN n l2)" +lemma LASTN_APPEND2: "ALL (n::nat) l2::'a::type list. + n <= length l2 --> + (ALL l1::'a::type list. LASTN n (l1 @ l2) = LASTN n l2)" by (import rich_list LASTN_APPEND2) -lemma LASTN_APPEND1: "ALL l2 n. +lemma LASTN_APPEND1: "ALL (l2::'a::type list) n::nat. length l2 <= n --> - (ALL l1. LASTN n (l1 @ l2) = LASTN (n - length l2) l1 @ l2)" + (ALL l1::'a::type list. + LASTN n (l1 @ l2) = LASTN (n - length l2) l1 @ l2)" by (import rich_list LASTN_APPEND1) -lemma LASTN_MAP: "ALL n l. n <= length l --> (ALL f. LASTN n (map f l) = map f (LASTN n l))" +lemma LASTN_MAP: "ALL (n::nat) l::'a::type list. + n <= length l --> + (ALL f::'a::type => 'b::type. LASTN n (map f l) = map f (LASTN n l))" by (import rich_list LASTN_MAP) -lemma BUTLASTN_MAP: "ALL n l. - n <= length l --> (ALL f. BUTLASTN n (map f l) = map f (BUTLASTN n l))" +lemma BUTLASTN_MAP: "ALL (n::nat) l::'a::type list. + n <= length l --> + (ALL f::'a::type => 'b::type. + BUTLASTN n (map f l) = map f (BUTLASTN n l))" by (import rich_list BUTLASTN_MAP) -lemma ALL_EL_LASTN: "(All::(('a => bool) => bool) => bool) - (%P::'a => bool. - (All::('a list => bool) => bool) - (%l::'a list. +lemma ALL_EL_LASTN: "(All::(('a::type => bool) => bool) => bool) + (%P::'a::type => bool. + (All::('a::type list => bool) => bool) + (%l::'a::type list. (op -->::bool => bool => bool) - ((list_all::('a => bool) => 'a list => bool) P l) + ((list_all::('a::type => bool) => 'a::type list => bool) P l) ((All::(nat => bool) => bool) (%m::nat. (op -->::bool => bool => bool) - ((op <=::nat => nat => bool) m ((size::'a list => nat) l)) - ((list_all::('a => bool) => 'a list => bool) P - ((LASTN::nat => 'a list => 'a list) m l))))))" + ((op <=::nat => nat => bool) m + ((size::'a::type list => nat) l)) + ((list_all::('a::type => bool) => 'a::type list => bool) P + ((LASTN::nat => 'a::type list => 'a::type list) m + l))))))" by (import rich_list ALL_EL_LASTN) -lemma ALL_EL_BUTLASTN: "(All::(('a => bool) => bool) => bool) - (%P::'a => bool. - (All::('a list => bool) => bool) - (%l::'a list. +lemma ALL_EL_BUTLASTN: "(All::(('a::type => bool) => bool) => bool) + (%P::'a::type => bool. + (All::('a::type list => bool) => bool) + (%l::'a::type list. (op -->::bool => bool => bool) - ((list_all::('a => bool) => 'a list => bool) P l) + ((list_all::('a::type => bool) => 'a::type list => bool) P l) ((All::(nat => bool) => bool) (%m::nat. (op -->::bool => bool => bool) - ((op <=::nat => nat => bool) m ((size::'a list => nat) l)) - ((list_all::('a => bool) => 'a list => bool) P - ((BUTLASTN::nat => 'a list => 'a list) m l))))))" + ((op <=::nat => nat => bool) m + ((size::'a::type list => nat) l)) + ((list_all::('a::type => bool) => 'a::type list => bool) P + ((BUTLASTN::nat => 'a::type list => 'a::type list) m + l))))))" by (import rich_list ALL_EL_BUTLASTN) -lemma LENGTH_FIRSTN: "ALL n l. n <= length l --> length (FIRSTN n l) = n" +lemma LENGTH_FIRSTN: "ALL (n::nat) l::'a::type list. n <= length l --> length (FIRSTN n l) = n" by (import rich_list LENGTH_FIRSTN) lemma FIRSTN_FIRSTN: "(All::(nat => bool) => bool) (%m::nat. - (All::('a list => bool) => bool) - (%l::'a list. + (All::('a::type list => bool) => bool) + (%l::'a::type list. (op -->::bool => bool => bool) - ((op <=::nat => nat => bool) m ((size::'a list => nat) l)) + ((op <=::nat => nat => bool) m ((size::'a::type list => nat) l)) ((All::(nat => bool) => bool) (%n::nat. (op -->::bool => bool => bool) ((op <=::nat => nat => bool) n m) - ((op =::'a list => 'a list => bool) - ((FIRSTN::nat => 'a list => 'a list) n - ((FIRSTN::nat => 'a list => 'a list) m l)) - ((FIRSTN::nat => 'a list => 'a list) n l))))))" + ((op =::'a::type list => 'a::type list => bool) + ((FIRSTN::nat => 'a::type list => 'a::type list) n + ((FIRSTN::nat => 'a::type list => 'a::type list) m l)) + ((FIRSTN::nat => 'a::type list => 'a::type list) n + l))))))" by (import rich_list FIRSTN_FIRSTN) -lemma LENGTH_BUTFIRSTN: "ALL n l. n <= length l --> length (BUTFIRSTN n l) = length l - n" +lemma LENGTH_BUTFIRSTN: "ALL (n::nat) l::'a::type list. + n <= length l --> length (BUTFIRSTN n l) = length l - n" by (import rich_list LENGTH_BUTFIRSTN) -lemma BUTFIRSTN_LENGTH_NIL: "ALL l. BUTFIRSTN (length l) l = []" +lemma BUTFIRSTN_LENGTH_NIL: "ALL l::'a::type list. BUTFIRSTN (length l) l = []" by (import rich_list BUTFIRSTN_LENGTH_NIL) -lemma BUTFIRSTN_APPEND1: "ALL n l1. - n <= length l1 --> (ALL l2. BUTFIRSTN n (l1 @ l2) = BUTFIRSTN n l1 @ l2)" +lemma BUTFIRSTN_APPEND1: "ALL (n::nat) l1::'a::type list. + n <= length l1 --> + (ALL l2::'a::type list. BUTFIRSTN n (l1 @ l2) = BUTFIRSTN n l1 @ l2)" by (import rich_list BUTFIRSTN_APPEND1) -lemma BUTFIRSTN_APPEND2: "ALL l1 n. +lemma BUTFIRSTN_APPEND2: "ALL (l1::'a::type list) n::nat. length l1 <= n --> - (ALL l2. BUTFIRSTN n (l1 @ l2) = BUTFIRSTN (n - length l1) l2)" + (ALL l2::'a::type list. + BUTFIRSTN n (l1 @ l2) = BUTFIRSTN (n - length l1) l2)" by (import rich_list BUTFIRSTN_APPEND2) -lemma BUTFIRSTN_BUTFIRSTN: "ALL n m l. +lemma BUTFIRSTN_BUTFIRSTN: "ALL (n::nat) (m::nat) l::'a::type list. n + m <= length l --> BUTFIRSTN n (BUTFIRSTN m l) = BUTFIRSTN (n + m) l" by (import rich_list BUTFIRSTN_BUTFIRSTN) -lemma APPEND_FIRSTN_BUTFIRSTN: "ALL n l. n <= length l --> FIRSTN n l @ BUTFIRSTN n l = l" +lemma APPEND_FIRSTN_BUTFIRSTN: "ALL (n::nat) l::'a::type list. + n <= length l --> FIRSTN n l @ BUTFIRSTN n l = l" by (import rich_list APPEND_FIRSTN_BUTFIRSTN) -lemma LASTN_SEG: "ALL n l. n <= length l --> LASTN n l = SEG n (length l - n) l" +lemma LASTN_SEG: "ALL (n::nat) l::'a::type list. + n <= length l --> LASTN n l = SEG n (length l - n) l" by (import rich_list LASTN_SEG) -lemma FIRSTN_SEG: "ALL n l. n <= length l --> FIRSTN n l = SEG n 0 l" +lemma FIRSTN_SEG: "ALL (n::nat) l::'a::type list. + n <= length l --> FIRSTN n l = SEG n (0::nat) l" by (import rich_list FIRSTN_SEG) -lemma BUTFIRSTN_SEG: "ALL n l. n <= length l --> BUTFIRSTN n l = SEG (length l - n) n l" +lemma BUTFIRSTN_SEG: "ALL (n::nat) l::'a::type list. + n <= length l --> BUTFIRSTN n l = SEG (length l - n) n l" by (import rich_list BUTFIRSTN_SEG) -lemma BUTFIRSTN_SNOC: "ALL n l. +lemma BUTFIRSTN_SNOC: "ALL (n::nat) l::'a::type list. n <= length l --> - (ALL x. BUTFIRSTN n (SNOC x l) = SNOC x (BUTFIRSTN n l))" + (ALL x::'a::type. BUTFIRSTN n (SNOC x l) = SNOC x (BUTFIRSTN n l))" by (import rich_list BUTFIRSTN_SNOC) -lemma APPEND_BUTLASTN_BUTFIRSTN: "ALL m n l. m + n = length l --> BUTLASTN m l @ BUTFIRSTN n l = l" +lemma APPEND_BUTLASTN_BUTFIRSTN: "ALL (m::nat) (n::nat) l::'a::type list. + m + n = length l --> BUTLASTN m l @ BUTFIRSTN n l = l" by (import rich_list APPEND_BUTLASTN_BUTFIRSTN) -lemma SEG_SEG: "ALL n1 m1 n2 m2 l. +lemma SEG_SEG: "ALL (n1::nat) (m1::nat) (n2::nat) (m2::nat) l::'a::type list. n1 + m1 <= length l & n2 + m2 <= n1 --> SEG n2 m2 (SEG n1 m1 l) = SEG n2 (m1 + m2) l" by (import rich_list SEG_SEG) -lemma SEG_APPEND1: "ALL n m l1. n + m <= length l1 --> (ALL l2. SEG n m (l1 @ l2) = SEG n m l1)" +lemma SEG_APPEND1: "ALL (n::nat) (m::nat) l1::'a::type list. + n + m <= length l1 --> + (ALL l2::'a::type list. SEG n m (l1 @ l2) = SEG n m l1)" by (import rich_list SEG_APPEND1) -lemma SEG_APPEND2: "ALL l1 m n l2. +lemma SEG_APPEND2: "ALL (l1::'a::type list) (m::nat) (n::nat) l2::'a::type list. length l1 <= m & n <= length l2 --> SEG n m (l1 @ l2) = SEG n (m - length l1) l2" by (import rich_list SEG_APPEND2) -lemma SEG_FIRSTN_BUTFISTN: "ALL n m l. n + m <= length l --> SEG n m l = FIRSTN n (BUTFIRSTN m l)" +lemma SEG_FIRSTN_BUTFISTN: "ALL (n::nat) (m::nat) l::'a::type list. + n + m <= length l --> SEG n m l = FIRSTN n (BUTFIRSTN m l)" by (import rich_list SEG_FIRSTN_BUTFISTN) -lemma SEG_APPEND: "ALL m l1 n l2. +lemma SEG_APPEND: "ALL (m::nat) (l1::'a::type list) (n::nat) l2::'a::type list. m < length l1 & length l1 <= n + m & n + m <= length l1 + length l2 --> SEG n m (l1 @ l2) = - SEG (length l1 - m) m l1 @ SEG (n + m - length l1) 0 l2" + SEG (length l1 - m) m l1 @ SEG (n + m - length l1) (0::nat) l2" by (import rich_list SEG_APPEND) -lemma SEG_LENGTH_SNOC: "ALL x xa. SEG 1 (length x) (SNOC xa x) = [xa]" +lemma SEG_LENGTH_SNOC: "ALL (x::'a::type list) xa::'a::type. + SEG (1::nat) (length x) (SNOC xa x) = [xa]" by (import rich_list SEG_LENGTH_SNOC) -lemma SEG_SNOC: "ALL n m l. n + m <= length l --> (ALL x. SEG n m (SNOC x l) = SEG n m l)" +lemma SEG_SNOC: "ALL (n::nat) (m::nat) l::'a::type list. + n + m <= length l --> (ALL x::'a::type. SEG n m (SNOC x l) = SEG n m l)" by (import rich_list SEG_SNOC) -lemma ELL_SEG: "ALL n l. n < length l --> ELL n l = hd (SEG 1 (PRE (length l - n)) l)" +lemma ELL_SEG: "ALL (n::nat) l::'a::type list. + n < length l --> ELL n l = hd (SEG (1::nat) (PRE (length l - n)) l)" by (import rich_list ELL_SEG) -lemma SNOC_FOLDR: "ALL x l. SNOC x l = foldr op # l [x]" +lemma SNOC_FOLDR: "ALL (x::'a::type) l::'a::type list. SNOC x l = foldr op # l [x]" by (import rich_list SNOC_FOLDR) -lemma IS_EL_FOLDR_MAP: "ALL x xa. x mem xa = foldr op | (map (op = x) xa) False" +lemma IS_EL_FOLDR_MAP: "ALL (x::'a::type) xa::'a::type list. + x mem xa = foldr op | (map (op = x) xa) False" by (import rich_list IS_EL_FOLDR_MAP) -lemma IS_EL_FOLDL_MAP: "ALL x xa. x mem xa = foldl op | False (map (op = x) xa)" +lemma IS_EL_FOLDL_MAP: "ALL (x::'a::type) xa::'a::type list. + x mem xa = foldl op | False (map (op = x) xa)" by (import rich_list IS_EL_FOLDL_MAP) -lemma FILTER_FILTER: "ALL P Q l. filter P (filter Q l) = [x:l. P x & Q x]" +lemma FILTER_FILTER: "ALL (P::'a::type => bool) (Q::'a::type => bool) l::'a::type list. + filter P (filter Q l) = [x::'a::type:l. P x & Q x]" by (import rich_list FILTER_FILTER) -lemma FCOMM_FOLDR_FLAT: "ALL g f. +lemma FCOMM_FOLDR_FLAT: "ALL (g::'a::type => 'a::type => 'a::type) + f::'b::type => 'a::type => 'a::type. FCOMM g f --> - (ALL e. + (ALL e::'a::type. LEFT_ID g e --> - (ALL l. foldr f (concat l) e = foldr g (map (FOLDR f e) l) e))" + (ALL l::'b::type list list. + foldr f (concat l) e = foldr g (map (FOLDR f e) l) e))" by (import rich_list FCOMM_FOLDR_FLAT) -lemma FCOMM_FOLDL_FLAT: "ALL f g. +lemma FCOMM_FOLDL_FLAT: "ALL (f::'a::type => 'b::type => 'a::type) + g::'a::type => 'a::type => 'a::type. FCOMM f g --> - (ALL e. + (ALL e::'a::type. RIGHT_ID g e --> - (ALL l. foldl f e (concat l) = foldl g e (map (foldl f e) l)))" + (ALL l::'b::type list list. + foldl f e (concat l) = foldl g e (map (foldl f e) l)))" by (import rich_list FCOMM_FOLDL_FLAT) -lemma FOLDR_MAP_REVERSE: "ALL f::'a => 'a => 'a. - (ALL (a::'a) (b::'a) c::'a. f a (f b c) = f b (f a c)) --> - (ALL (e::'a) (g::'b => 'a) l::'b list. +lemma FOLDR_MAP_REVERSE: "ALL f::'a::type => 'a::type => 'a::type. + (ALL (a::'a::type) (b::'a::type) c::'a::type. + f a (f b c) = f b (f a c)) --> + (ALL (e::'a::type) (g::'b::type => 'a::type) l::'b::type list. foldr f (map g (rev l)) e = foldr f (map g l) e)" by (import rich_list FOLDR_MAP_REVERSE) -lemma FOLDR_FILTER_REVERSE: "ALL f::'a => 'a => 'a. - (ALL (a::'a) (b::'a) c::'a. f a (f b c) = f b (f a c)) --> - (ALL (e::'a) (P::'a => bool) l::'a list. +lemma FOLDR_FILTER_REVERSE: "ALL f::'a::type => 'a::type => 'a::type. + (ALL (a::'a::type) (b::'a::type) c::'a::type. + f a (f b c) = f b (f a c)) --> + (ALL (e::'a::type) (P::'a::type => bool) l::'a::type list. foldr f (filter P (rev l)) e = foldr f (filter P l) e)" by (import rich_list FOLDR_FILTER_REVERSE) -lemma COMM_ASSOC_FOLDR_REVERSE: "ALL f. COMM f --> ASSOC f --> (ALL e l. foldr f (rev l) e = foldr f l e)" +lemma COMM_ASSOC_FOLDR_REVERSE: "ALL f::'a::type => 'a::type => 'a::type. + COMM f --> + ASSOC f --> + (ALL (e::'a::type) l::'a::type list. foldr f (rev l) e = foldr f l e)" by (import rich_list COMM_ASSOC_FOLDR_REVERSE) -lemma COMM_ASSOC_FOLDL_REVERSE: "ALL f. COMM f --> ASSOC f --> (ALL e l. foldl f e (rev l) = foldl f e l)" +lemma COMM_ASSOC_FOLDL_REVERSE: "ALL f::'a::type => 'a::type => 'a::type. + COMM f --> + ASSOC f --> + (ALL (e::'a::type) l::'a::type list. foldl f e (rev l) = foldl f e l)" by (import rich_list COMM_ASSOC_FOLDL_REVERSE) -lemma ELL_LAST: "ALL l. ~ null l --> ELL 0 l = last l" +lemma ELL_LAST: "ALL l::'a::type list. ~ null l --> ELL (0::nat) l = last l" by (import rich_list ELL_LAST) -lemma ELL_0_SNOC: "ALL l x. ELL 0 (SNOC x l) = x" +lemma ELL_0_SNOC: "ALL (l::'a::type list) x::'a::type. ELL (0::nat) (SNOC x l) = x" by (import rich_list ELL_0_SNOC) -lemma ELL_SNOC: "ALL n>0. ALL x l. ELL n (SNOC x l) = ELL (PRE n) l" +lemma ELL_SNOC: "ALL n>0::nat. + ALL (x::'a::type) l::'a::type list. ELL n (SNOC x l) = ELL (PRE n) l" by (import rich_list ELL_SNOC) -lemma ELL_SUC_SNOC: "ALL n x xa. ELL (Suc n) (SNOC x xa) = ELL n xa" +lemma ELL_SUC_SNOC: "ALL (n::nat) (x::'a::type) xa::'a::type list. + ELL (Suc n) (SNOC x xa) = ELL n xa" by (import rich_list ELL_SUC_SNOC) -lemma ELL_CONS: "ALL n l. n < length l --> (ALL x. ELL n (x # l) = ELL n l)" +lemma ELL_CONS: "ALL (n::nat) l::'a::type list. + n < length l --> (ALL x::'a::type. ELL n (x # l) = ELL n l)" by (import rich_list ELL_CONS) -lemma ELL_LENGTH_CONS: "ALL l x. ELL (length l) (x # l) = x" +lemma ELL_LENGTH_CONS: "ALL (l::'a::type list) x::'a::type. ELL (length l) (x # l) = x" by (import rich_list ELL_LENGTH_CONS) -lemma ELL_LENGTH_SNOC: "ALL l x. ELL (length l) (SNOC x l) = (if null l then x else hd l)" +lemma ELL_LENGTH_SNOC: "ALL (l::'a::type list) x::'a::type. + ELL (length l) (SNOC x l) = (if null l then x else hd l)" by (import rich_list ELL_LENGTH_SNOC) -lemma ELL_APPEND2: "ALL n l2. n < length l2 --> (ALL l1. ELL n (l1 @ l2) = ELL n l2)" +lemma ELL_APPEND2: "ALL (n::nat) l2::'a::type list. + n < length l2 --> (ALL l1::'a::type list. ELL n (l1 @ l2) = ELL n l2)" by (import rich_list ELL_APPEND2) -lemma ELL_APPEND1: "ALL l2 n. - length l2 <= n --> (ALL l1. ELL n (l1 @ l2) = ELL (n - length l2) l1)" +lemma ELL_APPEND1: "ALL (l2::'a::type list) n::nat. + length l2 <= n --> + (ALL l1::'a::type list. ELL n (l1 @ l2) = ELL (n - length l2) l1)" by (import rich_list ELL_APPEND1) -lemma ELL_PRE_LENGTH: "ALL l. l ~= [] --> ELL (PRE (length l)) l = hd l" +lemma ELL_PRE_LENGTH: "ALL l::'a::type list. l ~= [] --> ELL (PRE (length l)) l = hd l" by (import rich_list ELL_PRE_LENGTH) -lemma EL_LENGTH_SNOC: "ALL l x. EL (length l) (SNOC x l) = x" +lemma EL_LENGTH_SNOC: "ALL (l::'a::type list) x::'a::type. EL (length l) (SNOC x l) = x" by (import rich_list EL_LENGTH_SNOC) -lemma EL_PRE_LENGTH: "ALL l. l ~= [] --> EL (PRE (length l)) l = last l" +lemma EL_PRE_LENGTH: "ALL l::'a::type list. l ~= [] --> EL (PRE (length l)) l = last l" by (import rich_list EL_PRE_LENGTH) -lemma EL_SNOC: "ALL n l. n < length l --> (ALL x. EL n (SNOC x l) = EL n l)" +lemma EL_SNOC: "ALL (n::nat) l::'a::type list. + n < length l --> (ALL x::'a::type. EL n (SNOC x l) = EL n l)" by (import rich_list EL_SNOC) -lemma EL_ELL: "ALL n l. n < length l --> EL n l = ELL (PRE (length l - n)) l" +lemma EL_ELL: "ALL (n::nat) l::'a::type list. + n < length l --> EL n l = ELL (PRE (length l - n)) l" by (import rich_list EL_ELL) -lemma EL_LENGTH_APPEND: "ALL l2 l1. ~ null l2 --> EL (length l1) (l1 @ l2) = hd l2" +lemma EL_LENGTH_APPEND: "ALL (l2::'a::type list) l1::'a::type list. + ~ null l2 --> EL (length l1) (l1 @ l2) = hd l2" by (import rich_list EL_LENGTH_APPEND) -lemma ELL_EL: "ALL n l. n < length l --> ELL n l = EL (PRE (length l - n)) l" +lemma ELL_EL: "ALL (n::nat) l::'a::type list. + n < length l --> ELL n l = EL (PRE (length l - n)) l" by (import rich_list ELL_EL) -lemma ELL_MAP: "ALL n l f. n < length l --> ELL n (map f l) = f (ELL n l)" +lemma ELL_MAP: "ALL (n::nat) (l::'a::type list) f::'a::type => 'b::type. + n < length l --> ELL n (map f l) = f (ELL n l)" by (import rich_list ELL_MAP) -lemma LENGTH_BUTLAST: "ALL l. l ~= [] --> length (butlast l) = PRE (length l)" +lemma LENGTH_BUTLAST: "ALL l::'a::type list. l ~= [] --> length (butlast l) = PRE (length l)" by (import rich_list LENGTH_BUTLAST) -lemma BUTFIRSTN_LENGTH_APPEND: "ALL l1 l2. BUTFIRSTN (length l1) (l1 @ l2) = l2" +lemma BUTFIRSTN_LENGTH_APPEND: "ALL (l1::'a::type list) l2::'a::type list. + BUTFIRSTN (length l1) (l1 @ l2) = l2" by (import rich_list BUTFIRSTN_LENGTH_APPEND) -lemma FIRSTN_APPEND1: "ALL n l1. n <= length l1 --> (ALL l2. FIRSTN n (l1 @ l2) = FIRSTN n l1)" +lemma FIRSTN_APPEND1: "ALL (n::nat) l1::'a::type list. + n <= length l1 --> + (ALL l2::'a::type list. FIRSTN n (l1 @ l2) = FIRSTN n l1)" by (import rich_list FIRSTN_APPEND1) -lemma FIRSTN_APPEND2: "ALL l1 n. +lemma FIRSTN_APPEND2: "ALL (l1::'a::type list) n::nat. length l1 <= n --> - (ALL l2. FIRSTN n (l1 @ l2) = l1 @ FIRSTN (n - length l1) l2)" + (ALL l2::'a::type list. + FIRSTN n (l1 @ l2) = l1 @ FIRSTN (n - length l1) l2)" by (import rich_list FIRSTN_APPEND2) -lemma FIRSTN_LENGTH_APPEND: "ALL l1 l2. FIRSTN (length l1) (l1 @ l2) = l1" +lemma FIRSTN_LENGTH_APPEND: "ALL (l1::'a::type list) l2::'a::type list. FIRSTN (length l1) (l1 @ l2) = l1" by (import rich_list FIRSTN_LENGTH_APPEND) -lemma REVERSE_FLAT: "ALL l. rev (concat l) = concat (rev (map rev l))" +lemma REVERSE_FLAT: "ALL l::'a::type list list. rev (concat l) = concat (rev (map rev l))" by (import rich_list REVERSE_FLAT) -lemma MAP_FILTER: "ALL f P l. - (ALL x. P (f x) = P x) --> map f (filter P l) = filter P (map f l)" +lemma MAP_FILTER: "ALL (f::'a::type => 'a::type) (P::'a::type => bool) l::'a::type list. + (ALL x::'a::type. P (f x) = P x) --> + map f (filter P l) = filter P (map f l)" by (import rich_list MAP_FILTER) -lemma FLAT_REVERSE: "ALL l. concat (rev l) = rev (concat (map rev l))" +lemma FLAT_REVERSE: "ALL l::'a::type list list. concat (rev l) = rev (concat (map rev l))" by (import rich_list FLAT_REVERSE) -lemma FLAT_FLAT: "ALL l. concat (concat l) = concat (map concat l)" +lemma FLAT_FLAT: "ALL l::'a::type list list list. concat (concat l) = concat (map concat l)" by (import rich_list FLAT_FLAT) -lemma SOME_EL_REVERSE: "ALL P l. list_exists P (rev l) = list_exists P l" +lemma SOME_EL_REVERSE: "ALL (P::'a::type => bool) l::'a::type list. + list_exists P (rev l) = list_exists P l" by (import rich_list SOME_EL_REVERSE) -lemma ALL_EL_SEG: "ALL P l. - list_all P l --> (ALL m k. m + k <= length l --> list_all P (SEG m k l))" +lemma ALL_EL_SEG: "ALL (P::'a::type => bool) l::'a::type list. + list_all P l --> + (ALL (m::nat) k::nat. m + k <= length l --> list_all P (SEG m k l))" by (import rich_list ALL_EL_SEG) -lemma ALL_EL_FIRSTN: "(All::(('a => bool) => bool) => bool) - (%P::'a => bool. - (All::('a list => bool) => bool) - (%l::'a list. +lemma ALL_EL_FIRSTN: "(All::(('a::type => bool) => bool) => bool) + (%P::'a::type => bool. + (All::('a::type list => bool) => bool) + (%l::'a::type list. (op -->::bool => bool => bool) - ((list_all::('a => bool) => 'a list => bool) P l) - ((All::(nat => bool) => bool) - (%m::nat. - (op -->::bool => bool => bool) - ((op <=::nat => nat => bool) m ((size::'a list => nat) l)) - ((list_all::('a => bool) => 'a list => bool) P - ((FIRSTN::nat => 'a list => 'a list) m l))))))" - by (import rich_list ALL_EL_FIRSTN) - -lemma ALL_EL_BUTFIRSTN: "(All::(('a => bool) => bool) => bool) - (%P::'a => bool. - (All::('a list => bool) => bool) - (%l::'a list. - (op -->::bool => bool => bool) - ((list_all::('a => bool) => 'a list => bool) P l) + ((list_all::('a::type => bool) => 'a::type list => bool) P l) ((All::(nat => bool) => bool) (%m::nat. (op -->::bool => bool => bool) - ((op <=::nat => nat => bool) m ((size::'a list => nat) l)) - ((list_all::('a => bool) => 'a list => bool) P - ((BUTFIRSTN::nat => 'a list => 'a list) m l))))))" + ((op <=::nat => nat => bool) m + ((size::'a::type list => nat) l)) + ((list_all::('a::type => bool) => 'a::type list => bool) P + ((FIRSTN::nat => 'a::type list => 'a::type list) m + l))))))" + by (import rich_list ALL_EL_FIRSTN) + +lemma ALL_EL_BUTFIRSTN: "(All::(('a::type => bool) => bool) => bool) + (%P::'a::type => bool. + (All::('a::type list => bool) => bool) + (%l::'a::type list. + (op -->::bool => bool => bool) + ((list_all::('a::type => bool) => 'a::type list => bool) P l) + ((All::(nat => bool) => bool) + (%m::nat. + (op -->::bool => bool => bool) + ((op <=::nat => nat => bool) m + ((size::'a::type list => nat) l)) + ((list_all::('a::type => bool) => 'a::type list => bool) P + ((BUTFIRSTN::nat => 'a::type list => 'a::type list) m + l))))))" by (import rich_list ALL_EL_BUTFIRSTN) -lemma SOME_EL_SEG: "ALL m k l. +lemma SOME_EL_SEG: "ALL (m::nat) (k::nat) l::'a::type list. m + k <= length l --> - (ALL P. list_exists P (SEG m k l) --> list_exists P l)" + (ALL P::'a::type => bool. list_exists P (SEG m k l) --> list_exists P l)" by (import rich_list SOME_EL_SEG) -lemma SOME_EL_FIRSTN: "ALL m l. - m <= length l --> (ALL P. list_exists P (FIRSTN m l) --> list_exists P l)" +lemma SOME_EL_FIRSTN: "ALL (m::nat) l::'a::type list. + m <= length l --> + (ALL P::'a::type => bool. list_exists P (FIRSTN m l) --> list_exists P l)" by (import rich_list SOME_EL_FIRSTN) -lemma SOME_EL_BUTFIRSTN: "ALL m l. +lemma SOME_EL_BUTFIRSTN: "ALL (m::nat) l::'a::type list. m <= length l --> - (ALL P. list_exists P (BUTFIRSTN m l) --> list_exists P l)" + (ALL P::'a::type => bool. + list_exists P (BUTFIRSTN m l) --> list_exists P l)" by (import rich_list SOME_EL_BUTFIRSTN) -lemma SOME_EL_LASTN: "ALL m l. - m <= length l --> (ALL P. list_exists P (LASTN m l) --> list_exists P l)" +lemma SOME_EL_LASTN: "ALL (m::nat) l::'a::type list. + m <= length l --> + (ALL P::'a::type => bool. list_exists P (LASTN m l) --> list_exists P l)" by (import rich_list SOME_EL_LASTN) -lemma SOME_EL_BUTLASTN: "ALL m l. +lemma SOME_EL_BUTLASTN: "ALL (m::nat) l::'a::type list. m <= length l --> - (ALL P. list_exists P (BUTLASTN m l) --> list_exists P l)" + (ALL P::'a::type => bool. + list_exists P (BUTLASTN m l) --> list_exists P l)" by (import rich_list SOME_EL_BUTLASTN) -lemma IS_EL_REVERSE: "ALL x l. x mem rev l = x mem l" +lemma IS_EL_REVERSE: "ALL (x::'a::type) l::'a::type list. x mem rev l = x mem l" by (import rich_list IS_EL_REVERSE) -lemma IS_EL_FILTER: "ALL P x. P x --> (ALL l. x mem filter P l = x mem l)" +lemma IS_EL_FILTER: "ALL (P::'a::type => bool) x::'a::type. + P x --> (ALL l::'a::type list. x mem filter P l = x mem l)" by (import rich_list IS_EL_FILTER) -lemma IS_EL_SEG: "ALL n m l. n + m <= length l --> (ALL x. x mem SEG n m l --> x mem l)" +lemma IS_EL_SEG: "ALL (n::nat) (m::nat) l::'a::type list. + n + m <= length l --> (ALL x::'a::type. x mem SEG n m l --> x mem l)" by (import rich_list IS_EL_SEG) -lemma IS_EL_SOME_EL: "ALL x l. x mem l = list_exists (op = x) l" +lemma IS_EL_SOME_EL: "ALL (x::'a::type) l::'a::type list. x mem l = list_exists (op = x) l" by (import rich_list IS_EL_SOME_EL) -lemma IS_EL_FIRSTN: "ALL x xa. x <= length xa --> (ALL xb. xb mem FIRSTN x xa --> xb mem xa)" +lemma IS_EL_FIRSTN: "ALL (x::nat) xa::'a::type list. + x <= length xa --> (ALL xb::'a::type. xb mem FIRSTN x xa --> xb mem xa)" by (import rich_list IS_EL_FIRSTN) -lemma IS_EL_BUTFIRSTN: "ALL x xa. x <= length xa --> (ALL xb. xb mem BUTFIRSTN x xa --> xb mem xa)" +lemma IS_EL_BUTFIRSTN: "ALL (x::nat) xa::'a::type list. + x <= length xa --> + (ALL xb::'a::type. xb mem BUTFIRSTN x xa --> xb mem xa)" by (import rich_list IS_EL_BUTFIRSTN) -lemma IS_EL_BUTLASTN: "ALL x xa. x <= length xa --> (ALL xb. xb mem BUTLASTN x xa --> xb mem xa)" +lemma IS_EL_BUTLASTN: "ALL (x::nat) xa::'a::type list. + x <= length xa --> (ALL xb::'a::type. xb mem BUTLASTN x xa --> xb mem xa)" by (import rich_list IS_EL_BUTLASTN) -lemma IS_EL_LASTN: "ALL x xa. x <= length xa --> (ALL xb. xb mem LASTN x xa --> xb mem xa)" +lemma IS_EL_LASTN: "ALL (x::nat) xa::'a::type list. + x <= length xa --> (ALL xb::'a::type. xb mem LASTN x xa --> xb mem xa)" by (import rich_list IS_EL_LASTN) -lemma ZIP_SNOC: "ALL l1 l2. +lemma ZIP_SNOC: "ALL (l1::'a::type list) l2::'b::type list. length l1 = length l2 --> - (ALL x1 x2. zip (SNOC x1 l1) (SNOC x2 l2) = SNOC (x1, x2) (zip l1 l2))" + (ALL (x1::'a::type) x2::'b::type. + zip (SNOC x1 l1) (SNOC x2 l2) = SNOC (x1, x2) (zip l1 l2))" by (import rich_list ZIP_SNOC) -lemma UNZIP_SNOC: "ALL x l. +lemma UNZIP_SNOC: "ALL (x::'a::type * 'b::type) l::('a::type * 'b::type) list. unzip (SNOC x l) = (SNOC (fst x) (fst (unzip l)), SNOC (snd x) (snd (unzip l)))" by (import rich_list UNZIP_SNOC) -lemma LENGTH_UNZIP_FST: "ALL x. length (UNZIP_FST x) = length x" +lemma LENGTH_UNZIP_FST: "ALL x::('a::type * 'b::type) list. length (UNZIP_FST x) = length x" by (import rich_list LENGTH_UNZIP_FST) -lemma LENGTH_UNZIP_SND: "ALL x. length (UNZIP_SND x) = length x" +lemma LENGTH_UNZIP_SND: "ALL x::('a::type * 'b::type) list. length (UNZIP_SND x) = length x" by (import rich_list LENGTH_UNZIP_SND) -lemma SUM_APPEND: "ALL l1 l2. sum (l1 @ l2) = sum l1 + sum l2" +lemma SUM_APPEND: "ALL (l1::nat list) l2::nat list. sum (l1 @ l2) = sum l1 + sum l2" by (import rich_list SUM_APPEND) -lemma SUM_REVERSE: "ALL l. sum (rev l) = sum l" +lemma SUM_REVERSE: "ALL l::nat list. sum (rev l) = sum l" by (import rich_list SUM_REVERSE) -lemma SUM_FLAT: "ALL l. sum (concat l) = sum (map sum l)" +lemma SUM_FLAT: "ALL l::nat list list. sum (concat l) = sum (map sum l)" by (import rich_list SUM_FLAT) -lemma EL_APPEND1: "ALL n l1 l2. n < length l1 --> EL n (l1 @ l2) = EL n l1" +lemma EL_APPEND1: "ALL (n::nat) (l1::'a::type list) l2::'a::type list. + n < length l1 --> EL n (l1 @ l2) = EL n l1" by (import rich_list EL_APPEND1) -lemma EL_APPEND2: "ALL l1 n. - length l1 <= n --> (ALL l2. EL n (l1 @ l2) = EL (n - length l1) l2)" +lemma EL_APPEND2: "ALL (l1::'a::type list) n::nat. + length l1 <= n --> + (ALL l2::'a::type list. EL n (l1 @ l2) = EL (n - length l1) l2)" by (import rich_list EL_APPEND2) -lemma EL_MAP: "ALL n l. n < length l --> (ALL f. EL n (map f l) = f (EL n l))" +lemma EL_MAP: "ALL (n::nat) l::'a::type list. + n < length l --> + (ALL f::'a::type => 'b::type. EL n (map f l) = f (EL n l))" by (import rich_list EL_MAP) -lemma EL_CONS: "ALL n>0. ALL x l. EL n (x # l) = EL (PRE n) l" +lemma EL_CONS: "ALL n>0::nat. + ALL (x::'a::type) l::'a::type list. EL n (x # l) = EL (PRE n) l" by (import rich_list EL_CONS) -lemma EL_SEG: "ALL n l. n < length l --> EL n l = hd (SEG 1 n l)" +lemma EL_SEG: "ALL (n::nat) l::'a::type list. + n < length l --> EL n l = hd (SEG (1::nat) n l)" by (import rich_list EL_SEG) -lemma EL_IS_EL: "ALL n l. n < length l --> EL n l mem l" +lemma EL_IS_EL: "ALL (n::nat) l::'a::type list. n < length l --> EL n l mem l" by (import rich_list EL_IS_EL) -lemma TL_SNOC: "ALL x l. tl (SNOC x l) = (if null l then [] else SNOC x (tl l))" +lemma TL_SNOC: "ALL (x::'a::type) l::'a::type list. + tl (SNOC x l) = (if null l then [] else SNOC x (tl l))" by (import rich_list TL_SNOC) -lemma EL_REVERSE: "ALL n l. n < length l --> EL n (rev l) = EL (PRE (length l - n)) l" +lemma EL_REVERSE: "ALL (n::nat) l::'a::type list. + n < length l --> EL n (rev l) = EL (PRE (length l - n)) l" by (import rich_list EL_REVERSE) -lemma EL_REVERSE_ELL: "ALL n l. n < length l --> EL n (rev l) = ELL n l" +lemma EL_REVERSE_ELL: "ALL (n::nat) l::'a::type list. n < length l --> EL n (rev l) = ELL n l" by (import rich_list EL_REVERSE_ELL) -lemma ELL_LENGTH_APPEND: "ALL l1 l2. ~ null l1 --> ELL (length l2) (l1 @ l2) = last l1" +lemma ELL_LENGTH_APPEND: "ALL (l1::'a::type list) l2::'a::type list. + ~ null l1 --> ELL (length l2) (l1 @ l2) = last l1" by (import rich_list ELL_LENGTH_APPEND) -lemma ELL_IS_EL: "ALL n l. n < length l --> ELL n l mem l" +lemma ELL_IS_EL: "ALL (n::nat) l::'a::type list. n < length l --> ELL n l mem l" by (import rich_list ELL_IS_EL) -lemma ELL_REVERSE: "ALL n l. n < length l --> ELL n (rev l) = ELL (PRE (length l - n)) l" +lemma ELL_REVERSE: "ALL (n::nat) l::'a::type list. + n < length l --> ELL n (rev l) = ELL (PRE (length l - n)) l" by (import rich_list ELL_REVERSE) -lemma ELL_REVERSE_EL: "ALL n l. n < length l --> ELL n (rev l) = EL n l" +lemma ELL_REVERSE_EL: "ALL (n::nat) l::'a::type list. n < length l --> ELL n (rev l) = EL n l" by (import rich_list ELL_REVERSE_EL) -lemma FIRSTN_BUTLASTN: "ALL n l. n <= length l --> FIRSTN n l = BUTLASTN (length l - n) l" +lemma FIRSTN_BUTLASTN: "ALL (n::nat) l::'a::type list. + n <= length l --> FIRSTN n l = BUTLASTN (length l - n) l" by (import rich_list FIRSTN_BUTLASTN) -lemma BUTLASTN_FIRSTN: "ALL n l. n <= length l --> BUTLASTN n l = FIRSTN (length l - n) l" +lemma BUTLASTN_FIRSTN: "ALL (n::nat) l::'a::type list. + n <= length l --> BUTLASTN n l = FIRSTN (length l - n) l" by (import rich_list BUTLASTN_FIRSTN) -lemma LASTN_BUTFIRSTN: "ALL n l. n <= length l --> LASTN n l = BUTFIRSTN (length l - n) l" +lemma LASTN_BUTFIRSTN: "ALL (n::nat) l::'a::type list. + n <= length l --> LASTN n l = BUTFIRSTN (length l - n) l" by (import rich_list LASTN_BUTFIRSTN) -lemma BUTFIRSTN_LASTN: "ALL n l. n <= length l --> BUTFIRSTN n l = LASTN (length l - n) l" +lemma BUTFIRSTN_LASTN: "ALL (n::nat) l::'a::type list. + n <= length l --> BUTFIRSTN n l = LASTN (length l - n) l" by (import rich_list BUTFIRSTN_LASTN) -lemma SEG_LASTN_BUTLASTN: "ALL n m l. +lemma SEG_LASTN_BUTLASTN: "ALL (n::nat) (m::nat) l::'a::type list. n + m <= length l --> SEG n m l = LASTN n (BUTLASTN (length l - (n + m)) l)" by (import rich_list SEG_LASTN_BUTLASTN) -lemma BUTFIRSTN_REVERSE: "ALL n l. n <= length l --> BUTFIRSTN n (rev l) = rev (BUTLASTN n l)" +lemma BUTFIRSTN_REVERSE: "ALL (n::nat) l::'a::type list. + n <= length l --> BUTFIRSTN n (rev l) = rev (BUTLASTN n l)" by (import rich_list BUTFIRSTN_REVERSE) -lemma BUTLASTN_REVERSE: "ALL n l. n <= length l --> BUTLASTN n (rev l) = rev (BUTFIRSTN n l)" +lemma BUTLASTN_REVERSE: "ALL (n::nat) l::'a::type list. + n <= length l --> BUTLASTN n (rev l) = rev (BUTFIRSTN n l)" by (import rich_list BUTLASTN_REVERSE) -lemma LASTN_REVERSE: "ALL n l. n <= length l --> LASTN n (rev l) = rev (FIRSTN n l)" +lemma LASTN_REVERSE: "ALL (n::nat) l::'a::type list. + n <= length l --> LASTN n (rev l) = rev (FIRSTN n l)" by (import rich_list LASTN_REVERSE) -lemma FIRSTN_REVERSE: "ALL n l. n <= length l --> FIRSTN n (rev l) = rev (LASTN n l)" +lemma FIRSTN_REVERSE: "ALL (n::nat) l::'a::type list. + n <= length l --> FIRSTN n (rev l) = rev (LASTN n l)" by (import rich_list FIRSTN_REVERSE) -lemma SEG_REVERSE: "ALL n m l. +lemma SEG_REVERSE: "ALL (n::nat) (m::nat) l::'a::type list. n + m <= length l --> SEG n m (rev l) = rev (SEG n (length l - (n + m)) l)" by (import rich_list SEG_REVERSE) -lemma LENGTH_GENLIST: "ALL f n. length (GENLIST f n) = n" +lemma LENGTH_GENLIST: "ALL (f::nat => 'a::type) n::nat. length (GENLIST f n) = n" by (import rich_list LENGTH_GENLIST) -lemma LENGTH_REPLICATE: "ALL n x. length (REPLICATE n x) = n" +lemma LENGTH_REPLICATE: "ALL (n::nat) x::'a::type. length (REPLICATE n x) = n" by (import rich_list LENGTH_REPLICATE) -lemma IS_EL_REPLICATE: "ALL n>0. ALL x. x mem REPLICATE n x" +lemma IS_EL_REPLICATE: "ALL n>0::nat. ALL x::'a::type. x mem REPLICATE n x" by (import rich_list IS_EL_REPLICATE) -lemma ALL_EL_REPLICATE: "ALL x n. list_all (op = x) (REPLICATE n x)" +lemma ALL_EL_REPLICATE: "ALL (x::'a::type) n::nat. list_all (op = x) (REPLICATE n x)" by (import rich_list ALL_EL_REPLICATE) -lemma AND_EL_FOLDL: "ALL l. AND_EL l = foldl op & True l" +lemma AND_EL_FOLDL: "ALL l::bool list. AND_EL l = foldl op & True l" by (import rich_list AND_EL_FOLDL) -lemma AND_EL_FOLDR: "ALL l. AND_EL l = foldr op & l True" +lemma AND_EL_FOLDR: "ALL l::bool list. AND_EL l = foldr op & l True" by (import rich_list AND_EL_FOLDR) -lemma OR_EL_FOLDL: "ALL l. OR_EL l = foldl op | False l" +lemma OR_EL_FOLDL: "ALL l::bool list. OR_EL l = foldl op | False l" by (import rich_list OR_EL_FOLDL) -lemma OR_EL_FOLDR: "ALL l. OR_EL l = foldr op | l False" +lemma OR_EL_FOLDR: "ALL l::bool list. OR_EL l = foldr op | l False" by (import rich_list OR_EL_FOLDR) ;end_setup @@ -4861,71 +5749,106 @@ ;setup_theory state_transformer constdefs - UNIT :: "'b => 'a => 'b * 'a" - "(op ==::('b => 'a => 'b * 'a) => ('b => 'a => 'b * 'a) => prop) - (UNIT::'b => 'a => 'b * 'a) (Pair::'b => 'a => 'b * 'a)" - -lemma UNIT_DEF: "ALL x::'b. UNIT x = Pair x" + UNIT :: "'b::type => 'a::type => 'b::type * 'a::type" + "(op ==::('b::type => 'a::type => 'b::type * 'a::type) + => ('b::type => 'a::type => 'b::type * 'a::type) => prop) + (UNIT::'b::type => 'a::type => 'b::type * 'a::type) + (Pair::'b::type => 'a::type => 'b::type * 'a::type)" + +lemma UNIT_DEF: "ALL x::'b::type. UNIT x = Pair x" by (import state_transformer UNIT_DEF) constdefs - BIND :: "('a => 'b * 'a) => ('b => 'a => 'c * 'a) => 'a => 'c * 'a" - "(op ==::(('a => 'b * 'a) => ('b => 'a => 'c * 'a) => 'a => 'c * 'a) - => (('a => 'b * 'a) => ('b => 'a => 'c * 'a) => 'a => 'c * 'a) + BIND :: "('a::type => 'b::type * 'a::type) +=> ('b::type => 'a::type => 'c::type * 'a::type) + => 'a::type => 'c::type * 'a::type" + "(op ==::(('a::type => 'b::type * 'a::type) + => ('b::type => 'a::type => 'c::type * 'a::type) + => 'a::type => 'c::type * 'a::type) + => (('a::type => 'b::type * 'a::type) + => ('b::type => 'a::type => 'c::type * 'a::type) + => 'a::type => 'c::type * 'a::type) => prop) - (BIND::('a => 'b * 'a) => ('b => 'a => 'c * 'a) => 'a => 'c * 'a) - (%(g::'a => 'b * 'a) f::'b => 'a => 'c * 'a. - (op o::('b * 'a => 'c * 'a) => ('a => 'b * 'a) => 'a => 'c * 'a) - ((split::('b => 'a => 'c * 'a) => 'b * 'a => 'c * 'a) f) g)" - -lemma BIND_DEF: "(All::(('a => 'b * 'a) => bool) => bool) - (%g::'a => 'b * 'a. - (All::(('b => 'a => 'c * 'a) => bool) => bool) - (%f::'b => 'a => 'c * 'a. - (op =::('a => 'c * 'a) => ('a => 'c * 'a) => bool) - ((BIND::('a => 'b * 'a) - => ('b => 'a => 'c * 'a) => 'a => 'c * 'a) + (BIND::('a::type => 'b::type * 'a::type) + => ('b::type => 'a::type => 'c::type * 'a::type) + => 'a::type => 'c::type * 'a::type) + (%(g::'a::type => 'b::type * 'a::type) + f::'b::type => 'a::type => 'c::type * 'a::type. + (op o::('b::type * 'a::type => 'c::type * 'a::type) + => ('a::type => 'b::type * 'a::type) + => 'a::type => 'c::type * 'a::type) + ((split::('b::type => 'a::type => 'c::type * 'a::type) + => 'b::type * 'a::type => 'c::type * 'a::type) + f) + g)" + +lemma BIND_DEF: "(All::(('a::type => 'b::type * 'a::type) => bool) => bool) + (%g::'a::type => 'b::type * 'a::type. + (All::(('b::type => 'a::type => 'c::type * 'a::type) => bool) => bool) + (%f::'b::type => 'a::type => 'c::type * 'a::type. + (op =::('a::type => 'c::type * 'a::type) + => ('a::type => 'c::type * 'a::type) => bool) + ((BIND::('a::type => 'b::type * 'a::type) + => ('b::type => 'a::type => 'c::type * 'a::type) + => 'a::type => 'c::type * 'a::type) g f) - ((op o::('b * 'a => 'c * 'a) => ('a => 'b * 'a) => 'a => 'c * 'a) - ((split::('b => 'a => 'c * 'a) => 'b * 'a => 'c * 'a) f) g)))" + ((op o::('b::type * 'a::type => 'c::type * 'a::type) + => ('a::type => 'b::type * 'a::type) + => 'a::type => 'c::type * 'a::type) + ((split::('b::type => 'a::type => 'c::type * 'a::type) + => 'b::type * 'a::type => 'c::type * 'a::type) + f) + g)))" by (import state_transformer BIND_DEF) constdefs - MMAP :: "('c => 'b) => ('a => 'c * 'a) => 'a => 'b * 'a" - "MMAP == %(f::'c => 'b) m::'a => 'c * 'a. BIND m (UNIT o f)" - -lemma MMAP_DEF: "ALL (f::'c => 'b) m::'a => 'c * 'a. MMAP f m = BIND m (UNIT o f)" + MMAP :: "('c::type => 'b::type) +=> ('a::type => 'c::type * 'a::type) => 'a::type => 'b::type * 'a::type" + "MMAP == +%(f::'c::type => 'b::type) m::'a::type => 'c::type * 'a::type. + BIND m (UNIT o f)" + +lemma MMAP_DEF: "ALL (f::'c::type => 'b::type) m::'a::type => 'c::type * 'a::type. + MMAP f m = BIND m (UNIT o f)" by (import state_transformer MMAP_DEF) constdefs - JOIN :: "('a => ('a => 'b * 'a) * 'a) => 'a => 'b * 'a" - "JOIN == %z. BIND z I" - -lemma JOIN_DEF: "ALL z. JOIN z = BIND z I" + JOIN :: "('a::type => ('a::type => 'b::type * 'a::type) * 'a::type) +=> 'a::type => 'b::type * 'a::type" + "JOIN == +%z::'a::type => ('a::type => 'b::type * 'a::type) * 'a::type. BIND z I" + +lemma JOIN_DEF: "ALL z::'a::type => ('a::type => 'b::type * 'a::type) * 'a::type. + JOIN z = BIND z I" by (import state_transformer JOIN_DEF) -lemma BIND_LEFT_UNIT: "ALL k x. BIND (UNIT x) k = k x" +lemma BIND_LEFT_UNIT: "ALL (k::'a::type => 'b::type => 'c::type * 'b::type) x::'a::type. + BIND (UNIT x) k = k x" by (import state_transformer BIND_LEFT_UNIT) -lemma UNIT_UNCURRY: "ALL x. split UNIT x = x" +lemma UNIT_UNCURRY: "ALL x::'a::type * 'b::type. split UNIT x = x" by (import state_transformer UNIT_UNCURRY) -lemma BIND_RIGHT_UNIT: "ALL k. BIND k UNIT = k" +lemma BIND_RIGHT_UNIT: "ALL k::'a::type => 'b::type * 'a::type. BIND k UNIT = k" by (import state_transformer BIND_RIGHT_UNIT) -lemma BIND_ASSOC: "ALL x xa xb. BIND x (%a. BIND (xa a) xb) = BIND (BIND x xa) xb" +lemma BIND_ASSOC: "ALL (x::'a::type => 'b::type * 'a::type) + (xa::'b::type => 'a::type => 'c::type * 'a::type) + xb::'c::type => 'a::type => 'd::type * 'a::type. + BIND x (%a::'b::type. BIND (xa a) xb) = BIND (BIND x xa) xb" by (import state_transformer BIND_ASSOC) lemma MMAP_ID: "MMAP I = I" by (import state_transformer MMAP_ID) -lemma MMAP_COMP: "ALL (f::'c => 'd) g::'b => 'c. MMAP (f o g) = MMAP f o MMAP g" +lemma MMAP_COMP: "ALL (f::'c::type => 'd::type) g::'b::type => 'c::type. + MMAP (f o g) = MMAP f o MMAP g" by (import state_transformer MMAP_COMP) -lemma MMAP_UNIT: "ALL f::'b => 'c. MMAP f o UNIT = UNIT o f" +lemma MMAP_UNIT: "ALL f::'b::type => 'c::type. MMAP f o UNIT = UNIT o f" by (import state_transformer MMAP_UNIT) -lemma MMAP_JOIN: "ALL f::'b => 'c. MMAP f o JOIN = JOIN o MMAP (MMAP f)" +lemma MMAP_JOIN: "ALL f::'b::type => 'c::type. MMAP f o JOIN = JOIN o MMAP (MMAP f)" by (import state_transformer MMAP_JOIN) lemma JOIN_UNIT: "JOIN o UNIT = I" @@ -4937,16 +5860,19 @@ lemma JOIN_MAP_JOIN: "JOIN o MMAP JOIN = JOIN o JOIN" by (import state_transformer JOIN_MAP_JOIN) -lemma JOIN_MAP: "ALL x xa. BIND x xa = JOIN (MMAP xa x)" +lemma JOIN_MAP: "ALL (x::'a::type => 'b::type * 'a::type) + xa::'b::type => 'a::type => 'c::type * 'a::type. + BIND x xa = JOIN (MMAP xa x)" by (import state_transformer JOIN_MAP) -lemma FST_o_UNIT: "ALL x. fst o UNIT x = K x" +lemma FST_o_UNIT: "ALL x::'a::type. fst o UNIT x = K x" by (import state_transformer FST_o_UNIT) -lemma SND_o_UNIT: "ALL x. snd o UNIT x = I" +lemma SND_o_UNIT: "ALL x::'a::type. snd o UNIT x = I" by (import state_transformer SND_o_UNIT) -lemma FST_o_MMAP: "ALL x xa. fst o MMAP x xa = x o (fst o xa)" +lemma FST_o_MMAP: "ALL (x::'a::type => 'b::type) xa::'c::type => 'a::type * 'c::type. + fst o MMAP x xa = x o (fst o xa)" by (import state_transformer FST_o_MMAP) ;end_setup diff -r 7e417a7cbf9f -r bd59bfd4bf37 src/HOL/Import/HOL/HOL4Prob.thy --- a/src/HOL/Import/HOL/HOL4Prob.thy Mon Sep 26 02:06:44 2005 +0200 +++ b/src/HOL/Import/HOL/HOL4Prob.thy Mon Sep 26 02:27:14 2005 +0200 @@ -4,16 +4,22 @@ ;setup_theory prob_extra -lemma BOOL_BOOL_CASES_THM: "ALL f. f = (%b. False) | f = (%b. True) | f = (%b. b) | f = Not" +lemma BOOL_BOOL_CASES_THM: "ALL f::bool => bool. + f = (%b::bool. False) | + f = (%b::bool. True) | f = (%b::bool. b) | f = Not" by (import prob_extra BOOL_BOOL_CASES_THM) -lemma EVEN_ODD_BASIC: "EVEN 0 & ~ EVEN 1 & EVEN 2 & ~ ODD 0 & ODD 1 & ~ ODD 2" +lemma EVEN_ODD_BASIC: "EVEN (0::nat) & +~ EVEN (1::nat) & +EVEN (2::nat) & ~ ODD (0::nat) & ODD (1::nat) & ~ ODD (2::nat)" by (import prob_extra EVEN_ODD_BASIC) -lemma EVEN_ODD_EXISTS_EQ: "ALL n. EVEN n = (EX m. n = 2 * m) & ODD n = (EX m. n = Suc (2 * m))" +lemma EVEN_ODD_EXISTS_EQ: "ALL n::nat. + EVEN n = (EX m::nat. n = (2::nat) * m) & + ODD n = (EX m::nat. n = Suc ((2::nat) * m))" by (import prob_extra EVEN_ODD_EXISTS_EQ) -lemma DIV_THEN_MULT: "ALL p q. Suc q * (p div Suc q) <= p" +lemma DIV_THEN_MULT: "ALL (p::nat) q::nat. Suc q * (p div Suc q) <= p" by (import prob_extra DIV_THEN_MULT) lemma DIV_TWO_UNIQUE: "(All::(nat => bool) => bool) @@ -62,7 +68,8 @@ lemma DIV_TWO: "ALL n::nat. n = (2::nat) * (n div (2::nat)) + n mod (2::nat)" by (import prob_extra DIV_TWO) -lemma MOD_TWO: "ALL n. n mod 2 = (if EVEN n then 0 else 1)" +lemma MOD_TWO: "(ALL (n::nat). + ((n mod (2::nat)) = (if (EVEN n) then (0::nat) else (1::nat))))" by (import prob_extra MOD_TWO) lemma DIV_TWO_BASIC: "(0::nat) div (2::nat) = (0::nat) & @@ -112,13 +119,14 @@ ((op <::nat => nat => bool) m n))))" by (import prob_extra DIV_TWO_MONO_EVEN) -lemma DIV_TWO_CANCEL: "ALL n. 2 * n div 2 = n & Suc (2 * n) div 2 = n" +lemma DIV_TWO_CANCEL: "ALL n::nat. + (2::nat) * n div (2::nat) = n & Suc ((2::nat) * n) div (2::nat) = n" by (import prob_extra DIV_TWO_CANCEL) lemma EXP_DIV_TWO: "ALL n::nat. (2::nat) ^ Suc n div (2::nat) = (2::nat) ^ n" by (import prob_extra EXP_DIV_TWO) -lemma EVEN_EXP_TWO: "ALL n. EVEN (2 ^ n) = (n ~= 0)" +lemma EVEN_EXP_TWO: "ALL n::nat. EVEN ((2::nat) ^ n) = (n ~= (0::nat))" by (import prob_extra EVEN_EXP_TWO) lemma DIV_TWO_EXP: "ALL (n::nat) k::nat. @@ -129,12 +137,12 @@ inf :: "(real => bool) => real" defs - inf_primdef: "inf == %P. - sup (IMAGE uminus P)" - -lemma inf_def: "ALL P. inf P = - sup (IMAGE uminus P)" + inf_primdef: "inf == %P::real => bool. - sup (IMAGE uminus P)" + +lemma inf_def: "ALL P::real => bool. inf P = - sup (IMAGE uminus P)" by (import prob_extra inf_def) -lemma INF_DEF_ALT: "ALL P. inf P = - sup (%r. P (- r))" +lemma INF_DEF_ALT: "ALL P::real => bool. inf P = - sup (%r::real. P (- r))" by (import prob_extra INF_DEF_ALT) lemma REAL_SUP_EXISTS_UNIQUE: "(All::((real => bool) => bool) => bool) @@ -275,13 +283,15 @@ lemma POW_HALF_EXP: "ALL n::nat. ((1::real) / (2::real)) ^ n = inverse (real ((2::nat) ^ n))" by (import prob_extra POW_HALF_EXP) -lemma INV_SUC_POS: "ALL n. 0 < 1 / real (Suc n)" +lemma INV_SUC_POS: "ALL n::nat. (0::real) < (1::real) / real (Suc n)" by (import prob_extra INV_SUC_POS) -lemma INV_SUC_MAX: "ALL x. 1 / real (Suc x) <= 1" +lemma INV_SUC_MAX: "ALL x::nat. (1::real) / real (Suc x) <= (1::real)" by (import prob_extra INV_SUC_MAX) -lemma INV_SUC: "ALL n. 0 < 1 / real (Suc n) & 1 / real (Suc n) <= 1" +lemma INV_SUC: "ALL n::nat. + (0::real) < (1::real) / real (Suc n) & + (1::real) / real (Suc n) <= (1::real)" by (import prob_extra INV_SUC) lemma ABS_UNIT_INTERVAL: "(All::(real => bool) => bool) @@ -301,181 +311,200 @@ (1::real))))" by (import prob_extra ABS_UNIT_INTERVAL) -lemma MEM_NIL: "ALL l. (ALL x. ~ x mem l) = (l = [])" +lemma MEM_NIL: "ALL l::'a::type list. (ALL x::'a::type. ~ x mem l) = (l = [])" by (import prob_extra MEM_NIL) -lemma MAP_MEM: "ALL f l x. x mem map f l = (EX y. y mem l & x = f y)" +lemma MAP_MEM: "ALL (f::'a::type => 'b::type) (l::'a::type list) x::'b::type. + x mem map f l = (EX y::'a::type. y mem l & x = f y)" by (import prob_extra MAP_MEM) -lemma MEM_NIL_MAP_CONS: "ALL x l. ~ [] mem map (op # x) l" +lemma MEM_NIL_MAP_CONS: "ALL (x::'a::type) l::'a::type list list. ~ [] mem map (op # x) l" by (import prob_extra MEM_NIL_MAP_CONS) -lemma FILTER_TRUE: "ALL l. [x:l. True] = l" +lemma FILTER_TRUE: "ALL l::'a::type list. [x::'a::type:l. True] = l" by (import prob_extra FILTER_TRUE) -lemma FILTER_FALSE: "ALL l. [x:l. False] = []" +lemma FILTER_FALSE: "ALL l::'a::type list. [x::'a::type:l. False] = []" by (import prob_extra FILTER_FALSE) -lemma FILTER_MEM: "(All::(('a => bool) => bool) => bool) - (%P::'a => bool. - (All::('a => bool) => bool) - (%x::'a. - (All::('a list => bool) => bool) - (%l::'a list. +lemma FILTER_MEM: "(All::(('a::type => bool) => bool) => bool) + (%P::'a::type => bool. + (All::('a::type => bool) => bool) + (%x::'a::type. + (All::('a::type list => bool) => bool) + (%l::'a::type list. (op -->::bool => bool => bool) - ((op mem::'a => 'a list => bool) x - ((filter::('a => bool) => 'a list => 'a list) P l)) + ((op mem::'a::type => 'a::type list => bool) x + ((filter::('a::type => bool) + => 'a::type list => 'a::type list) + P l)) (P x))))" by (import prob_extra FILTER_MEM) -lemma MEM_FILTER: "(All::(('a => bool) => bool) => bool) - (%P::'a => bool. - (All::('a list => bool) => bool) - (%l::'a list. - (All::('a => bool) => bool) - (%x::'a. +lemma MEM_FILTER: "(All::(('a::type => bool) => bool) => bool) + (%P::'a::type => bool. + (All::('a::type list => bool) => bool) + (%l::'a::type list. + (All::('a::type => bool) => bool) + (%x::'a::type. (op -->::bool => bool => bool) - ((op mem::'a => 'a list => bool) x - ((filter::('a => bool) => 'a list => 'a list) P l)) - ((op mem::'a => 'a list => bool) x l))))" + ((op mem::'a::type => 'a::type list => bool) x + ((filter::('a::type => bool) + => 'a::type list => 'a::type list) + P l)) + ((op mem::'a::type => 'a::type list => bool) x l))))" by (import prob_extra MEM_FILTER) -lemma FILTER_OUT_ELT: "ALL x l. x mem l | [y:l. y ~= x] = l" +lemma FILTER_OUT_ELT: "ALL (x::'a::type) l::'a::type list. x mem l | [y::'a::type:l. y ~= x] = l" by (import prob_extra FILTER_OUT_ELT) -lemma IS_PREFIX_NIL: "ALL x. IS_PREFIX x [] & IS_PREFIX [] x = (x = [])" +lemma IS_PREFIX_NIL: "ALL x::'a::type list. IS_PREFIX x [] & IS_PREFIX [] x = (x = [])" by (import prob_extra IS_PREFIX_NIL) -lemma IS_PREFIX_REFL: "ALL x. IS_PREFIX x x" +lemma IS_PREFIX_REFL: "ALL x::'a::type list. IS_PREFIX x x" by (import prob_extra IS_PREFIX_REFL) -lemma IS_PREFIX_ANTISYM: "(All::('a list => bool) => bool) - (%x::'a list. - (All::('a list => bool) => bool) - (%y::'a list. +lemma IS_PREFIX_ANTISYM: "(All::('a::type list => bool) => bool) + (%x::'a::type list. + (All::('a::type list => bool) => bool) + (%y::'a::type list. (op -->::bool => bool => bool) ((op &::bool => bool => bool) - ((IS_PREFIX::'a list => 'a list => bool) y x) - ((IS_PREFIX::'a list => 'a list => bool) x y)) - ((op =::'a list => 'a list => bool) x y)))" + ((IS_PREFIX::'a::type list => 'a::type list => bool) y x) + ((IS_PREFIX::'a::type list => 'a::type list => bool) x y)) + ((op =::'a::type list => 'a::type list => bool) x y)))" by (import prob_extra IS_PREFIX_ANTISYM) -lemma IS_PREFIX_TRANS: "(All::('a list => bool) => bool) - (%x::'a list. - (All::('a list => bool) => bool) - (%y::'a list. - (All::('a list => bool) => bool) - (%z::'a list. +lemma IS_PREFIX_TRANS: "(All::('a::type list => bool) => bool) + (%x::'a::type list. + (All::('a::type list => bool) => bool) + (%y::'a::type list. + (All::('a::type list => bool) => bool) + (%z::'a::type list. (op -->::bool => bool => bool) ((op &::bool => bool => bool) - ((IS_PREFIX::'a list => 'a list => bool) x y) - ((IS_PREFIX::'a list => 'a list => bool) y z)) - ((IS_PREFIX::'a list => 'a list => bool) x z))))" + ((IS_PREFIX::'a::type list => 'a::type list => bool) x y) + ((IS_PREFIX::'a::type list => 'a::type list => bool) y z)) + ((IS_PREFIX::'a::type list => 'a::type list => bool) x z))))" by (import prob_extra IS_PREFIX_TRANS) -lemma IS_PREFIX_BUTLAST: "ALL x y. IS_PREFIX (x # y) (butlast (x # y))" +lemma IS_PREFIX_BUTLAST: "ALL (x::'a::type) y::'a::type list. IS_PREFIX (x # y) (butlast (x # y))" by (import prob_extra IS_PREFIX_BUTLAST) -lemma IS_PREFIX_LENGTH: "(All::('a list => bool) => bool) - (%x::'a list. - (All::('a list => bool) => bool) - (%y::'a list. +lemma IS_PREFIX_LENGTH: "(All::('a::type list => bool) => bool) + (%x::'a::type list. + (All::('a::type list => bool) => bool) + (%y::'a::type list. (op -->::bool => bool => bool) - ((IS_PREFIX::'a list => 'a list => bool) y x) - ((op <=::nat => nat => bool) ((size::'a list => nat) x) - ((size::'a list => nat) y))))" + ((IS_PREFIX::'a::type list => 'a::type list => bool) y x) + ((op <=::nat => nat => bool) ((size::'a::type list => nat) x) + ((size::'a::type list => nat) y))))" by (import prob_extra IS_PREFIX_LENGTH) -lemma IS_PREFIX_LENGTH_ANTI: "(All::('a list => bool) => bool) - (%x::'a list. - (All::('a list => bool) => bool) - (%y::'a list. +lemma IS_PREFIX_LENGTH_ANTI: "(All::('a::type list => bool) => bool) + (%x::'a::type list. + (All::('a::type list => bool) => bool) + (%y::'a::type list. (op -->::bool => bool => bool) ((op &::bool => bool => bool) - ((IS_PREFIX::'a list => 'a list => bool) y x) - ((op =::nat => nat => bool) ((size::'a list => nat) x) - ((size::'a list => nat) y))) - ((op =::'a list => 'a list => bool) x y)))" + ((IS_PREFIX::'a::type list => 'a::type list => bool) y x) + ((op =::nat => nat => bool) ((size::'a::type list => nat) x) + ((size::'a::type list => nat) y))) + ((op =::'a::type list => 'a::type list => bool) x y)))" by (import prob_extra IS_PREFIX_LENGTH_ANTI) -lemma IS_PREFIX_SNOC: "ALL x y z. IS_PREFIX (SNOC x y) z = (IS_PREFIX y z | z = SNOC x y)" +lemma IS_PREFIX_SNOC: "ALL (x::'a::type) (y::'a::type list) z::'a::type list. + IS_PREFIX (SNOC x y) z = (IS_PREFIX y z | z = SNOC x y)" by (import prob_extra IS_PREFIX_SNOC) -lemma FOLDR_MAP: "ALL (f::'b => 'c => 'c) (e::'c) (g::'a => 'b) l::'a list. - foldr f (map g l) e = foldr (%x::'a. f (g x)) l e" +lemma FOLDR_MAP: "ALL (f::'b::type => 'c::type => 'c::type) (e::'c::type) + (g::'a::type => 'b::type) l::'a::type list. + foldr f (map g l) e = foldr (%x::'a::type. f (g x)) l e" by (import prob_extra FOLDR_MAP) -lemma LAST_MEM: "ALL h t. last (h # t) mem h # t" +lemma LAST_MEM: "ALL (h::'a::type) t::'a::type list. last (h # t) mem h # t" by (import prob_extra LAST_MEM) lemma LAST_MAP_CONS: "ALL (b::bool) (h::bool list) t::bool list list. EX x::bool list. last (map (op # b) (h # t)) = b # x" by (import prob_extra LAST_MAP_CONS) -lemma EXISTS_LONGEST: "(All::('a list => bool) => bool) - (%x::'a list. - (All::('a list list => bool) => bool) - (%y::'a list list. - (Ex::('a list => bool) => bool) - (%z::'a list. +lemma EXISTS_LONGEST: "(All::('a::type list => bool) => bool) + (%x::'a::type list. + (All::('a::type list list => bool) => bool) + (%y::'a::type list list. + (Ex::('a::type list => bool) => bool) + (%z::'a::type list. (op &::bool => bool => bool) - ((op mem::'a list => 'a list list => bool) z - ((op #::'a list => 'a list list => 'a list list) x y)) - ((All::('a list => bool) => bool) - (%w::'a list. + ((op mem::'a::type list => 'a::type list list => bool) z + ((op #::'a::type list + => 'a::type list list => 'a::type list list) + x y)) + ((All::('a::type list => bool) => bool) + (%w::'a::type list. (op -->::bool => bool => bool) - ((op mem::'a list => 'a list list => bool) w - ((op #::'a list => 'a list list => 'a list list) x - y)) + ((op mem::'a::type list + => 'a::type list list => bool) + w ((op #::'a::type list + => 'a::type list list +=> 'a::type list list) + x y)) ((op <=::nat => nat => bool) - ((size::'a list => nat) w) - ((size::'a list => nat) z)))))))" + ((size::'a::type list => nat) w) + ((size::'a::type list => nat) z)))))))" by (import prob_extra EXISTS_LONGEST) -lemma UNION_DEF_ALT: "ALL s t. pred_set.UNION s t = (%x. s x | t x)" +lemma UNION_DEF_ALT: "ALL (s::'a::type => bool) t::'a::type => bool. + pred_set.UNION s t = (%x::'a::type. s x | t x)" by (import prob_extra UNION_DEF_ALT) -lemma INTER_UNION_RDISTRIB: "ALL p q r. +lemma INTER_UNION_RDISTRIB: "ALL (p::'a::type => bool) (q::'a::type => bool) r::'a::type => bool. pred_set.INTER (pred_set.UNION p q) r = pred_set.UNION (pred_set.INTER p r) (pred_set.INTER q r)" by (import prob_extra INTER_UNION_RDISTRIB) -lemma SUBSET_EQ: "ALL x xa. (x = xa) = (SUBSET x xa & SUBSET xa x)" +lemma SUBSET_EQ: "ALL (x::'a::type => bool) xa::'a::type => bool. + (x = xa) = (SUBSET x xa & SUBSET xa x)" by (import prob_extra SUBSET_EQ) -lemma INTER_IS_EMPTY: "ALL s t. (pred_set.INTER s t = EMPTY) = (ALL x. ~ s x | ~ t x)" +lemma INTER_IS_EMPTY: "ALL (s::'a::type => bool) t::'a::type => bool. + (pred_set.INTER s t = EMPTY) = (ALL x::'a::type. ~ s x | ~ t x)" by (import prob_extra INTER_IS_EMPTY) -lemma UNION_DISJOINT_SPLIT: "(All::(('a => bool) => bool) => bool) - (%s::'a => bool. - (All::(('a => bool) => bool) => bool) - (%t::'a => bool. - (All::(('a => bool) => bool) => bool) - (%u::'a => bool. +lemma UNION_DISJOINT_SPLIT: "(All::(('a::type => bool) => bool) => bool) + (%s::'a::type => bool. + (All::(('a::type => bool) => bool) => bool) + (%t::'a::type => bool. + (All::(('a::type => bool) => bool) => bool) + (%u::'a::type => bool. (op -->::bool => bool => bool) ((op &::bool => bool => bool) - ((op =::('a => bool) => ('a => bool) => bool) - ((pred_set.UNION::('a => bool) -=> ('a => bool) => 'a => bool) + ((op =::('a::type => bool) => ('a::type => bool) => bool) + ((pred_set.UNION::('a::type => bool) +=> ('a::type => bool) => 'a::type => bool) s t) - ((pred_set.UNION::('a => bool) -=> ('a => bool) => 'a => bool) + ((pred_set.UNION::('a::type => bool) +=> ('a::type => bool) => 'a::type => bool) s u)) ((op &::bool => bool => bool) - ((op =::('a => bool) => ('a => bool) => bool) - ((pred_set.INTER::('a => bool) - => ('a => bool) => 'a => bool) + ((op =::('a::type => bool) + => ('a::type => bool) => bool) + ((pred_set.INTER::('a::type => bool) + => ('a::type => bool) => 'a::type => bool) s t) - (EMPTY::'a => bool)) - ((op =::('a => bool) => ('a => bool) => bool) - ((pred_set.INTER::('a => bool) - => ('a => bool) => 'a => bool) + (EMPTY::'a::type => bool)) + ((op =::('a::type => bool) + => ('a::type => bool) => bool) + ((pred_set.INTER::('a::type => bool) + => ('a::type => bool) => 'a::type => bool) s u) - (EMPTY::'a => bool)))) - ((op =::('a => bool) => ('a => bool) => bool) t u))))" + (EMPTY::'a::type => bool)))) + ((op =::('a::type => bool) => ('a::type => bool) => bool) t + u))))" by (import prob_extra UNION_DISJOINT_SPLIT) -lemma GSPEC_DEF_ALT: "ALL f. GSPEC f = (%v. EX x. (v, True) = f x)" +lemma GSPEC_DEF_ALT: "ALL f::'a::type => 'b::type * bool. + GSPEC f = (%v::'b::type. EX x::'a::type. (v, True) = f x)" by (import prob_extra GSPEC_DEF_ALT) ;end_setup @@ -486,9 +515,12 @@ alg_twin :: "bool list => bool list => bool" defs - alg_twin_primdef: "alg_twin == %x y. EX l. x = SNOC True l & y = SNOC False l" - -lemma alg_twin_def: "ALL x y. alg_twin x y = (EX l. x = SNOC True l & y = SNOC False l)" + alg_twin_primdef: "alg_twin == +%(x::bool list) y::bool list. + EX l::bool list. x = SNOC True l & y = SNOC False l" + +lemma alg_twin_def: "ALL (x::bool list) y::bool list. + alg_twin x y = (EX l::bool list. x = SNOC True l & y = SNOC False l)" by (import prob_canon alg_twin_def) constdefs @@ -618,9 +650,9 @@ alg_order :: "bool list => bool list => bool" defs - alg_order_primdef: "alg_order == %x x1. alg_order_tupled (x, x1)" - -lemma alg_order_curried_def: "ALL x x1. alg_order x x1 = alg_order_tupled (x, x1)" + alg_order_primdef: "alg_order == %(x::bool list) x1::bool list. alg_order_tupled (x, x1)" + +lemma alg_order_curried_def: "ALL (x::bool list) x1::bool list. alg_order x x1 = alg_order_tupled (x, x1)" by (import prob_canon alg_order_curried_def) lemma alg_order_ind: "(All::((bool list => bool list => bool) => bool) => bool) @@ -658,10 +690,10 @@ (%x::bool list. (All::(bool list => bool) => bool) (P x))))" by (import prob_canon alg_order_ind) -lemma alg_order_def: "alg_order [] (v6 # v7) = True & +lemma alg_order_def: "alg_order [] ((v6::bool) # (v7::bool list)) = True & alg_order [] [] = True & -alg_order (v2 # v3) [] = False & -alg_order (h # t) (h' # t') = +alg_order ((v2::bool) # (v3::bool list)) [] = False & +alg_order ((h::bool) # (t::bool list)) ((h'::bool) # (t'::bool list)) = (h = True & h' = False | h = h' & alg_order t t')" by (import prob_canon alg_order_def) @@ -670,18 +702,30 @@ defs alg_sorted_primdef: "alg_sorted == -WFREC (SOME R. WF R & (ALL x z y. R (y # z) (x # y # z))) - (%alg_sorted. +WFREC + (SOME R::bool list list => bool list list => bool. + WF R & + (ALL (x::bool list) (z::bool list list) y::bool list. + R (y # z) (x # y # z))) + (%alg_sorted::bool list list => bool. list_case True - (%v2. list_case True - (%v6 v7. alg_order v2 v6 & alg_sorted (v6 # v7))))" + (%v2::bool list. + list_case True + (%(v6::bool list) v7::bool list list. + alg_order v2 v6 & alg_sorted (v6 # v7))))" lemma alg_sorted_primitive_def: "alg_sorted = -WFREC (SOME R. WF R & (ALL x z y. R (y # z) (x # y # z))) - (%alg_sorted. +WFREC + (SOME R::bool list list => bool list list => bool. + WF R & + (ALL (x::bool list) (z::bool list list) y::bool list. + R (y # z) (x # y # z))) + (%alg_sorted::bool list list => bool. list_case True - (%v2. list_case True - (%v6 v7. alg_order v2 v6 & alg_sorted (v6 # v7))))" + (%v2::bool list. + list_case True + (%(v6::bool list) v7::bool list list. + alg_order v2 v6 & alg_sorted (v6 # v7))))" by (import prob_canon alg_sorted_primitive_def) lemma alg_sorted_ind: "(All::((bool list list => bool) => bool) => bool) @@ -712,8 +756,9 @@ ((All::(bool list list => bool) => bool) P))" by (import prob_canon alg_sorted_ind) -lemma alg_sorted_def: "alg_sorted (x # y # z) = (alg_order x y & alg_sorted (y # z)) & -alg_sorted [v] = True & alg_sorted [] = True" +lemma alg_sorted_def: "alg_sorted ((x::bool list) # (y::bool list) # (z::bool list list)) = +(alg_order x y & alg_sorted (y # z)) & +alg_sorted [v::bool list] = True & alg_sorted [] = True" by (import prob_canon alg_sorted_def) consts @@ -721,18 +766,30 @@ defs alg_prefixfree_primdef: "alg_prefixfree == -WFREC (SOME R. WF R & (ALL x z y. R (y # z) (x # y # z))) - (%alg_prefixfree. +WFREC + (SOME R::bool list list => bool list list => bool. + WF R & + (ALL (x::bool list) (z::bool list list) y::bool list. + R (y # z) (x # y # z))) + (%alg_prefixfree::bool list list => bool. list_case True - (%v2. list_case True - (%v6 v7. ~ IS_PREFIX v6 v2 & alg_prefixfree (v6 # v7))))" + (%v2::bool list. + list_case True + (%(v6::bool list) v7::bool list list. + ~ IS_PREFIX v6 v2 & alg_prefixfree (v6 # v7))))" lemma alg_prefixfree_primitive_def: "alg_prefixfree = -WFREC (SOME R. WF R & (ALL x z y. R (y # z) (x # y # z))) - (%alg_prefixfree. +WFREC + (SOME R::bool list list => bool list list => bool. + WF R & + (ALL (x::bool list) (z::bool list list) y::bool list. + R (y # z) (x # y # z))) + (%alg_prefixfree::bool list list => bool. list_case True - (%v2. list_case True - (%v6 v7. ~ IS_PREFIX v6 v2 & alg_prefixfree (v6 # v7))))" + (%v2::bool list. + list_case True + (%(v6::bool list) v7::bool list list. + ~ IS_PREFIX v6 v2 & alg_prefixfree (v6 # v7))))" by (import prob_canon alg_prefixfree_primitive_def) lemma alg_prefixfree_ind: "(All::((bool list list => bool) => bool) => bool) @@ -763,8 +820,9 @@ ((All::(bool list list => bool) => bool) P))" by (import prob_canon alg_prefixfree_ind) -lemma alg_prefixfree_def: "alg_prefixfree (x # y # z) = (~ IS_PREFIX y x & alg_prefixfree (y # z)) & -alg_prefixfree [v] = True & alg_prefixfree [] = True" +lemma alg_prefixfree_def: "alg_prefixfree ((x::bool list) # (y::bool list) # (z::bool list list)) = +(~ IS_PREFIX y x & alg_prefixfree (y # z)) & +alg_prefixfree [v::bool list] = True & alg_prefixfree [] = True" by (import prob_canon alg_prefixfree_def) consts @@ -772,18 +830,30 @@ defs alg_twinfree_primdef: "alg_twinfree == -WFREC (SOME R. WF R & (ALL x z y. R (y # z) (x # y # z))) - (%alg_twinfree. +WFREC + (SOME R::bool list list => bool list list => bool. + WF R & + (ALL (x::bool list) (z::bool list list) y::bool list. + R (y # z) (x # y # z))) + (%alg_twinfree::bool list list => bool. list_case True - (%v2. list_case True - (%v6 v7. ~ alg_twin v2 v6 & alg_twinfree (v6 # v7))))" + (%v2::bool list. + list_case True + (%(v6::bool list) v7::bool list list. + ~ alg_twin v2 v6 & alg_twinfree (v6 # v7))))" lemma alg_twinfree_primitive_def: "alg_twinfree = -WFREC (SOME R. WF R & (ALL x z y. R (y # z) (x # y # z))) - (%alg_twinfree. +WFREC + (SOME R::bool list list => bool list list => bool. + WF R & + (ALL (x::bool list) (z::bool list list) y::bool list. + R (y # z) (x # y # z))) + (%alg_twinfree::bool list list => bool. list_case True - (%v2. list_case True - (%v6 v7. ~ alg_twin v2 v6 & alg_twinfree (v6 # v7))))" + (%v2::bool list. + list_case True + (%(v6::bool list) v7::bool list list. + ~ alg_twin v2 v6 & alg_twinfree (v6 # v7))))" by (import prob_canon alg_twinfree_primitive_def) lemma alg_twinfree_ind: "(All::((bool list list => bool) => bool) => bool) @@ -814,24 +884,29 @@ ((All::(bool list list => bool) => bool) P))" by (import prob_canon alg_twinfree_ind) -lemma alg_twinfree_def: "alg_twinfree (x # y # z) = (~ alg_twin x y & alg_twinfree (y # z)) & -alg_twinfree [v] = True & alg_twinfree [] = True" +lemma alg_twinfree_def: "alg_twinfree ((x::bool list) # (y::bool list) # (z::bool list list)) = +(~ alg_twin x y & alg_twinfree (y # z)) & +alg_twinfree [v::bool list] = True & alg_twinfree [] = True" by (import prob_canon alg_twinfree_def) consts alg_longest :: "bool list list => nat" defs - alg_longest_primdef: "alg_longest == FOLDR (%h t. if t <= length h then length h else t) 0" - -lemma alg_longest_def: "alg_longest = FOLDR (%h t. if t <= length h then length h else t) 0" + alg_longest_primdef: "alg_longest == +FOLDR (%(h::bool list) t::nat. if t <= length h then length h else t) + (0::nat)" + +lemma alg_longest_def: "alg_longest = +FOLDR (%(h::bool list) t::nat. if t <= length h then length h else t) + (0::nat)" by (import prob_canon alg_longest_def) consts alg_canon_prefs :: "bool list => bool list list => bool list list" -specification (alg_canon_prefs_primdef: alg_canon_prefs) alg_canon_prefs_def: "(ALL l. alg_canon_prefs l [] = [l]) & -(ALL l h t. +specification (alg_canon_prefs_primdef: alg_canon_prefs) alg_canon_prefs_def: "(ALL l::bool list. alg_canon_prefs l [] = [l]) & +(ALL (l::bool list) (h::bool list) t::bool list list. alg_canon_prefs l (h # t) = (if IS_PREFIX h l then alg_canon_prefs l t else l # h # t))" by (import prob_canon alg_canon_prefs_def) @@ -839,8 +914,8 @@ consts alg_canon_find :: "bool list => bool list list => bool list list" -specification (alg_canon_find_primdef: alg_canon_find) alg_canon_find_def: "(ALL l. alg_canon_find l [] = [l]) & -(ALL l h t. +specification (alg_canon_find_primdef: alg_canon_find) alg_canon_find_def: "(ALL l::bool list. alg_canon_find l [] = [l]) & +(ALL (l::bool list) (h::bool list) t::bool list list. alg_canon_find l (h # t) = (if alg_order h l then if IS_PREFIX l h then h # t else h # alg_canon_find l t @@ -859,8 +934,8 @@ consts alg_canon_merge :: "bool list => bool list list => bool list list" -specification (alg_canon_merge_primdef: alg_canon_merge) alg_canon_merge_def: "(ALL l. alg_canon_merge l [] = [l]) & -(ALL l h t. +specification (alg_canon_merge_primdef: alg_canon_merge) alg_canon_merge_def: "(ALL l::bool list. alg_canon_merge l [] = [l]) & +(ALL (l::bool list) (h::bool list) t::bool list list. alg_canon_merge l (h # t) = (if alg_twin l h then alg_canon_merge (butlast h) t else l # h # t))" by (import prob_canon alg_canon_merge_def) @@ -878,34 +953,35 @@ alg_canon :: "bool list list => bool list list" defs - alg_canon_primdef: "alg_canon == %l. alg_canon2 (alg_canon1 l)" - -lemma alg_canon_def: "ALL l. alg_canon l = alg_canon2 (alg_canon1 l)" + alg_canon_primdef: "alg_canon == %l::bool list list. alg_canon2 (alg_canon1 l)" + +lemma alg_canon_def: "ALL l::bool list list. alg_canon l = alg_canon2 (alg_canon1 l)" by (import prob_canon alg_canon_def) consts algebra_canon :: "bool list list => bool" defs - algebra_canon_primdef: "algebra_canon == %l. alg_canon l = l" - -lemma algebra_canon_def: "ALL l. algebra_canon l = (alg_canon l = l)" + algebra_canon_primdef: "algebra_canon == %l::bool list list. alg_canon l = l" + +lemma algebra_canon_def: "ALL l::bool list list. algebra_canon l = (alg_canon l = l)" by (import prob_canon algebra_canon_def) -lemma ALG_TWIN_NIL: "ALL l. ~ alg_twin l [] & ~ alg_twin [] l" +lemma ALG_TWIN_NIL: "ALL l::bool list. ~ alg_twin l [] & ~ alg_twin [] l" by (import prob_canon ALG_TWIN_NIL) -lemma ALG_TWIN_SING: "ALL x l. +lemma ALG_TWIN_SING: "ALL (x::bool) l::bool list. alg_twin [x] l = (x = True & l = [False]) & alg_twin l [x] = (l = [True] & x = False)" by (import prob_canon ALG_TWIN_SING) -lemma ALG_TWIN_CONS: "ALL x y z h t. +lemma ALG_TWIN_CONS: "ALL (x::bool) (y::bool) (z::bool list) (h::bool) t::bool list. alg_twin (x # y # z) (h # t) = (x = h & alg_twin (y # z) t) & alg_twin (h # t) (x # y # z) = (x = h & alg_twin t (y # z))" by (import prob_canon ALG_TWIN_CONS) -lemma ALG_TWIN_REDUCE: "ALL h t t'. alg_twin (h # t) (h # t') = alg_twin t t'" +lemma ALG_TWIN_REDUCE: "ALL (h::bool) (t::bool list) t'::bool list. + alg_twin (h # t) (h # t') = alg_twin t t'" by (import prob_canon ALG_TWIN_REDUCE) lemma ALG_TWINS_PREFIX: "(All::(bool list => bool) => bool) @@ -924,10 +1000,10 @@ l))))))" by (import prob_canon ALG_TWINS_PREFIX) -lemma ALG_ORDER_NIL: "ALL x. alg_order [] x & alg_order x [] = (x = [])" +lemma ALG_ORDER_NIL: "ALL x::bool list. alg_order [] x & alg_order x [] = (x = [])" by (import prob_canon ALG_ORDER_NIL) -lemma ALG_ORDER_REFL: "ALL x. alg_order x x" +lemma ALG_ORDER_REFL: "ALL x::bool list. alg_order x x" by (import prob_canon ALG_ORDER_REFL) lemma ALG_ORDER_ANTISYM: "(All::(bool list => bool) => bool) @@ -954,7 +1030,7 @@ ((alg_order::bool list => bool list => bool) x z))))" by (import prob_canon ALG_ORDER_TRANS) -lemma ALG_ORDER_TOTAL: "ALL x y. alg_order x y | alg_order y x" +lemma ALG_ORDER_TOTAL: "ALL (x::bool list) y::bool list. alg_order x y | alg_order y x" by (import prob_canon ALG_ORDER_TOTAL) lemma ALG_ORDER_PREFIX: "(All::(bool list => bool) => bool) @@ -1007,7 +1083,7 @@ ((IS_PREFIX::bool list => bool list => bool) x z)))))" by (import prob_canon ALG_ORDER_PREFIX_TRANS) -lemma ALG_ORDER_SNOC: "ALL x l. ~ alg_order (SNOC x l) l" +lemma ALG_ORDER_SNOC: "ALL (x::bool) l::bool list. ~ alg_order (SNOC x l) l" by (import prob_canon ALG_ORDER_SNOC) lemma ALG_SORTED_MIN: "(All::(bool list => bool) => bool) @@ -1066,15 +1142,15 @@ z)))))" by (import prob_canon ALG_SORTED_MONO) -lemma ALG_SORTED_TLS: "ALL l b. alg_sorted (map (op # b) l) = alg_sorted l" +lemma ALG_SORTED_TLS: "ALL (l::bool list list) b::bool. alg_sorted (map (op # b) l) = alg_sorted l" by (import prob_canon ALG_SORTED_TLS) -lemma ALG_SORTED_STEP: "ALL l1 l2. +lemma ALG_SORTED_STEP: "ALL (l1::bool list list) l2::bool list list. alg_sorted (map (op # True) l1 @ map (op # False) l2) = (alg_sorted l1 & alg_sorted l2)" by (import prob_canon ALG_SORTED_STEP) -lemma ALG_SORTED_APPEND: "ALL h h' t t'. +lemma ALG_SORTED_APPEND: "ALL (h::bool list) (h'::bool list) (t::bool list list) t'::bool list list. alg_sorted ((h # t) @ h' # t') = (alg_sorted (h # t) & alg_sorted (h' # t') & alg_order (last (h # t)) h')" by (import prob_canon ALG_SORTED_APPEND) @@ -1144,15 +1220,16 @@ x)))))))" by (import prob_canon ALG_PREFIXFREE_ELT) -lemma ALG_PREFIXFREE_TLS: "ALL l b. alg_prefixfree (map (op # b) l) = alg_prefixfree l" +lemma ALG_PREFIXFREE_TLS: "ALL (l::bool list list) b::bool. + alg_prefixfree (map (op # b) l) = alg_prefixfree l" by (import prob_canon ALG_PREFIXFREE_TLS) -lemma ALG_PREFIXFREE_STEP: "ALL l1 l2. +lemma ALG_PREFIXFREE_STEP: "ALL (l1::bool list list) l2::bool list list. alg_prefixfree (map (op # True) l1 @ map (op # False) l2) = (alg_prefixfree l1 & alg_prefixfree l2)" by (import prob_canon ALG_PREFIXFREE_STEP) -lemma ALG_PREFIXFREE_APPEND: "ALL h h' t t'. +lemma ALG_PREFIXFREE_APPEND: "ALL (h::bool list) (h'::bool list) (t::bool list list) t'::bool list list. alg_prefixfree ((h # t) @ h' # t') = (alg_prefixfree (h # t) & alg_prefixfree (h' # t') & ~ IS_PREFIX h' (last (h # t)))" @@ -1182,7 +1259,8 @@ ((alg_twinfree::bool list list => bool) t)))" by (import prob_canon ALG_TWINFREE_TL) -lemma ALG_TWINFREE_TLS: "ALL l b. alg_twinfree (map (op # b) l) = alg_twinfree l" +lemma ALG_TWINFREE_TLS: "ALL (l::bool list list) b::bool. + alg_twinfree (map (op # b) l) = alg_twinfree l" by (import prob_canon ALG_TWINFREE_TLS) lemma ALG_TWINFREE_STEP1: "(All::(bool list list => bool) => bool) @@ -1258,21 +1336,22 @@ ((alg_twinfree::bool list list => bool) l2)))))" by (import prob_canon ALG_TWINFREE_STEP) -lemma ALG_LONGEST_HD: "ALL h t. length h <= alg_longest (h # t)" +lemma ALG_LONGEST_HD: "ALL (h::bool list) t::bool list list. length h <= alg_longest (h # t)" by (import prob_canon ALG_LONGEST_HD) -lemma ALG_LONGEST_TL: "ALL h t. alg_longest t <= alg_longest (h # t)" +lemma ALG_LONGEST_TL: "ALL (h::bool list) t::bool list list. alg_longest t <= alg_longest (h # t)" by (import prob_canon ALG_LONGEST_TL) -lemma ALG_LONGEST_TLS: "ALL h t b. alg_longest (map (op # b) (h # t)) = Suc (alg_longest (h # t))" +lemma ALG_LONGEST_TLS: "ALL (h::bool list) (t::bool list list) b::bool. + alg_longest (map (op # b) (h # t)) = Suc (alg_longest (h # t))" by (import prob_canon ALG_LONGEST_TLS) -lemma ALG_LONGEST_APPEND: "ALL l1 l2. +lemma ALG_LONGEST_APPEND: "ALL (l1::bool list list) l2::bool list list. alg_longest l1 <= alg_longest (l1 @ l2) & alg_longest l2 <= alg_longest (l1 @ l2)" by (import prob_canon ALG_LONGEST_APPEND) -lemma ALG_CANON_PREFS_HD: "ALL l b. hd (alg_canon_prefs l b) = l" +lemma ALG_CANON_PREFS_HD: "ALL (l::bool list) b::bool list list. hd (alg_canon_prefs l b) = l" by (import prob_canon ALG_CANON_PREFS_HD) lemma ALG_CANON_PREFS_DELETES: "(All::(bool list => bool) => bool) @@ -1332,7 +1411,7 @@ ((op #::bool list => bool list list => bool list list) l b))))" by (import prob_canon ALG_CANON_PREFS_CONSTANT) -lemma ALG_CANON_FIND_HD: "ALL l h t. +lemma ALG_CANON_FIND_HD: "ALL (l::bool list) (h::bool list) t::bool list list. hd (alg_canon_find l (h # t)) = l | hd (alg_canon_find l (h # t)) = h" by (import prob_canon ALG_CANON_FIND_HD) @@ -1395,10 +1474,10 @@ ((op #::bool list => bool list list => bool list list) l b))))" by (import prob_canon ALG_CANON_FIND_CONSTANT) -lemma ALG_CANON1_SORTED: "ALL x. alg_sorted (alg_canon1 x)" +lemma ALG_CANON1_SORTED: "ALL x::bool list list. alg_sorted (alg_canon1 x)" by (import prob_canon ALG_CANON1_SORTED) -lemma ALG_CANON1_PREFIXFREE: "ALL l. alg_prefixfree (alg_canon1 l)" +lemma ALG_CANON1_PREFIXFREE: "ALL l::bool list list. alg_prefixfree (alg_canon1 l)" by (import prob_canon ALG_CANON1_PREFIXFREE) lemma ALG_CANON1_CONSTANT: "(All::(bool list list => bool) => bool) @@ -1577,7 +1656,7 @@ ((alg_canon2::bool list list => bool list list) l) l))" by (import prob_canon ALG_CANON2_CONSTANT) -lemma ALG_CANON_SORTED_PREFIXFREE_TWINFREE: "ALL l. +lemma ALG_CANON_SORTED_PREFIXFREE_TWINFREE: "ALL l::bool list list. alg_sorted (alg_canon l) & alg_prefixfree (alg_canon l) & alg_twinfree (alg_canon l)" by (import prob_canon ALG_CANON_SORTED_PREFIXFREE_TWINFREE) @@ -1593,16 +1672,19 @@ ((alg_canon::bool list list => bool list list) l) l))" by (import prob_canon ALG_CANON_CONSTANT) -lemma ALG_CANON_IDEMPOT: "ALL l. alg_canon (alg_canon l) = alg_canon l" +lemma ALG_CANON_IDEMPOT: "ALL l::bool list list. alg_canon (alg_canon l) = alg_canon l" by (import prob_canon ALG_CANON_IDEMPOT) -lemma ALGEBRA_CANON_DEF_ALT: "ALL l. algebra_canon l = (alg_sorted l & alg_prefixfree l & alg_twinfree l)" +lemma ALGEBRA_CANON_DEF_ALT: "ALL l::bool list list. + algebra_canon l = (alg_sorted l & alg_prefixfree l & alg_twinfree l)" by (import prob_canon ALGEBRA_CANON_DEF_ALT) -lemma ALGEBRA_CANON_BASIC: "algebra_canon [] & algebra_canon [[]] & (ALL x. algebra_canon [x])" +lemma ALGEBRA_CANON_BASIC: "algebra_canon [] & +algebra_canon [[]] & (ALL x::bool list. algebra_canon [x])" by (import prob_canon ALGEBRA_CANON_BASIC) -lemma ALG_CANON_BASIC: "alg_canon [] = [] & alg_canon [[]] = [[]] & (ALL x. alg_canon [x] = [x])" +lemma ALG_CANON_BASIC: "alg_canon [] = [] & +alg_canon [[]] = [[]] & (ALL x::bool list. alg_canon [x] = [x])" by (import prob_canon ALG_CANON_BASIC) lemma ALGEBRA_CANON_TL: "(All::(bool list => bool) => bool) @@ -1615,10 +1697,11 @@ ((algebra_canon::bool list list => bool) t)))" by (import prob_canon ALGEBRA_CANON_TL) -lemma ALGEBRA_CANON_NIL_MEM: "ALL l. (algebra_canon l & [] mem l) = (l = [[]])" +lemma ALGEBRA_CANON_NIL_MEM: "ALL l::bool list list. (algebra_canon l & [] mem l) = (l = [[]])" by (import prob_canon ALGEBRA_CANON_NIL_MEM) -lemma ALGEBRA_CANON_TLS: "ALL l b. algebra_canon (map (op # b) l) = algebra_canon l" +lemma ALGEBRA_CANON_TLS: "ALL (l::bool list list) b::bool. + algebra_canon (map (op # b) l) = algebra_canon l" by (import prob_canon ALGEBRA_CANON_TLS) lemma ALGEBRA_CANON_STEP1: "(All::(bool list list => bool) => bool) @@ -1828,10 +1911,12 @@ ((algebra_canon::bool list list => bool) l) (P l))))" by (import prob_canon ALGEBRA_CANON_INDUCTION) -lemma MEM_NIL_STEP: "ALL l1 l2. ~ [] mem map (op # True) l1 @ map (op # False) l2" +lemma MEM_NIL_STEP: "ALL (l1::bool list list) l2::bool list list. + ~ [] mem map (op # True) l1 @ map (op # False) l2" by (import prob_canon MEM_NIL_STEP) -lemma ALG_SORTED_PREFIXFREE_MEM_NIL: "ALL l. (alg_sorted l & alg_prefixfree l & [] mem l) = (l = [[]])" +lemma ALG_SORTED_PREFIXFREE_MEM_NIL: "ALL l::bool list list. + (alg_sorted l & alg_prefixfree l & [] mem l) = (l = [[]])" by (import prob_canon ALG_SORTED_PREFIXFREE_MEM_NIL) lemma ALG_SORTED_PREFIXFREE_EQUALITY: "(All::(bool list list => bool) => bool) @@ -1863,33 +1948,34 @@ SHD :: "(nat => bool) => bool" defs - SHD_primdef: "SHD == %f. f 0" - -lemma SHD_def: "ALL f. SHD f = f 0" + SHD_primdef: "SHD == %f::nat => bool. f (0::nat)" + +lemma SHD_def: "ALL f::nat => bool. SHD f = f (0::nat)" by (import boolean_sequence SHD_def) consts STL :: "(nat => bool) => nat => bool" defs - STL_primdef: "STL == %f n. f (Suc n)" - -lemma STL_def: "ALL f n. STL f n = f (Suc n)" + STL_primdef: "STL == %(f::nat => bool) n::nat. f (Suc n)" + +lemma STL_def: "ALL (f::nat => bool) n::nat. STL f n = f (Suc n)" by (import boolean_sequence STL_def) consts SCONS :: "bool => (nat => bool) => nat => bool" -specification (SCONS_primdef: SCONS) SCONS_def: "(ALL h t. SCONS h t 0 = h) & (ALL h t n. SCONS h t (Suc n) = t n)" +specification (SCONS_primdef: SCONS) SCONS_def: "(ALL (h::bool) t::nat => bool. SCONS h t (0::nat) = h) & +(ALL (h::bool) (t::nat => bool) n::nat. SCONS h t (Suc n) = t n)" by (import boolean_sequence SCONS_def) consts SDEST :: "(nat => bool) => bool * (nat => bool)" defs - SDEST_primdef: "SDEST == %s. (SHD s, STL s)" - -lemma SDEST_def: "SDEST = (%s. (SHD s, STL s))" + SDEST_primdef: "SDEST == %s::nat => bool. (SHD s, STL s)" + +lemma SDEST_def: "SDEST = (%s::nat => bool. (SHD s, STL s))" by (import boolean_sequence SDEST_def) consts @@ -1904,32 +1990,32 @@ consts STAKE :: "nat => (nat => bool) => bool list" -specification (STAKE_primdef: STAKE) STAKE_def: "(ALL s. STAKE 0 s = []) & -(ALL n s. STAKE (Suc n) s = SHD s # STAKE n (STL s))" +specification (STAKE_primdef: STAKE) STAKE_def: "(ALL s::nat => bool. STAKE (0::nat) s = []) & +(ALL (n::nat) s::nat => bool. STAKE (Suc n) s = SHD s # STAKE n (STL s))" by (import boolean_sequence STAKE_def) consts SDROP :: "nat => (nat => bool) => nat => bool" -specification (SDROP_primdef: SDROP) SDROP_def: "SDROP 0 = I & (ALL n. SDROP (Suc n) = SDROP n o STL)" +specification (SDROP_primdef: SDROP) SDROP_def: "SDROP (0::nat) = I & (ALL n::nat. SDROP (Suc n) = SDROP n o STL)" by (import boolean_sequence SDROP_def) -lemma SCONS_SURJ: "ALL x. EX xa t. x = SCONS xa t" +lemma SCONS_SURJ: "ALL x::nat => bool. EX (xa::bool) t::nat => bool. x = SCONS xa t" by (import boolean_sequence SCONS_SURJ) -lemma SHD_STL_ISO: "ALL h t. EX x. SHD x = h & STL x = t" +lemma SHD_STL_ISO: "ALL (h::bool) t::nat => bool. EX x::nat => bool. SHD x = h & STL x = t" by (import boolean_sequence SHD_STL_ISO) -lemma SHD_SCONS: "ALL h t. SHD (SCONS h t) = h" +lemma SHD_SCONS: "ALL (h::bool) t::nat => bool. SHD (SCONS h t) = h" by (import boolean_sequence SHD_SCONS) -lemma STL_SCONS: "ALL h t. STL (SCONS h t) = t" +lemma STL_SCONS: "ALL (h::bool) t::nat => bool. STL (SCONS h t) = t" by (import boolean_sequence STL_SCONS) -lemma SHD_SCONST: "ALL b. SHD (SCONST b) = b" +lemma SHD_SCONST: "ALL b::bool. SHD (SCONST b) = b" by (import boolean_sequence SHD_SCONST) -lemma STL_SCONST: "ALL b. STL (SCONST b) = SCONST b" +lemma STL_SCONST: "ALL b::bool. STL (SCONST b) = SCONST b" by (import boolean_sequence STL_SCONST) ;end_setup @@ -1939,15 +2025,16 @@ consts alg_embed :: "bool list => (nat => bool) => bool" -specification (alg_embed_primdef: alg_embed) alg_embed_def: "(ALL s. alg_embed [] s = True) & -(ALL h t s. alg_embed (h # t) s = (h = SHD s & alg_embed t (STL s)))" +specification (alg_embed_primdef: alg_embed) alg_embed_def: "(ALL s::nat => bool. alg_embed [] s = True) & +(ALL (h::bool) (t::bool list) s::nat => bool. + alg_embed (h # t) s = (h = SHD s & alg_embed t (STL s)))" by (import prob_algebra alg_embed_def) consts algebra_embed :: "bool list list => (nat => bool) => bool" specification (algebra_embed_primdef: algebra_embed) algebra_embed_def: "algebra_embed [] = EMPTY & -(ALL h t. +(ALL (h::bool list) t::bool list list. algebra_embed (h # t) = pred_set.UNION (alg_embed h) (algebra_embed t))" by (import prob_algebra algebra_embed_def) @@ -1955,29 +2042,36 @@ measurable :: "((nat => bool) => bool) => bool" defs - measurable_primdef: "measurable == %s. EX b. s = algebra_embed b" - -lemma measurable_def: "ALL s. measurable s = (EX b. s = algebra_embed b)" + measurable_primdef: "measurable == +%s::(nat => bool) => bool. EX b::bool list list. s = algebra_embed b" + +lemma measurable_def: "ALL s::(nat => bool) => bool. + measurable s = (EX b::bool list list. s = algebra_embed b)" by (import prob_algebra measurable_def) -lemma HALVES_INTER: "pred_set.INTER (%x. SHD x = True) (%x. SHD x = False) = EMPTY" +lemma HALVES_INTER: "pred_set.INTER (%x::nat => bool. SHD x = True) + (%x::nat => bool. SHD x = False) = +EMPTY" by (import prob_algebra HALVES_INTER) -lemma INTER_STL: "ALL p q. pred_set.INTER p q o STL = pred_set.INTER (p o STL) (q o STL)" +lemma INTER_STL: "ALL (p::(nat => bool) => bool) q::(nat => bool) => bool. + pred_set.INTER p q o STL = pred_set.INTER (p o STL) (q o STL)" by (import prob_algebra INTER_STL) -lemma COMPL_SHD: "ALL b. COMPL (%x. SHD x = b) = (%x. SHD x = (~ b))" +lemma COMPL_SHD: "ALL b::bool. + COMPL (%x::nat => bool. SHD x = b) = (%x::nat => bool. SHD x = (~ b))" by (import prob_algebra COMPL_SHD) lemma ALG_EMBED_BASIC: "alg_embed [] = pred_set.UNIV & -(ALL h t. - alg_embed (h # t) = pred_set.INTER (%x. SHD x = h) (alg_embed t o STL))" +(ALL (h::bool) t::bool list. + alg_embed (h # t) = + pred_set.INTER (%x::nat => bool. SHD x = h) (alg_embed t o STL))" by (import prob_algebra ALG_EMBED_BASIC) -lemma ALG_EMBED_NIL: "ALL c. All (alg_embed c) = (c = [])" +lemma ALG_EMBED_NIL: "ALL c::bool list. All (alg_embed c) = (c = [])" by (import prob_algebra ALG_EMBED_NIL) -lemma ALG_EMBED_POPULATED: "ALL b. Ex (alg_embed b)" +lemma ALG_EMBED_POPULATED: "ALL b::bool list. Ex (alg_embed b)" by (import prob_algebra ALG_EMBED_POPULATED) lemma ALG_EMBED_PREFIX: "(All::(bool list => bool) => bool) @@ -1995,17 +2089,18 @@ ((IS_PREFIX::bool list => bool list => bool) c b)))))" by (import prob_algebra ALG_EMBED_PREFIX) -lemma ALG_EMBED_PREFIX_SUBSET: "ALL b c. SUBSET (alg_embed b) (alg_embed c) = IS_PREFIX b c" +lemma ALG_EMBED_PREFIX_SUBSET: "ALL (b::bool list) c::bool list. + SUBSET (alg_embed b) (alg_embed c) = IS_PREFIX b c" by (import prob_algebra ALG_EMBED_PREFIX_SUBSET) -lemma ALG_EMBED_TWINS: "ALL l. +lemma ALG_EMBED_TWINS: "ALL l::bool list. pred_set.UNION (alg_embed (SNOC True l)) (alg_embed (SNOC False l)) = alg_embed l" by (import prob_algebra ALG_EMBED_TWINS) lemma ALGEBRA_EMBED_BASIC: "algebra_embed [] = EMPTY & algebra_embed [[]] = pred_set.UNIV & -(ALL b. algebra_embed [[b]] = (%s. SHD s = b))" +(ALL b::bool. algebra_embed [[b]] = (%s::nat => bool. SHD s = b))" by (import prob_algebra ALGEBRA_EMBED_BASIC) lemma ALGEBRA_EMBED_MEM: "(All::(bool list list => bool) => bool) @@ -2021,31 +2116,35 @@ ((alg_embed::bool list => (nat => bool) => bool) l x)))))" by (import prob_algebra ALGEBRA_EMBED_MEM) -lemma ALGEBRA_EMBED_APPEND: "ALL l1 l2. +lemma ALGEBRA_EMBED_APPEND: "ALL (l1::bool list list) l2::bool list list. algebra_embed (l1 @ l2) = pred_set.UNION (algebra_embed l1) (algebra_embed l2)" by (import prob_algebra ALGEBRA_EMBED_APPEND) -lemma ALGEBRA_EMBED_TLS: "ALL l b. - algebra_embed (map (op # b) l) (SCONS h t) = (h = b & algebra_embed l t)" +lemma ALGEBRA_EMBED_TLS: "ALL (l::bool list list) b::bool. + algebra_embed (map (op # b) l) (SCONS (h::bool) (t::nat => bool)) = + (h = b & algebra_embed l t)" by (import prob_algebra ALGEBRA_EMBED_TLS) -lemma ALG_CANON_PREFS_EMBED: "ALL l b. algebra_embed (alg_canon_prefs l b) = algebra_embed (l # b)" +lemma ALG_CANON_PREFS_EMBED: "ALL (l::bool list) b::bool list list. + algebra_embed (alg_canon_prefs l b) = algebra_embed (l # b)" by (import prob_algebra ALG_CANON_PREFS_EMBED) -lemma ALG_CANON_FIND_EMBED: "ALL l b. algebra_embed (alg_canon_find l b) = algebra_embed (l # b)" +lemma ALG_CANON_FIND_EMBED: "ALL (l::bool list) b::bool list list. + algebra_embed (alg_canon_find l b) = algebra_embed (l # b)" by (import prob_algebra ALG_CANON_FIND_EMBED) -lemma ALG_CANON1_EMBED: "ALL x. algebra_embed (alg_canon1 x) = algebra_embed x" +lemma ALG_CANON1_EMBED: "ALL x::bool list list. algebra_embed (alg_canon1 x) = algebra_embed x" by (import prob_algebra ALG_CANON1_EMBED) -lemma ALG_CANON_MERGE_EMBED: "ALL l b. algebra_embed (alg_canon_merge l b) = algebra_embed (l # b)" +lemma ALG_CANON_MERGE_EMBED: "ALL (l::bool list) b::bool list list. + algebra_embed (alg_canon_merge l b) = algebra_embed (l # b)" by (import prob_algebra ALG_CANON_MERGE_EMBED) -lemma ALG_CANON2_EMBED: "ALL x. algebra_embed (alg_canon2 x) = algebra_embed x" +lemma ALG_CANON2_EMBED: "ALL x::bool list list. algebra_embed (alg_canon2 x) = algebra_embed x" by (import prob_algebra ALG_CANON2_EMBED) -lemma ALG_CANON_EMBED: "ALL l. algebra_embed (alg_canon l) = algebra_embed l" +lemma ALG_CANON_EMBED: "ALL l::bool list list. algebra_embed (alg_canon l) = algebra_embed l" by (import prob_algebra ALG_CANON_EMBED) lemma ALGEBRA_CANON_UNIV: "(All::(bool list list => bool) => bool) @@ -2061,7 +2160,8 @@ ([]::bool list) ([]::bool list list)))))" by (import prob_algebra ALGEBRA_CANON_UNIV) -lemma ALG_CANON_REP: "ALL b c. (alg_canon b = alg_canon c) = (algebra_embed b = algebra_embed c)" +lemma ALG_CANON_REP: "ALL (b::bool list list) c::bool list list. + (alg_canon b = alg_canon c) = (algebra_embed b = algebra_embed c)" by (import prob_algebra ALG_CANON_REP) lemma ALGEBRA_CANON_EMBED_EMPTY: "(All::(bool list list => bool) => bool) @@ -2090,20 +2190,22 @@ ([]::bool list) ([]::bool list list)))))" by (import prob_algebra ALGEBRA_CANON_EMBED_UNIV) -lemma MEASURABLE_ALGEBRA: "ALL b. measurable (algebra_embed b)" +lemma MEASURABLE_ALGEBRA: "ALL b::bool list list. measurable (algebra_embed b)" by (import prob_algebra MEASURABLE_ALGEBRA) lemma MEASURABLE_BASIC: "measurable EMPTY & -measurable pred_set.UNIV & (ALL b. measurable (%s. SHD s = b))" +measurable pred_set.UNIV & +(ALL b::bool. measurable (%s::nat => bool. SHD s = b))" by (import prob_algebra MEASURABLE_BASIC) -lemma MEASURABLE_SHD: "ALL b. measurable (%s. SHD s = b)" +lemma MEASURABLE_SHD: "ALL b::bool. measurable (%s::nat => bool. SHD s = b)" by (import prob_algebra MEASURABLE_SHD) -lemma ALGEBRA_EMBED_COMPL: "ALL l. EX l'. COMPL (algebra_embed l) = algebra_embed l'" +lemma ALGEBRA_EMBED_COMPL: "ALL l::bool list list. + EX l'::bool list list. COMPL (algebra_embed l) = algebra_embed l'" by (import prob_algebra ALGEBRA_EMBED_COMPL) -lemma MEASURABLE_COMPL: "ALL s. measurable (COMPL s) = measurable s" +lemma MEASURABLE_COMPL: "ALL s::(nat => bool) => bool. measurable (COMPL s) = measurable s" by (import prob_algebra MEASURABLE_COMPL) lemma MEASURABLE_UNION: "(All::(((nat => bool) => bool) => bool) => bool) @@ -2136,28 +2238,30 @@ s t))))" by (import prob_algebra MEASURABLE_INTER) -lemma MEASURABLE_STL: "ALL p. measurable (p o STL) = measurable p" +lemma MEASURABLE_STL: "ALL p::(nat => bool) => bool. measurable (p o STL) = measurable p" by (import prob_algebra MEASURABLE_STL) -lemma MEASURABLE_SDROP: "ALL n p. measurable (p o SDROP n) = measurable p" +lemma MEASURABLE_SDROP: "ALL (n::nat) p::(nat => bool) => bool. + measurable (p o SDROP n) = measurable p" by (import prob_algebra MEASURABLE_SDROP) -lemma MEASURABLE_INTER_HALVES: "ALL p. - (measurable (pred_set.INTER (%x. SHD x = True) p) & - measurable (pred_set.INTER (%x. SHD x = False) p)) = +lemma MEASURABLE_INTER_HALVES: "ALL p::(nat => bool) => bool. + (measurable (pred_set.INTER (%x::nat => bool. SHD x = True) p) & + measurable (pred_set.INTER (%x::nat => bool. SHD x = False) p)) = measurable p" by (import prob_algebra MEASURABLE_INTER_HALVES) -lemma MEASURABLE_HALVES: "ALL p q. +lemma MEASURABLE_HALVES: "ALL (p::(nat => bool) => bool) q::(nat => bool) => bool. measurable - (pred_set.UNION (pred_set.INTER (%x. SHD x = True) p) - (pred_set.INTER (%x. SHD x = False) q)) = - (measurable (pred_set.INTER (%x. SHD x = True) p) & - measurable (pred_set.INTER (%x. SHD x = False) q))" + (pred_set.UNION (pred_set.INTER (%x::nat => bool. SHD x = True) p) + (pred_set.INTER (%x::nat => bool. SHD x = False) q)) = + (measurable (pred_set.INTER (%x::nat => bool. SHD x = True) p) & + measurable (pred_set.INTER (%x::nat => bool. SHD x = False) q))" by (import prob_algebra MEASURABLE_HALVES) -lemma MEASURABLE_INTER_SHD: "ALL b p. - measurable (pred_set.INTER (%x. SHD x = b) (p o STL)) = measurable p" +lemma MEASURABLE_INTER_SHD: "ALL (b::bool) p::(nat => bool) => bool. + measurable (pred_set.INTER (%x::nat => bool. SHD x = b) (p o STL)) = + measurable p" by (import prob_algebra MEASURABLE_INTER_SHD) ;end_setup @@ -2167,8 +2271,10 @@ consts alg_measure :: "bool list list => real" -specification (alg_measure_primdef: alg_measure) alg_measure_def: "alg_measure [] = 0 & -(ALL l rest. alg_measure (l # rest) = (1 / 2) ^ length l + alg_measure rest)" +specification (alg_measure_primdef: alg_measure) alg_measure_def: "alg_measure [] = (0::real) & +(ALL (l::bool list) rest::bool list list. + alg_measure (l # rest) = + ((1::real) / (2::real)) ^ length l + alg_measure rest)" by (import prob alg_measure_def) consts @@ -2176,11 +2282,16 @@ defs algebra_measure_primdef: "algebra_measure == -%b. inf (%r. EX c. algebra_embed b = algebra_embed c & alg_measure c = r)" - -lemma algebra_measure_def: "ALL b. +%b::bool list list. + inf (%r::real. + EX c::bool list list. + algebra_embed b = algebra_embed c & alg_measure c = r)" + +lemma algebra_measure_def: "ALL b::bool list list. algebra_measure b = - inf (%r. EX c. algebra_embed b = algebra_embed c & alg_measure c = r)" + inf (%r::real. + EX c::bool list list. + algebra_embed b = algebra_embed c & alg_measure c = r)" by (import prob algebra_measure_def) consts @@ -2188,11 +2299,16 @@ defs prob_primdef: "prob == -%s. sup (%r. EX b. algebra_measure b = r & SUBSET (algebra_embed b) s)" - -lemma prob_def: "ALL s. +%s::(nat => bool) => bool. + sup (%r::real. + EX b::bool list list. + algebra_measure b = r & SUBSET (algebra_embed b) s)" + +lemma prob_def: "ALL s::(nat => bool) => bool. prob s = - sup (%r. EX b. algebra_measure b = r & SUBSET (algebra_embed b) s)" + sup (%r::real. + EX b::bool list list. + algebra_measure b = r & SUBSET (algebra_embed b) s)" by (import prob prob_def) lemma ALG_TWINS_MEASURE: "ALL l::bool list. @@ -2201,42 +2317,49 @@ ((1::real) / (2::real)) ^ length l" by (import prob ALG_TWINS_MEASURE) -lemma ALG_MEASURE_BASIC: "alg_measure [] = 0 & -alg_measure [[]] = 1 & (ALL b. alg_measure [[b]] = 1 / 2)" +lemma ALG_MEASURE_BASIC: "alg_measure [] = (0::real) & +alg_measure [[]] = (1::real) & +(ALL b::bool. alg_measure [[b]] = (1::real) / (2::real))" by (import prob ALG_MEASURE_BASIC) -lemma ALG_MEASURE_POS: "ALL l. 0 <= alg_measure l" +lemma ALG_MEASURE_POS: "ALL l::bool list list. (0::real) <= alg_measure l" by (import prob ALG_MEASURE_POS) -lemma ALG_MEASURE_APPEND: "ALL l1 l2. alg_measure (l1 @ l2) = alg_measure l1 + alg_measure l2" +lemma ALG_MEASURE_APPEND: "ALL (l1::bool list list) l2::bool list list. + alg_measure (l1 @ l2) = alg_measure l1 + alg_measure l2" by (import prob ALG_MEASURE_APPEND) -lemma ALG_MEASURE_TLS: "ALL l b. 2 * alg_measure (map (op # b) l) = alg_measure l" +lemma ALG_MEASURE_TLS: "ALL (l::bool list list) b::bool. + (2::real) * alg_measure (map (op # b) l) = alg_measure l" by (import prob ALG_MEASURE_TLS) -lemma ALG_CANON_PREFS_MONO: "ALL l b. alg_measure (alg_canon_prefs l b) <= alg_measure (l # b)" +lemma ALG_CANON_PREFS_MONO: "ALL (l::bool list) b::bool list list. + alg_measure (alg_canon_prefs l b) <= alg_measure (l # b)" by (import prob ALG_CANON_PREFS_MONO) -lemma ALG_CANON_FIND_MONO: "ALL l b. alg_measure (alg_canon_find l b) <= alg_measure (l # b)" +lemma ALG_CANON_FIND_MONO: "ALL (l::bool list) b::bool list list. + alg_measure (alg_canon_find l b) <= alg_measure (l # b)" by (import prob ALG_CANON_FIND_MONO) -lemma ALG_CANON1_MONO: "ALL x. alg_measure (alg_canon1 x) <= alg_measure x" +lemma ALG_CANON1_MONO: "ALL x::bool list list. alg_measure (alg_canon1 x) <= alg_measure x" by (import prob ALG_CANON1_MONO) -lemma ALG_CANON_MERGE_MONO: "ALL l b. alg_measure (alg_canon_merge l b) <= alg_measure (l # b)" +lemma ALG_CANON_MERGE_MONO: "ALL (l::bool list) b::bool list list. + alg_measure (alg_canon_merge l b) <= alg_measure (l # b)" by (import prob ALG_CANON_MERGE_MONO) -lemma ALG_CANON2_MONO: "ALL x. alg_measure (alg_canon2 x) <= alg_measure x" +lemma ALG_CANON2_MONO: "ALL x::bool list list. alg_measure (alg_canon2 x) <= alg_measure x" by (import prob ALG_CANON2_MONO) -lemma ALG_CANON_MONO: "ALL l. alg_measure (alg_canon l) <= alg_measure l" +lemma ALG_CANON_MONO: "ALL l::bool list list. alg_measure (alg_canon l) <= alg_measure l" by (import prob ALG_CANON_MONO) -lemma ALGEBRA_MEASURE_DEF_ALT: "ALL l. algebra_measure l = alg_measure (alg_canon l)" +lemma ALGEBRA_MEASURE_DEF_ALT: "ALL l::bool list list. algebra_measure l = alg_measure (alg_canon l)" by (import prob ALGEBRA_MEASURE_DEF_ALT) -lemma ALGEBRA_MEASURE_BASIC: "algebra_measure [] = 0 & -algebra_measure [[]] = 1 & (ALL b. algebra_measure [[b]] = 1 / 2)" +lemma ALGEBRA_MEASURE_BASIC: "algebra_measure [] = (0::real) & +algebra_measure [[]] = (1::real) & +(ALL b::bool. algebra_measure [[b]] = (1::real) / (2::real))" by (import prob ALGEBRA_MEASURE_BASIC) lemma ALGEBRA_CANON_MEASURE_MAX: "(All::(bool list list => bool) => bool) @@ -2247,7 +2370,7 @@ ((alg_measure::bool list list => real) l) (1::real)))" by (import prob ALGEBRA_CANON_MEASURE_MAX) -lemma ALGEBRA_MEASURE_MAX: "ALL l. algebra_measure l <= 1" +lemma ALGEBRA_MEASURE_MAX: "ALL l::bool list list. algebra_measure l <= (1::real)" by (import prob ALGEBRA_MEASURE_MAX) lemma ALGEBRA_MEASURE_MONO_EMBED: "(All::(bool list list => bool) => bool) @@ -2332,11 +2455,12 @@ ((alg_measure::bool list list => real) d)))))))))" by (import prob ALG_MEASURE_ADDITIVE) -lemma PROB_ALGEBRA: "ALL l. prob (algebra_embed l) = algebra_measure l" +lemma PROB_ALGEBRA: "ALL l::bool list list. prob (algebra_embed l) = algebra_measure l" by (import prob PROB_ALGEBRA) -lemma PROB_BASIC: "prob EMPTY = 0 & -prob pred_set.UNIV = 1 & (ALL b. prob (%s. SHD s = b) = 1 / 2)" +lemma PROB_BASIC: "prob EMPTY = (0::real) & +prob pred_set.UNIV = (1::real) & +(ALL b::bool. prob (%s::nat => bool. SHD s = b) = (1::real) / (2::real))" by (import prob PROB_BASIC) lemma PROB_ADDITIVE: "(All::(((nat => bool) => bool) => bool) => bool) @@ -2377,7 +2501,9 @@ ((prob::((nat => bool) => bool) => real) s))))" by (import prob PROB_COMPL) -lemma PROB_SUP_EXISTS1: "ALL s. EX x b. algebra_measure b = x & SUBSET (algebra_embed b) s" +lemma PROB_SUP_EXISTS1: "ALL s::(nat => bool) => bool. + EX (x::real) b::bool list list. + algebra_measure b = x & SUBSET (algebra_embed b) s" by (import prob PROB_SUP_EXISTS1) lemma PROB_SUP_EXISTS2: "(All::(((nat => bool) => bool) => bool) => bool) @@ -2452,7 +2578,7 @@ ((prob::((nat => bool) => bool) => real) t))))" by (import prob PROB_SUBSET_MONO) -lemma PROB_ALG: "ALL x. prob (alg_embed x) = (1 / 2) ^ length x" +lemma PROB_ALG: "ALL x::bool list. prob (alg_embed x) = ((1::real) / (2::real)) ^ length x" by (import prob PROB_ALG) lemma PROB_STL: "(All::(((nat => bool) => bool) => bool) => bool) @@ -2536,25 +2662,26 @@ ((prob::((nat => bool) => bool) => real) p)))))" by (import prob PROB_INTER_SHD) -lemma ALGEBRA_MEASURE_POS: "ALL l. 0 <= algebra_measure l" +lemma ALGEBRA_MEASURE_POS: "ALL l::bool list list. (0::real) <= algebra_measure l" by (import prob ALGEBRA_MEASURE_POS) -lemma ALGEBRA_MEASURE_RANGE: "ALL l. 0 <= algebra_measure l & algebra_measure l <= 1" +lemma ALGEBRA_MEASURE_RANGE: "ALL l::bool list list. + (0::real) <= algebra_measure l & algebra_measure l <= (1::real)" by (import prob ALGEBRA_MEASURE_RANGE) -lemma PROB_POS: "ALL p. 0 <= prob p" +lemma PROB_POS: "ALL p::(nat => bool) => bool. (0::real) <= prob p" by (import prob PROB_POS) -lemma PROB_MAX: "ALL p. prob p <= 1" +lemma PROB_MAX: "ALL p::(nat => bool) => bool. prob p <= (1::real)" by (import prob PROB_MAX) -lemma PROB_RANGE: "ALL p. 0 <= prob p & prob p <= 1" +lemma PROB_RANGE: "ALL p::(nat => bool) => bool. (0::real) <= prob p & prob p <= (1::real)" by (import prob PROB_RANGE) -lemma ABS_PROB: "ALL p. abs (prob p) = prob p" +lemma ABS_PROB: "ALL p::(nat => bool) => bool. abs (prob p) = prob p" by (import prob ABS_PROB) -lemma PROB_SHD: "ALL b. prob (%s. SHD s = b) = 1 / 2" +lemma PROB_SHD: "ALL b::bool. prob (%s::nat => bool. SHD s = b) = (1::real) / (2::real)" by (import prob PROB_SHD) lemma PROB_COMPL_LE1: "(All::(((nat => bool) => bool) => bool) => bool) @@ -2591,41 +2718,42 @@ pseudo_linear_tl :: "nat => nat => nat => nat => nat" defs - pseudo_linear_tl_primdef: "pseudo_linear_tl == %a b n x. (a * x + b) mod (2 * n + 1)" - -lemma pseudo_linear_tl_def: "ALL a b n x. pseudo_linear_tl a b n x = (a * x + b) mod (2 * n + 1)" + pseudo_linear_tl_primdef: "pseudo_linear_tl == +%(a::nat) (b::nat) (n::nat) x::nat. + (a * x + b) mod ((2::nat) * n + (1::nat))" + +lemma pseudo_linear_tl_def: "ALL (a::nat) (b::nat) (n::nat) x::nat. + pseudo_linear_tl a b n x = (a * x + b) mod ((2::nat) * n + (1::nat))" by (import prob_pseudo pseudo_linear_tl_def) -lemma PSEUDO_LINEAR1_EXECUTE: "EX x. (ALL xa. SHD (x xa) = pseudo_linear_hd xa) & - (ALL xa. - STL (x xa) = - x (pseudo_linear_tl - (NUMERAL - (NUMERAL_BIT1 - (NUMERAL_BIT1 - (NUMERAL_BIT1 - (NUMERAL_BIT2 - (NUMERAL_BIT1 (NUMERAL_BIT2 ALT_ZERO))))))) - (NUMERAL - (NUMERAL_BIT1 - (NUMERAL_BIT1 - (NUMERAL_BIT1 - (NUMERAL_BIT1 - (NUMERAL_BIT1 (NUMERAL_BIT2 ALT_ZERO))))))) - (NUMERAL - (NUMERAL_BIT1 - (NUMERAL_BIT1 - (NUMERAL_BIT1 - (NUMERAL_BIT1 - (NUMERAL_BIT2 (NUMERAL_BIT1 ALT_ZERO))))))) - xa))" +lemma PSEUDO_LINEAR1_EXECUTE: "EX x::nat => nat => bool. + (ALL xa::nat. SHD (x xa) = pseudo_linear_hd xa) & + (ALL xa::nat. + STL (x xa) = + x (pseudo_linear_tl + (NUMERAL + (NUMERAL_BIT1 + (NUMERAL_BIT1 + (NUMERAL_BIT1 + (NUMERAL_BIT2 (NUMERAL_BIT1 (NUMERAL_BIT2 ALT_ZERO))))))) + (NUMERAL + (NUMERAL_BIT1 + (NUMERAL_BIT1 + (NUMERAL_BIT1 + (NUMERAL_BIT1 (NUMERAL_BIT1 (NUMERAL_BIT2 ALT_ZERO))))))) + (NUMERAL + (NUMERAL_BIT1 + (NUMERAL_BIT1 + (NUMERAL_BIT1 + (NUMERAL_BIT1 (NUMERAL_BIT2 (NUMERAL_BIT1 ALT_ZERO))))))) + xa))" by (import prob_pseudo PSEUDO_LINEAR1_EXECUTE) consts pseudo_linear1 :: "nat => nat => bool" -specification (pseudo_linear1_primdef: pseudo_linear1) pseudo_linear1_def: "(ALL x. SHD (pseudo_linear1 x) = pseudo_linear_hd x) & -(ALL x. +specification (pseudo_linear1_primdef: pseudo_linear1) pseudo_linear1_def: "(ALL x::nat. SHD (pseudo_linear1 x) = pseudo_linear_hd x) & +(ALL x::nat. STL (pseudo_linear1 x) = pseudo_linear1 (pseudo_linear_tl @@ -2665,10 +2793,10 @@ defs indep_set_primdef: "indep_set == -%p q. measurable p & - measurable q & prob (pred_set.INTER p q) = prob p * prob q" - -lemma indep_set_def: "ALL p q. +%(p::(nat => bool) => bool) q::(nat => bool) => bool. + measurable p & measurable q & prob (pred_set.INTER p q) = prob p * prob q" + +lemma indep_set_def: "ALL (p::(nat => bool) => bool) q::(nat => bool) => bool. indep_set p q = (measurable p & measurable q & prob (pred_set.INTER p q) = prob p * prob q)" @@ -2679,9 +2807,10 @@ defs alg_cover_set_primdef: "alg_cover_set == -%l. alg_sorted l & alg_prefixfree l & algebra_embed l = pred_set.UNIV" - -lemma alg_cover_set_def: "ALL l. +%l::bool list list. + alg_sorted l & alg_prefixfree l & algebra_embed l = pred_set.UNIV" + +lemma alg_cover_set_def: "ALL l::bool list list. alg_cover_set l = (alg_sorted l & alg_prefixfree l & algebra_embed l = pred_set.UNIV)" by (import prob_indep alg_cover_set_def) @@ -2690,25 +2819,33 @@ alg_cover :: "bool list list => (nat => bool) => bool list" defs - alg_cover_primdef: "alg_cover == %l x. SOME b. b mem l & alg_embed b x" - -lemma alg_cover_def: "ALL l x. alg_cover l x = (SOME b. b mem l & alg_embed b x)" + alg_cover_primdef: "alg_cover == +%(l::bool list list) x::nat => bool. + SOME b::bool list. b mem l & alg_embed b x" + +lemma alg_cover_def: "ALL (l::bool list list) x::nat => bool. + alg_cover l x = (SOME b::bool list. b mem l & alg_embed b x)" by (import prob_indep alg_cover_def) consts - indep :: "((nat => bool) => 'a * (nat => bool)) => bool" + indep :: "((nat => bool) => 'a::type * (nat => bool)) => bool" defs indep_primdef: "indep == -%f. EX l r. +%f::(nat => bool) => 'a::type * (nat => bool). + EX (l::bool list list) r::bool list => 'a::type. + alg_cover_set l & + (ALL s::nat => bool. + f s = + (let c::bool list = alg_cover l s in (r c, SDROP (length c) s)))" + +lemma indep_def: "ALL f::(nat => bool) => 'a::type * (nat => bool). + indep f = + (EX (l::bool list list) r::bool list => 'a::type. alg_cover_set l & - (ALL s. f s = (let c = alg_cover l s in (r c, SDROP (length c) s)))" - -lemma indep_def: "ALL f. - indep f = - (EX l r. - alg_cover_set l & - (ALL s. f s = (let c = alg_cover l s in (r c, SDROP (length c) s))))" + (ALL s::nat => bool. + f s = + (let c::bool list = alg_cover l s in (r c, SDROP (length c) s))))" by (import prob_indep indep_def) lemma INDEP_SET_BASIC: "(All::(((nat => bool) => bool) => bool) => bool) @@ -2724,7 +2861,8 @@ (pred_set.UNIV::(nat => bool) => bool) p)))" by (import prob_indep INDEP_SET_BASIC) -lemma INDEP_SET_SYM: "ALL p q. indep_set p q = indep_set p q" +lemma INDEP_SET_SYM: "ALL (p::(nat => bool) => bool) q::(nat => bool) => bool. + indep_set p q = indep_set p q" by (import prob_indep INDEP_SET_SYM) lemma INDEP_SET_DISJOINT_DECOMP: "(All::(((nat => bool) => bool) => bool) => bool) @@ -2809,10 +2947,10 @@ l))))" by (import prob_indep MAP_CONS_TL_FILTER) -lemma ALG_COVER_SET_CASES_THM: "ALL l. +lemma ALG_COVER_SET_CASES_THM: "ALL l::bool list list. alg_cover_set l = (l = [[]] | - (EX x xa. + (EX (x::bool list list) xa::bool list list. alg_cover_set x & alg_cover_set xa & l = map (op # True) x @ map (op # False) xa))" by (import prob_indep ALG_COVER_SET_CASES_THM) @@ -3191,144 +3329,162 @@ q)))" by (import prob_indep INDEP_SET_LIST) -lemma INDEP_INDEP_SET: "(All::(((nat => bool) => 'a * (nat => bool)) => bool) => bool) - (%f::(nat => bool) => 'a * (nat => bool). - (All::(('a => bool) => bool) => bool) - (%p::'a => bool. +lemma INDEP_INDEP_SET: "(All::(((nat => bool) => 'a::type * (nat => bool)) => bool) => bool) + (%f::(nat => bool) => 'a::type * (nat => bool). + (All::(('a::type => bool) => bool) => bool) + (%p::'a::type => bool. (All::(((nat => bool) => bool) => bool) => bool) (%q::(nat => bool) => bool. (op -->::bool => bool => bool) ((op &::bool => bool => bool) - ((indep::((nat => bool) => 'a * (nat => bool)) => bool) f) + ((indep::((nat => bool) => 'a::type * (nat => bool)) + => bool) + f) ((measurable::((nat => bool) => bool) => bool) q)) ((indep_set::((nat => bool) => bool) => ((nat => bool) => bool) => bool) - ((op o::('a => bool) - => ((nat => bool) => 'a) => (nat => bool) => bool) - p ((op o::('a * (nat => bool) => 'a) - => ((nat => bool) => 'a * (nat => bool)) - => (nat => bool) => 'a) - (fst::'a * (nat => bool) => 'a) f)) + ((op o::('a::type => bool) + => ((nat => bool) => 'a::type) + => (nat => bool) => bool) + p ((op o::('a::type * (nat => bool) => 'a::type) + => ((nat => bool) => 'a::type * (nat => bool)) + => (nat => bool) => 'a::type) + (fst::'a::type * (nat => bool) => 'a::type) f)) ((op o::((nat => bool) => bool) => ((nat => bool) => nat => bool) => (nat => bool) => bool) - q ((op o::('a * (nat => bool) => nat => bool) - => ((nat => bool) => 'a * (nat => bool)) + q ((op o::('a::type * (nat => bool) => nat => bool) + => ((nat => bool) => 'a::type * (nat => bool)) => (nat => bool) => nat => bool) - (snd::'a * (nat => bool) => nat => bool) f))))))" + (snd::'a::type * (nat => bool) => nat => bool) + f))))))" by (import prob_indep INDEP_INDEP_SET) -lemma INDEP_UNIT: "ALL x. indep (UNIT x)" +lemma INDEP_UNIT: "ALL x::'a::type. indep (UNIT x)" by (import prob_indep INDEP_UNIT) lemma INDEP_SDEST: "indep SDEST" by (import prob_indep INDEP_SDEST) -lemma BIND_STEP: "ALL f. BIND SDEST (%k. f o SCONS k) = f" +lemma BIND_STEP: "ALL f::(nat => bool) => 'a::type * (nat => bool). + BIND SDEST (%k::bool. f o SCONS k) = f" by (import prob_indep BIND_STEP) -lemma INDEP_BIND_SDEST: "(All::((bool => (nat => bool) => 'a * (nat => bool)) => bool) => bool) - (%f::bool => (nat => bool) => 'a * (nat => bool). +lemma INDEP_BIND_SDEST: "(All::((bool => (nat => bool) => 'a::type * (nat => bool)) => bool) => bool) + (%f::bool => (nat => bool) => 'a::type * (nat => bool). (op -->::bool => bool => bool) ((All::(bool => bool) => bool) (%x::bool. - (indep::((nat => bool) => 'a * (nat => bool)) => bool) (f x))) - ((indep::((nat => bool) => 'a * (nat => bool)) => bool) + (indep::((nat => bool) => 'a::type * (nat => bool)) => bool) + (f x))) + ((indep::((nat => bool) => 'a::type * (nat => bool)) => bool) ((BIND::((nat => bool) => bool * (nat => bool)) - => (bool => (nat => bool) => 'a * (nat => bool)) - => (nat => bool) => 'a * (nat => bool)) + => (bool => (nat => bool) => 'a::type * (nat => bool)) + => (nat => bool) => 'a::type * (nat => bool)) (SDEST::(nat => bool) => bool * (nat => bool)) f)))" by (import prob_indep INDEP_BIND_SDEST) -lemma INDEP_BIND: "(All::(((nat => bool) => 'a * (nat => bool)) => bool) => bool) - (%f::(nat => bool) => 'a * (nat => bool). - (All::(('a => (nat => bool) => 'b * (nat => bool)) => bool) => bool) - (%g::'a => (nat => bool) => 'b * (nat => bool). +lemma INDEP_BIND: "(All::(((nat => bool) => 'a::type * (nat => bool)) => bool) => bool) + (%f::(nat => bool) => 'a::type * (nat => bool). + (All::(('a::type => (nat => bool) => 'b::type * (nat => bool)) => bool) + => bool) + (%g::'a::type => (nat => bool) => 'b::type * (nat => bool). (op -->::bool => bool => bool) ((op &::bool => bool => bool) - ((indep::((nat => bool) => 'a * (nat => bool)) => bool) f) - ((All::('a => bool) => bool) - (%x::'a. - (indep::((nat => bool) => 'b * (nat => bool)) => bool) + ((indep::((nat => bool) => 'a::type * (nat => bool)) => bool) + f) + ((All::('a::type => bool) => bool) + (%x::'a::type. + (indep::((nat => bool) => 'b::type * (nat => bool)) + => bool) (g x)))) - ((indep::((nat => bool) => 'b * (nat => bool)) => bool) - ((BIND::((nat => bool) => 'a * (nat => bool)) - => ('a => (nat => bool) => 'b * (nat => bool)) - => (nat => bool) => 'b * (nat => bool)) + ((indep::((nat => bool) => 'b::type * (nat => bool)) => bool) + ((BIND::((nat => bool) => 'a::type * (nat => bool)) + => ('a::type + => (nat => bool) => 'b::type * (nat => bool)) + => (nat => bool) => 'b::type * (nat => bool)) f g))))" by (import prob_indep INDEP_BIND) -lemma INDEP_PROB: "(All::(((nat => bool) => 'a * (nat => bool)) => bool) => bool) - (%f::(nat => bool) => 'a * (nat => bool). - (All::(('a => bool) => bool) => bool) - (%p::'a => bool. +lemma INDEP_PROB: "(All::(((nat => bool) => 'a::type * (nat => bool)) => bool) => bool) + (%f::(nat => bool) => 'a::type * (nat => bool). + (All::(('a::type => bool) => bool) => bool) + (%p::'a::type => bool. (All::(((nat => bool) => bool) => bool) => bool) (%q::(nat => bool) => bool. (op -->::bool => bool => bool) ((op &::bool => bool => bool) - ((indep::((nat => bool) => 'a * (nat => bool)) => bool) f) + ((indep::((nat => bool) => 'a::type * (nat => bool)) + => bool) + f) ((measurable::((nat => bool) => bool) => bool) q)) ((op =::real => real => bool) ((prob::((nat => bool) => bool) => real) ((pred_set.INTER::((nat => bool) => bool) => ((nat => bool) => bool) => (nat => bool) => bool) - ((op o::('a => bool) - => ((nat => bool) => 'a) + ((op o::('a::type => bool) + => ((nat => bool) => 'a::type) => (nat => bool) => bool) - p ((op o::('a * (nat => bool) => 'a) - => ((nat => bool) => 'a * (nat => bool)) - => (nat => bool) => 'a) - (fst::'a * (nat => bool) => 'a) f)) + p ((op o::('a::type * (nat => bool) => 'a::type) + => ((nat => bool) +=> 'a::type * (nat => bool)) + => (nat => bool) => 'a::type) + (fst::'a::type * (nat => bool) => 'a::type) f)) ((op o::((nat => bool) => bool) => ((nat => bool) => nat => bool) => (nat => bool) => bool) - q ((op o::('a * (nat => bool) => nat => bool) - => ((nat => bool) => 'a * (nat => bool)) + q ((op o::('a::type * (nat => bool) => nat => bool) + => ((nat => bool) +=> 'a::type * (nat => bool)) => (nat => bool) => nat => bool) - (snd::'a * (nat => bool) => nat => bool) f)))) + (snd::'a::type * (nat => bool) => nat => bool) + f)))) ((op *::real => real => real) ((prob::((nat => bool) => bool) => real) - ((op o::('a => bool) - => ((nat => bool) => 'a) + ((op o::('a::type => bool) + => ((nat => bool) => 'a::type) => (nat => bool) => bool) - p ((op o::('a * (nat => bool) => 'a) - => ((nat => bool) => 'a * (nat => bool)) - => (nat => bool) => 'a) - (fst::'a * (nat => bool) => 'a) f))) + p ((op o::('a::type * (nat => bool) => 'a::type) + => ((nat => bool) +=> 'a::type * (nat => bool)) + => (nat => bool) => 'a::type) + (fst::'a::type * (nat => bool) => 'a::type) f))) ((prob::((nat => bool) => bool) => real) q))))))" by (import prob_indep INDEP_PROB) -lemma INDEP_MEASURABLE1: "(All::(((nat => bool) => 'a * (nat => bool)) => bool) => bool) - (%f::(nat => bool) => 'a * (nat => bool). - (All::(('a => bool) => bool) => bool) - (%p::'a => bool. +lemma INDEP_MEASURABLE1: "(All::(((nat => bool) => 'a::type * (nat => bool)) => bool) => bool) + (%f::(nat => bool) => 'a::type * (nat => bool). + (All::(('a::type => bool) => bool) => bool) + (%p::'a::type => bool. (op -->::bool => bool => bool) - ((indep::((nat => bool) => 'a * (nat => bool)) => bool) f) + ((indep::((nat => bool) => 'a::type * (nat => bool)) => bool) f) ((measurable::((nat => bool) => bool) => bool) - ((op o::('a => bool) - => ((nat => bool) => 'a) => (nat => bool) => bool) - p ((op o::('a * (nat => bool) => 'a) - => ((nat => bool) => 'a * (nat => bool)) - => (nat => bool) => 'a) - (fst::'a * (nat => bool) => 'a) f)))))" + ((op o::('a::type => bool) + => ((nat => bool) => 'a::type) + => (nat => bool) => bool) + p ((op o::('a::type * (nat => bool) => 'a::type) + => ((nat => bool) => 'a::type * (nat => bool)) + => (nat => bool) => 'a::type) + (fst::'a::type * (nat => bool) => 'a::type) f)))))" by (import prob_indep INDEP_MEASURABLE1) -lemma INDEP_MEASURABLE2: "(All::(((nat => bool) => 'a * (nat => bool)) => bool) => bool) - (%f::(nat => bool) => 'a * (nat => bool). +lemma INDEP_MEASURABLE2: "(All::(((nat => bool) => 'a::type * (nat => bool)) => bool) => bool) + (%f::(nat => bool) => 'a::type * (nat => bool). (All::(((nat => bool) => bool) => bool) => bool) (%q::(nat => bool) => bool. (op -->::bool => bool => bool) ((op &::bool => bool => bool) - ((indep::((nat => bool) => 'a * (nat => bool)) => bool) f) + ((indep::((nat => bool) => 'a::type * (nat => bool)) => bool) + f) ((measurable::((nat => bool) => bool) => bool) q)) ((measurable::((nat => bool) => bool) => bool) ((op o::((nat => bool) => bool) => ((nat => bool) => nat => bool) => (nat => bool) => bool) - q ((op o::('a * (nat => bool) => nat => bool) - => ((nat => bool) => 'a * (nat => bool)) + q ((op o::('a::type * (nat => bool) => nat => bool) + => ((nat => bool) => 'a::type * (nat => bool)) => (nat => bool) => nat => bool) - (snd::'a * (nat => bool) => nat => bool) f)))))" + (snd::'a::type * (nat => bool) => nat => bool) f)))))" by (import prob_indep INDEP_MEASURABLE2) lemma PROB_INDEP_BOUND: "(All::(((nat => bool) => nat * (nat => bool)) => bool) => bool) @@ -3363,15 +3519,22 @@ defs unif_bound_primdef: "unif_bound == -WFREC (SOME R. WF R & (ALL v. R (Suc v div 2) (Suc v))) - (%unif_bound. nat_case 0 (%v1. Suc (unif_bound (Suc v1 div 2))))" +WFREC + (SOME R::nat => nat => bool. + WF R & (ALL v::nat. R (Suc v div (2::nat)) (Suc v))) + (%unif_bound::nat => nat. + nat_case (0::nat) (%v1::nat. Suc (unif_bound (Suc v1 div (2::nat)))))" lemma unif_bound_primitive_def: "unif_bound = -WFREC (SOME R. WF R & (ALL v. R (Suc v div 2) (Suc v))) - (%unif_bound. nat_case 0 (%v1. Suc (unif_bound (Suc v1 div 2))))" +WFREC + (SOME R::nat => nat => bool. + WF R & (ALL v::nat. R (Suc v div (2::nat)) (Suc v))) + (%unif_bound::nat => nat. + nat_case (0::nat) (%v1::nat. Suc (unif_bound (Suc v1 div (2::nat)))))" by (import prob_uniform unif_bound_primitive_def) -lemma unif_bound_def: "unif_bound 0 = 0 & unif_bound (Suc v) = Suc (unif_bound (Suc v div 2))" +lemma unif_bound_def: "unif_bound (0::nat) = (0::nat) & +unif_bound (Suc (v::nat)) = Suc (unif_bound (Suc v div (2::nat)))" by (import prob_uniform unif_bound_def) lemma unif_bound_ind: "(All::((nat => bool) => bool) => bool) @@ -3394,35 +3557,47 @@ constdefs unif_tupled :: "nat * (nat => bool) => nat * (nat => bool)" "unif_tupled == -WFREC (SOME R. WF R & (ALL s v2. R (Suc v2 div 2, s) (Suc v2, s))) - (%unif_tupled (v, v1). - case v of 0 => (0, v1) - | Suc v3 => - let (m, s') = unif_tupled (Suc v3 div 2, v1) - in (if SHD s' then 2 * m + 1 else 2 * m, STL s'))" +WFREC + (SOME R::nat * (nat => bool) => nat * (nat => bool) => bool. + WF R & + (ALL (s::nat => bool) v2::nat. R (Suc v2 div (2::nat), s) (Suc v2, s))) + (%(unif_tupled::nat * (nat => bool) => nat * (nat => bool)) (v::nat, + v1::nat => bool). + case v of 0 => (0::nat, v1) + | Suc (v3::nat) => + let (m::nat, s'::nat => bool) = + unif_tupled (Suc v3 div (2::nat), v1) + in (if SHD s' then (2::nat) * m + (1::nat) else (2::nat) * m, + STL s'))" lemma unif_tupled_primitive_def: "unif_tupled = -WFREC (SOME R. WF R & (ALL s v2. R (Suc v2 div 2, s) (Suc v2, s))) - (%unif_tupled (v, v1). - case v of 0 => (0, v1) - | Suc v3 => - let (m, s') = unif_tupled (Suc v3 div 2, v1) - in (if SHD s' then 2 * m + 1 else 2 * m, STL s'))" +WFREC + (SOME R::nat * (nat => bool) => nat * (nat => bool) => bool. + WF R & + (ALL (s::nat => bool) v2::nat. R (Suc v2 div (2::nat), s) (Suc v2, s))) + (%(unif_tupled::nat * (nat => bool) => nat * (nat => bool)) (v::nat, + v1::nat => bool). + case v of 0 => (0::nat, v1) + | Suc (v3::nat) => + let (m::nat, s'::nat => bool) = + unif_tupled (Suc v3 div (2::nat), v1) + in (if SHD s' then (2::nat) * m + (1::nat) else (2::nat) * m, + STL s'))" by (import prob_uniform unif_tupled_primitive_def) consts unif :: "nat => (nat => bool) => nat * (nat => bool)" defs - unif_primdef: "unif == %x x1. unif_tupled (x, x1)" - -lemma unif_curried_def: "ALL x x1. unif x x1 = unif_tupled (x, x1)" + unif_primdef: "unif == %(x::nat) x1::nat => bool. unif_tupled (x, x1)" + +lemma unif_curried_def: "ALL (x::nat) x1::nat => bool. unif x x1 = unif_tupled (x, x1)" by (import prob_uniform unif_curried_def) -lemma unif_def: "unif 0 s = (0, s) & -unif (Suc v2) s = -(let (m, s') = unif (Suc v2 div 2) s - in (if SHD s' then 2 * m + 1 else 2 * m, STL s'))" +lemma unif_def: "unif (0::nat) (s::nat => bool) = (0::nat, s) & +unif (Suc (v2::nat)) s = +(let (m::nat, s'::nat => bool) = unif (Suc v2 div (2::nat)) s + in (if SHD s' then (2::nat) * m + (1::nat) else (2::nat) * m, STL s'))" by (import prob_uniform unif_def) lemma unif_ind: "(All::((nat => (nat => bool) => bool) => bool) => bool) @@ -3648,9 +3823,10 @@ uniform :: "nat => nat => (nat => bool) => nat * (nat => bool)" defs - uniform_primdef: "uniform == %x x1 x2. uniform_tupled (x, x1, x2)" - -lemma uniform_curried_def: "ALL x x1 x2. uniform x x1 x2 = uniform_tupled (x, x1, x2)" + uniform_primdef: "uniform == %(x::nat) (x1::nat) x2::nat => bool. uniform_tupled (x, x1, x2)" + +lemma uniform_curried_def: "ALL (x::nat) (x1::nat) x2::nat => bool. + uniform x x1 x2 = uniform_tupled (x, x1, x2)" by (import prob_uniform uniform_curried_def) lemma uniform_ind: "(All::((nat => nat => (nat => bool) => bool) => bool) => bool) @@ -3695,19 +3871,19 @@ (%xa::nat. (All::((nat => bool) => bool) => bool) (P x xa)))))" by (import prob_uniform uniform_ind) -lemma uniform_def: "uniform 0 (Suc n) s = (0, s) & -uniform (Suc t) (Suc n) s = -(let (xa, x) = unif n s +lemma uniform_def: "uniform (0::nat) (Suc (n::nat)) (s::nat => bool) = (0::nat, s) & +uniform (Suc (t::nat)) (Suc n) s = +(let (xa::nat, x::nat => bool) = unif n s in if xa < Suc n then (xa, x) else uniform t (Suc n) x)" by (import prob_uniform uniform_def) -lemma SUC_DIV_TWO_ZERO: "ALL n. (Suc n div 2 = 0) = (n = 0)" +lemma SUC_DIV_TWO_ZERO: "ALL n::nat. (Suc n div (2::nat) = (0::nat)) = (n = (0::nat))" by (import prob_uniform SUC_DIV_TWO_ZERO) -lemma UNIF_BOUND_LOWER: "ALL n. n < 2 ^ unif_bound n" +lemma UNIF_BOUND_LOWER: "ALL n::nat. n < (2::nat) ^ unif_bound n" by (import prob_uniform UNIF_BOUND_LOWER) -lemma UNIF_BOUND_LOWER_SUC: "ALL n. Suc n <= 2 ^ unif_bound n" +lemma UNIF_BOUND_LOWER_SUC: "ALL n::nat. Suc n <= (2::nat) ^ unif_bound n" by (import prob_uniform UNIF_BOUND_LOWER_SUC) lemma UNIF_BOUND_UPPER: "(All::(nat => bool) => bool) @@ -3729,39 +3905,71 @@ n)))" by (import prob_uniform UNIF_BOUND_UPPER) -lemma UNIF_BOUND_UPPER_SUC: "ALL n. 2 ^ unif_bound n <= Suc (2 * n)" +lemma UNIF_BOUND_UPPER_SUC: "ALL n::nat. (2::nat) ^ unif_bound n <= Suc ((2::nat) * n)" by (import prob_uniform UNIF_BOUND_UPPER_SUC) -lemma UNIF_DEF_MONAD: "unif 0 = UNIT 0 & -(ALL n. +lemma UNIF_DEF_MONAD: "unif (0::nat) = UNIT (0::nat) & +(ALL n::nat. unif (Suc n) = - BIND (unif (Suc n div 2)) - (%m. BIND SDEST (%b. UNIT (if b then 2 * m + 1 else 2 * m))))" + BIND (unif (Suc n div (2::nat))) + (%m::nat. + BIND SDEST + (%b::bool. + UNIT (if b then (2::nat) * m + (1::nat) else (2::nat) * m))))" by (import prob_uniform UNIF_DEF_MONAD) -lemma UNIFORM_DEF_MONAD: "(ALL x. uniform 0 (Suc x) = UNIT 0) & -(ALL x xa. +lemma UNIFORM_DEF_MONAD: "(ALL x::nat. uniform (0::nat) (Suc x) = UNIT (0::nat)) & +(ALL (x::nat) xa::nat. uniform (Suc x) (Suc xa) = - BIND (unif xa) (%m. if m < Suc xa then UNIT m else uniform x (Suc xa)))" + BIND (unif xa) + (%m::nat. if m < Suc xa then UNIT m else uniform x (Suc xa)))" by (import prob_uniform UNIFORM_DEF_MONAD) -lemma INDEP_UNIF: "ALL n. indep (unif n)" +lemma INDEP_UNIF: "ALL n::nat. indep (unif n)" by (import prob_uniform INDEP_UNIF) -lemma INDEP_UNIFORM: "ALL t n. indep (uniform t (Suc n))" +lemma INDEP_UNIFORM: "ALL (t::nat) n::nat. indep (uniform t (Suc n))" by (import prob_uniform INDEP_UNIFORM) -lemma PROB_UNIF: "ALL n k. - prob (%s. fst (unif n s) = k) = - (if k < 2 ^ unif_bound n then (1 / 2) ^ unif_bound n else 0)" +lemma PROB_UNIF: "(All::(nat => bool) => bool) + (%n::nat. + (All::(nat => bool) => bool) + (%k::nat. + (op =::real => real => bool) + ((prob::((nat => bool) => bool) => real) + (%s::nat => bool. + (op =::nat => nat => bool) + ((fst::nat * (nat => bool) => nat) + ((unif::nat => (nat => bool) => nat * (nat => bool)) n + s)) + k)) + ((If::bool => real => real => real) + ((op <::nat => nat => bool) k + ((op ^::nat => nat => nat) + ((number_of::bin => nat) + ((op BIT::bin => bit => bin) + ((op BIT::bin => bit => bin) (Numeral.Pls::bin) + (bit.B1::bit)) + (bit.B0::bit))) + ((unif_bound::nat => nat) n))) + ((op ^::real => nat => real) + ((op /::real => real => real) (1::real) + ((number_of::bin => real) + ((op BIT::bin => bit => bin) + ((op BIT::bin => bit => bin) (Numeral.Pls::bin) + (bit.B1::bit)) + (bit.B0::bit)))) + ((unif_bound::nat => nat) n)) + (0::real))))" by (import prob_uniform PROB_UNIF) -lemma UNIF_RANGE: "ALL n s. fst (unif n s) < 2 ^ unif_bound n" +lemma UNIF_RANGE: "ALL (n::nat) s::nat => bool. fst (unif n s) < (2::nat) ^ unif_bound n" by (import prob_uniform UNIF_RANGE) -lemma PROB_UNIF_PAIR: "ALL n k k'. - (prob (%s. fst (unif n s) = k) = prob (%s. fst (unif n s) = k')) = - ((k < 2 ^ unif_bound n) = (k' < 2 ^ unif_bound n))" +lemma PROB_UNIF_PAIR: "ALL (n::nat) (k::nat) k'::nat. + (prob (%s::nat => bool. fst (unif n s) = k) = + prob (%s::nat => bool. fst (unif n s) = k')) = + ((k < (2::nat) ^ unif_bound n) = (k' < (2::nat) ^ unif_bound n))" by (import prob_uniform PROB_UNIF_PAIR) lemma PROB_UNIF_BOUND: "(All::(nat => bool) => bool) @@ -3796,10 +4004,11 @@ ((unif_bound::nat => nat) n))))))" by (import prob_uniform PROB_UNIF_BOUND) -lemma PROB_UNIF_GOOD: "ALL n. 1 / 2 <= prob (%s. fst (unif n s) < Suc n)" +lemma PROB_UNIF_GOOD: "ALL n::nat. + (1::real) / (2::real) <= prob (%s::nat => bool. fst (unif n s) < Suc n)" by (import prob_uniform PROB_UNIF_GOOD) -lemma UNIFORM_RANGE: "ALL t n s. fst (uniform t (Suc n) s) < Suc n" +lemma UNIFORM_RANGE: "ALL (t::nat) (n::nat) s::nat => bool. fst (uniform t (Suc n) s) < Suc n" by (import prob_uniform UNIFORM_RANGE) lemma PROB_UNIFORM_LOWER_BOUND: "(All::(real => bool) => bool) diff -r 7e417a7cbf9f -r bd59bfd4bf37 src/HOL/Import/HOL/HOL4Real.thy --- a/src/HOL/Import/HOL/HOL4Real.thy Mon Sep 26 02:06:44 2005 +0200 +++ b/src/HOL/Import/HOL/HOL4Real.thy Mon Sep 26 02:27:14 2005 +0200 @@ -4,37 +4,39 @@ ;setup_theory realax -lemma HREAL_RDISTRIB: "ALL x y z. +lemma HREAL_RDISTRIB: "ALL (x::hreal) (y::hreal) z::hreal. hreal_mul (hreal_add x y) z = hreal_add (hreal_mul x z) (hreal_mul y z)" by (import realax HREAL_RDISTRIB) -lemma HREAL_EQ_ADDL: "ALL x y. x ~= hreal_add x y" +lemma HREAL_EQ_ADDL: "ALL (x::hreal) y::hreal. x ~= hreal_add x y" by (import realax HREAL_EQ_ADDL) -lemma HREAL_EQ_LADD: "ALL x y z. (hreal_add x y = hreal_add x z) = (y = z)" +lemma HREAL_EQ_LADD: "ALL (x::hreal) (y::hreal) z::hreal. + (hreal_add x y = hreal_add x z) = (y = z)" by (import realax HREAL_EQ_LADD) -lemma HREAL_LT_REFL: "ALL x. ~ hreal_lt x x" +lemma HREAL_LT_REFL: "ALL x::hreal. ~ hreal_lt x x" by (import realax HREAL_LT_REFL) -lemma HREAL_LT_ADDL: "ALL x y. hreal_lt x (hreal_add x y)" +lemma HREAL_LT_ADDL: "ALL (x::hreal) y::hreal. hreal_lt x (hreal_add x y)" by (import realax HREAL_LT_ADDL) -lemma HREAL_LT_NE: "ALL x y. hreal_lt x y --> x ~= y" +lemma HREAL_LT_NE: "ALL (x::hreal) y::hreal. hreal_lt x y --> x ~= y" by (import realax HREAL_LT_NE) -lemma HREAL_LT_ADDR: "ALL x y. ~ hreal_lt (hreal_add x y) x" +lemma HREAL_LT_ADDR: "ALL (x::hreal) y::hreal. ~ hreal_lt (hreal_add x y) x" by (import realax HREAL_LT_ADDR) -lemma HREAL_LT_GT: "ALL x y. hreal_lt x y --> ~ hreal_lt y x" +lemma HREAL_LT_GT: "ALL (x::hreal) y::hreal. hreal_lt x y --> ~ hreal_lt y x" by (import realax HREAL_LT_GT) -lemma HREAL_LT_ADD2: "ALL x1 x2 y1 y2. +lemma HREAL_LT_ADD2: "ALL (x1::hreal) (x2::hreal) (y1::hreal) y2::hreal. hreal_lt x1 y1 & hreal_lt x2 y2 --> hreal_lt (hreal_add x1 x2) (hreal_add y1 y2)" by (import realax HREAL_LT_ADD2) -lemma HREAL_LT_LADD: "ALL x y z. hreal_lt (hreal_add x y) (hreal_add x z) = hreal_lt y z" +lemma HREAL_LT_LADD: "ALL (x::hreal) (y::hreal) z::hreal. + hreal_lt (hreal_add x y) (hreal_add x z) = hreal_lt y z" by (import realax HREAL_LT_LADD) constdefs @@ -53,27 +55,29 @@ constdefs treal_neg :: "hreal * hreal => hreal * hreal" - "treal_neg == %(x, y). (y, x)" - -lemma treal_neg: "ALL x y. treal_neg (x, y) = (y, x)" + "treal_neg == %(x::hreal, y::hreal). (y, x)" + +lemma treal_neg: "ALL (x::hreal) y::hreal. treal_neg (x, y) = (y, x)" by (import realax treal_neg) constdefs treal_add :: "hreal * hreal => hreal * hreal => hreal * hreal" - "treal_add == %(x1, y1) (x2, y2). (hreal_add x1 x2, hreal_add y1 y2)" - -lemma treal_add: "ALL x1 y1 x2 y2. + "treal_add == +%(x1::hreal, y1::hreal) (x2::hreal, y2::hreal). + (hreal_add x1 x2, hreal_add y1 y2)" + +lemma treal_add: "ALL (x1::hreal) (y1::hreal) (x2::hreal) y2::hreal. treal_add (x1, y1) (x2, y2) = (hreal_add x1 x2, hreal_add y1 y2)" by (import realax treal_add) constdefs treal_mul :: "hreal * hreal => hreal * hreal => hreal * hreal" "treal_mul == -%(x1, y1) (x2, y2). +%(x1::hreal, y1::hreal) (x2::hreal, y2::hreal). (hreal_add (hreal_mul x1 x2) (hreal_mul y1 y2), hreal_add (hreal_mul x1 y2) (hreal_mul y1 x2))" -lemma treal_mul: "ALL x1 y1 x2 y2. +lemma treal_mul: "ALL (x1::hreal) (y1::hreal) (x2::hreal) y2::hreal. treal_mul (x1, y1) (x2, y2) = (hreal_add (hreal_mul x1 x2) (hreal_mul y1 y2), hreal_add (hreal_mul x1 y2) (hreal_mul y1 x2))" @@ -81,22 +85,24 @@ constdefs treal_lt :: "hreal * hreal => hreal * hreal => bool" - "treal_lt == %(x1, y1) (x2, y2). hreal_lt (hreal_add x1 y2) (hreal_add x2 y1)" - -lemma treal_lt: "ALL x1 y1 x2 y2. + "treal_lt == +%(x1::hreal, y1::hreal) (x2::hreal, y2::hreal). + hreal_lt (hreal_add x1 y2) (hreal_add x2 y1)" + +lemma treal_lt: "ALL (x1::hreal) (y1::hreal) (x2::hreal) y2::hreal. treal_lt (x1, y1) (x2, y2) = hreal_lt (hreal_add x1 y2) (hreal_add x2 y1)" by (import realax treal_lt) constdefs treal_inv :: "hreal * hreal => hreal * hreal" "treal_inv == -%(x, y). +%(x::hreal, y::hreal). if x = y then treal_0 else if hreal_lt y x then (hreal_add (hreal_inv (hreal_sub x y)) hreal_1, hreal_1) else (hreal_1, hreal_add (hreal_inv (hreal_sub y x)) hreal_1)" -lemma treal_inv: "ALL x y. +lemma treal_inv: "ALL (x::hreal) y::hreal. treal_inv (x, y) = (if x = y then treal_0 else if hreal_lt y x @@ -106,132 +112,155 @@ constdefs treal_eq :: "hreal * hreal => hreal * hreal => bool" - "treal_eq == %(x1, y1) (x2, y2). hreal_add x1 y2 = hreal_add x2 y1" - -lemma treal_eq: "ALL x1 y1 x2 y2. + "treal_eq == +%(x1::hreal, y1::hreal) (x2::hreal, y2::hreal). + hreal_add x1 y2 = hreal_add x2 y1" + +lemma treal_eq: "ALL (x1::hreal) (y1::hreal) (x2::hreal) y2::hreal. treal_eq (x1, y1) (x2, y2) = (hreal_add x1 y2 = hreal_add x2 y1)" by (import realax treal_eq) -lemma TREAL_EQ_REFL: "ALL x. treal_eq x x" +lemma TREAL_EQ_REFL: "ALL x::hreal * hreal. treal_eq x x" by (import realax TREAL_EQ_REFL) -lemma TREAL_EQ_SYM: "ALL x y. treal_eq x y = treal_eq y x" +lemma TREAL_EQ_SYM: "ALL (x::hreal * hreal) y::hreal * hreal. treal_eq x y = treal_eq y x" by (import realax TREAL_EQ_SYM) -lemma TREAL_EQ_TRANS: "ALL x y z. treal_eq x y & treal_eq y z --> treal_eq x z" +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" by (import realax TREAL_EQ_TRANS) -lemma TREAL_EQ_EQUIV: "ALL p q. treal_eq p q = (treal_eq p = treal_eq q)" +lemma TREAL_EQ_EQUIV: "ALL (p::hreal * hreal) q::hreal * hreal. + treal_eq p q = (treal_eq p = treal_eq q)" by (import realax TREAL_EQ_EQUIV) -lemma TREAL_EQ_AP: "ALL p q. p = q --> treal_eq p q" +lemma TREAL_EQ_AP: "ALL (p::hreal * hreal) q::hreal * hreal. p = q --> treal_eq p q" by (import realax TREAL_EQ_AP) lemma TREAL_10: "~ treal_eq treal_1 treal_0" by (import realax TREAL_10) -lemma TREAL_ADD_SYM: "ALL x y. treal_add x y = treal_add y x" +lemma TREAL_ADD_SYM: "ALL (x::hreal * hreal) y::hreal * hreal. treal_add x y = treal_add y x" by (import realax TREAL_ADD_SYM) -lemma TREAL_MUL_SYM: "ALL x y. treal_mul x y = treal_mul y x" +lemma TREAL_MUL_SYM: "ALL (x::hreal * hreal) y::hreal * hreal. treal_mul x y = treal_mul y x" by (import realax TREAL_MUL_SYM) -lemma TREAL_ADD_ASSOC: "ALL x y z. treal_add x (treal_add y z) = treal_add (treal_add x y) z" +lemma TREAL_ADD_ASSOC: "ALL (x::hreal * hreal) (y::hreal * hreal) z::hreal * hreal. + treal_add x (treal_add y z) = treal_add (treal_add x y) z" by (import realax TREAL_ADD_ASSOC) -lemma TREAL_MUL_ASSOC: "ALL x y z. treal_mul x (treal_mul y z) = treal_mul (treal_mul x y) z" +lemma TREAL_MUL_ASSOC: "ALL (x::hreal * hreal) (y::hreal * hreal) z::hreal * hreal. + treal_mul x (treal_mul y z) = treal_mul (treal_mul x y) z" by (import realax TREAL_MUL_ASSOC) -lemma TREAL_LDISTRIB: "ALL x y z. +lemma TREAL_LDISTRIB: "ALL (x::hreal * hreal) (y::hreal * hreal) z::hreal * hreal. treal_mul x (treal_add y z) = treal_add (treal_mul x y) (treal_mul x z)" by (import realax TREAL_LDISTRIB) -lemma TREAL_ADD_LID: "ALL x. treal_eq (treal_add treal_0 x) x" +lemma TREAL_ADD_LID: "ALL x::hreal * hreal. treal_eq (treal_add treal_0 x) x" by (import realax TREAL_ADD_LID) -lemma TREAL_MUL_LID: "ALL x. treal_eq (treal_mul treal_1 x) x" +lemma TREAL_MUL_LID: "ALL x::hreal * hreal. treal_eq (treal_mul treal_1 x) x" by (import realax TREAL_MUL_LID) -lemma TREAL_ADD_LINV: "ALL x. treal_eq (treal_add (treal_neg x) x) treal_0" +lemma TREAL_ADD_LINV: "ALL x::hreal * hreal. treal_eq (treal_add (treal_neg x) x) treal_0" by (import realax TREAL_ADD_LINV) lemma TREAL_INV_0: "treal_eq (treal_inv treal_0) treal_0" by (import realax TREAL_INV_0) -lemma TREAL_MUL_LINV: "ALL x. ~ treal_eq x treal_0 --> treal_eq (treal_mul (treal_inv x) x) treal_1" +lemma TREAL_MUL_LINV: "ALL x::hreal * hreal. + ~ treal_eq x treal_0 --> treal_eq (treal_mul (treal_inv x) x) treal_1" by (import realax TREAL_MUL_LINV) -lemma TREAL_LT_TOTAL: "ALL x y. treal_eq x y | treal_lt x y | treal_lt y x" +lemma TREAL_LT_TOTAL: "ALL (x::hreal * hreal) y::hreal * hreal. + treal_eq x y | treal_lt x y | treal_lt y x" by (import realax TREAL_LT_TOTAL) -lemma TREAL_LT_REFL: "ALL x. ~ treal_lt x x" +lemma TREAL_LT_REFL: "ALL x::hreal * hreal. ~ treal_lt x x" by (import realax TREAL_LT_REFL) -lemma TREAL_LT_TRANS: "ALL x y z. treal_lt x y & treal_lt y z --> treal_lt x z" +lemma TREAL_LT_TRANS: "ALL (x::hreal * hreal) (y::hreal * hreal) z::hreal * hreal. + treal_lt x y & treal_lt y z --> treal_lt x z" by (import realax TREAL_LT_TRANS) -lemma TREAL_LT_ADD: "ALL x y z. treal_lt y z --> treal_lt (treal_add x y) (treal_add x z)" +lemma TREAL_LT_ADD: "ALL (x::hreal * hreal) (y::hreal * hreal) z::hreal * hreal. + treal_lt y z --> treal_lt (treal_add x y) (treal_add x z)" by (import realax TREAL_LT_ADD) -lemma TREAL_LT_MUL: "ALL x y. +lemma TREAL_LT_MUL: "ALL (x::hreal * hreal) y::hreal * hreal. treal_lt treal_0 x & treal_lt treal_0 y --> treal_lt treal_0 (treal_mul x y)" by (import realax TREAL_LT_MUL) constdefs treal_of_hreal :: "hreal => hreal * hreal" - "treal_of_hreal == %x. (hreal_add x hreal_1, hreal_1)" - -lemma treal_of_hreal: "ALL x. treal_of_hreal x = (hreal_add x hreal_1, hreal_1)" + "treal_of_hreal == %x::hreal. (hreal_add x hreal_1, hreal_1)" + +lemma treal_of_hreal: "ALL x::hreal. treal_of_hreal x = (hreal_add x hreal_1, hreal_1)" by (import realax treal_of_hreal) constdefs hreal_of_treal :: "hreal * hreal => hreal" - "hreal_of_treal == %(x, y). SOME d. x = hreal_add y d" - -lemma hreal_of_treal: "ALL x y. hreal_of_treal (x, y) = (SOME d. x = hreal_add y d)" + "hreal_of_treal == %(x::hreal, y::hreal). SOME d::hreal. x = hreal_add y d" + +lemma hreal_of_treal: "ALL (x::hreal) y::hreal. + hreal_of_treal (x, y) = (SOME d::hreal. x = hreal_add y d)" by (import realax hreal_of_treal) -lemma TREAL_BIJ: "(ALL h. hreal_of_treal (treal_of_hreal h) = h) & -(ALL r. treal_lt treal_0 r = treal_eq (treal_of_hreal (hreal_of_treal r)) r)" +lemma TREAL_BIJ: "(ALL h::hreal. hreal_of_treal (treal_of_hreal h) = h) & +(ALL r::hreal * hreal. + treal_lt treal_0 r = treal_eq (treal_of_hreal (hreal_of_treal r)) r)" by (import realax TREAL_BIJ) -lemma TREAL_ISO: "ALL h i. hreal_lt h i --> treal_lt (treal_of_hreal h) (treal_of_hreal i)" +lemma TREAL_ISO: "ALL (h::hreal) i::hreal. + hreal_lt h i --> treal_lt (treal_of_hreal h) (treal_of_hreal i)" by (import realax TREAL_ISO) -lemma TREAL_BIJ_WELLDEF: "ALL h i. treal_eq h i --> hreal_of_treal h = hreal_of_treal i" +lemma TREAL_BIJ_WELLDEF: "ALL (h::hreal * hreal) i::hreal * hreal. + treal_eq h i --> hreal_of_treal h = hreal_of_treal i" by (import realax TREAL_BIJ_WELLDEF) -lemma TREAL_NEG_WELLDEF: "ALL x1 x2. treal_eq x1 x2 --> treal_eq (treal_neg x1) (treal_neg x2)" +lemma TREAL_NEG_WELLDEF: "ALL (x1::hreal * hreal) x2::hreal * hreal. + treal_eq x1 x2 --> treal_eq (treal_neg x1) (treal_neg x2)" by (import realax TREAL_NEG_WELLDEF) -lemma TREAL_ADD_WELLDEFR: "ALL x1 x2 y. treal_eq x1 x2 --> treal_eq (treal_add x1 y) (treal_add x2 y)" +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)" by (import realax TREAL_ADD_WELLDEFR) -lemma TREAL_ADD_WELLDEF: "ALL x1 x2 y1 y2. +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)" by (import realax TREAL_ADD_WELLDEF) -lemma TREAL_MUL_WELLDEFR: "ALL x1 x2 y. treal_eq x1 x2 --> treal_eq (treal_mul x1 y) (treal_mul x2 y)" +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)" by (import realax TREAL_MUL_WELLDEFR) -lemma TREAL_MUL_WELLDEF: "ALL x1 x2 y1 y2. +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)" by (import realax TREAL_MUL_WELLDEF) -lemma TREAL_LT_WELLDEFR: "ALL x1 x2 y. treal_eq x1 x2 --> treal_lt x1 y = treal_lt x2 y" +lemma TREAL_LT_WELLDEFR: "ALL (x1::hreal * hreal) (x2::hreal * hreal) y::hreal * hreal. + treal_eq x1 x2 --> treal_lt x1 y = treal_lt x2 y" by (import realax TREAL_LT_WELLDEFR) -lemma TREAL_LT_WELLDEFL: "ALL x y1 y2. treal_eq y1 y2 --> treal_lt x y1 = treal_lt x y2" +lemma TREAL_LT_WELLDEFL: "ALL (x::hreal * hreal) (y1::hreal * hreal) y2::hreal * hreal. + treal_eq y1 y2 --> treal_lt x y1 = treal_lt x y2" by (import realax TREAL_LT_WELLDEFL) -lemma TREAL_LT_WELLDEF: "ALL x1 x2 y1 y2. +lemma TREAL_LT_WELLDEF: "ALL (x1::hreal * hreal) (x2::hreal * hreal) (y1::hreal * hreal) + y2::hreal * hreal. treal_eq x1 x2 & treal_eq y1 y2 --> treal_lt x1 y1 = treal_lt x2 y2" by (import realax TREAL_LT_WELLDEF) -lemma TREAL_INV_WELLDEF: "ALL x1 x2. treal_eq x1 x2 --> treal_eq (treal_inv x1) (treal_inv x2)" +lemma TREAL_INV_WELLDEF: "ALL (x1::hreal * hreal) x2::hreal * hreal. + treal_eq x1 x2 --> treal_eq (treal_inv x1) (treal_inv x2)" by (import realax TREAL_INV_WELLDEF) ;end_setup @@ -413,7 +442,7 @@ x ~= (0::real) & x * xa = x * xb --> xa = xb" by (import real REAL_EQ_LMUL_IMP) -lemma REAL_FACT_NZ: "ALL n. real (FACT n) ~= 0" +lemma REAL_FACT_NZ: "ALL n::nat. real (FACT n) ~= (0::real)" by (import real REAL_FACT_NZ) lemma REAL_DIFFSQ: "ALL (x::real) y::real. (x + y) * (x - y) = x * x - y * y" @@ -477,7 +506,7 @@ x < y" by (import real ABS_BETWEEN2) -lemma POW_PLUS1: "ALL e>0. ALL n. 1 + real n * e <= (1 + e) ^ n" +lemma POW_PLUS1: "ALL e>0::real. ALL n::nat. (1::real) + real n * e <= ((1::real) + e) ^ n" by (import real POW_PLUS1) lemma POW_M1: "ALL n::nat. abs ((- (1::real)) ^ n) = (1::real)" @@ -533,17 +562,22 @@ constdefs sup :: "(real => bool) => real" - "sup == %P. SOME s. ALL y. (EX x. P x & y < x) = (y < s)" - -lemma sup: "ALL P. sup P = (SOME s. ALL y. (EX x. P x & y < x) = (y < s))" + "sup == +%P::real => bool. + SOME s::real. ALL y::real. (EX x::real. P x & y < x) = (y < s)" + +lemma sup: "ALL P::real => bool. + sup P = (SOME s::real. ALL y::real. (EX x::real. P x & y < x) = (y < s))" by (import real sup) -lemma REAL_SUP: "ALL P. - Ex P & (EX z. ALL x. P x --> x < z) --> - (ALL y. (EX x. P x & y < x) = (y < sup P))" +lemma REAL_SUP: "ALL P::real => bool. + Ex P & (EX z::real. ALL x::real. P x --> x < z) --> + (ALL y::real. (EX x::real. P x & y < x) = (y < sup P))" by (import real REAL_SUP) -lemma REAL_SUP_UBOUND: "ALL P. Ex P & (EX z. ALL x. P x --> x < z) --> (ALL y. P y --> y <= sup P)" +lemma REAL_SUP_UBOUND: "ALL P::real => bool. + Ex P & (EX z::real. ALL x::real. P x --> x < z) --> + (ALL y::real. P y --> y <= sup P)" by (import real REAL_SUP_UBOUND) lemma SETOK_LE_LT: "ALL P::real => bool. @@ -551,22 +585,26 @@ (Ex P & (EX z::real. ALL x::real. P x --> x < z))" by (import real SETOK_LE_LT) -lemma REAL_SUP_LE: "ALL P. - Ex P & (EX z. ALL x. P x --> x <= z) --> - (ALL y. (EX x. P x & y < x) = (y < sup P))" +lemma REAL_SUP_LE: "ALL P::real => bool. + Ex P & (EX z::real. ALL x::real. P x --> x <= z) --> + (ALL y::real. (EX x::real. P x & y < x) = (y < sup P))" by (import real REAL_SUP_LE) -lemma REAL_SUP_UBOUND_LE: "ALL P. Ex P & (EX z. ALL x. P x --> x <= z) --> (ALL y. P y --> y <= sup P)" +lemma REAL_SUP_UBOUND_LE: "ALL P::real => bool. + Ex P & (EX z::real. ALL x::real. P x --> x <= z) --> + (ALL y::real. P y --> y <= sup P)" by (import real REAL_SUP_UBOUND_LE) -lemma REAL_ARCH_LEAST: "ALL y>0. ALL x>=0. EX n. real n * y <= x & x < real (Suc n) * y" +lemma REAL_ARCH_LEAST: "ALL y>0::real. + ALL x>=0::real. EX n::nat. real n * y <= x & x < real (Suc n) * y" by (import real REAL_ARCH_LEAST) consts sumc :: "nat => nat => (nat => real) => real" -specification (sumc) sumc: "(ALL n f. sumc n 0 f = 0) & -(ALL n m f. sumc n (Suc m) f = sumc n m f + f (n + m))" +specification (sumc) sumc: "(ALL (n::nat) f::nat => real. sumc n (0::nat) f = (0::real)) & +(ALL (n::nat) (m::nat) f::nat => real. + sumc n (Suc m) f = sumc n m f + f (n + m))" by (import real sumc) consts @@ -580,97 +618,114 @@ => nat * nat => (nat => real) => real) (sumc::nat => nat => (nat => real) => real))" -lemma SUM_DEF: "ALL m n f. real.sum (m, n) f = sumc m n f" +lemma SUM_DEF: "ALL (m::nat) (n::nat) f::nat => real. real.sum (m, n) f = sumc m n f" by (import real SUM_DEF) -lemma sum: "ALL x xa xb. - real.sum (xa, 0) x = 0 & +lemma sum: "ALL (x::nat => real) (xa::nat) xb::nat. + real.sum (xa, 0::nat) x = (0::real) & real.sum (xa, Suc xb) x = real.sum (xa, xb) x + x (xa + xb)" by (import real sum) -lemma SUM_TWO: "ALL f n p. real.sum (0, n) f + real.sum (n, p) f = real.sum (0, n + p) f" +lemma SUM_TWO: "ALL (f::nat => real) (n::nat) p::nat. + real.sum (0::nat, n) f + real.sum (n, p) f = real.sum (0::nat, n + p) f" by (import real SUM_TWO) -lemma SUM_DIFF: "ALL f m n. real.sum (m, n) f = real.sum (0, m + n) f - real.sum (0, m) f" +lemma SUM_DIFF: "ALL (f::nat => real) (m::nat) n::nat. + real.sum (m, n) f = real.sum (0::nat, m + n) f - real.sum (0::nat, m) f" by (import real SUM_DIFF) -lemma ABS_SUM: "ALL f m n. abs (real.sum (m, n) f) <= real.sum (m, n) (%n. abs (f n))" +lemma ABS_SUM: "ALL (f::nat => real) (m::nat) n::nat. + abs (real.sum (m, n) f) <= real.sum (m, n) (%n::nat. abs (f n))" by (import real ABS_SUM) -lemma SUM_LE: "ALL f g m n. - (ALL r. m <= r & r < n + m --> f r <= g r) --> +lemma SUM_LE: "ALL (f::nat => real) (g::nat => real) (m::nat) n::nat. + (ALL r::nat. m <= r & r < n + m --> f r <= g r) --> real.sum (m, n) f <= real.sum (m, n) g" by (import real SUM_LE) -lemma SUM_EQ: "ALL f g m n. - (ALL r. m <= r & r < n + m --> f r = g r) --> +lemma SUM_EQ: "ALL (f::nat => real) (g::nat => real) (m::nat) n::nat. + (ALL r::nat. m <= r & r < n + m --> f r = g r) --> real.sum (m, n) f = real.sum (m, n) g" by (import real SUM_EQ) -lemma SUM_POS: "ALL f. (ALL n. 0 <= f n) --> (ALL m n. 0 <= real.sum (m, n) f)" +lemma SUM_POS: "ALL f::nat => real. + (ALL n::nat. (0::real) <= f n) --> + (ALL (m::nat) n::nat. (0::real) <= real.sum (m, n) f)" by (import real SUM_POS) -lemma SUM_POS_GEN: "ALL f m. (ALL n. m <= n --> 0 <= f n) --> (ALL n. 0 <= real.sum (m, n) f)" +lemma SUM_POS_GEN: "ALL (f::nat => real) m::nat. + (ALL n::nat. m <= n --> (0::real) <= f n) --> + (ALL n::nat. (0::real) <= real.sum (m, n) f)" by (import real SUM_POS_GEN) -lemma SUM_ABS: "ALL f m x. - abs (real.sum (m, x) (%m. abs (f m))) = real.sum (m, x) (%m. abs (f m))" +lemma SUM_ABS: "ALL (f::nat => real) (m::nat) x::nat. + abs (real.sum (m, x) (%m::nat. abs (f m))) = + real.sum (m, x) (%m::nat. abs (f m))" by (import real SUM_ABS) -lemma SUM_ABS_LE: "ALL f m n. abs (real.sum (m, n) f) <= real.sum (m, n) (%n. abs (f n))" +lemma SUM_ABS_LE: "ALL (f::nat => real) (m::nat) n::nat. + abs (real.sum (m, n) f) <= real.sum (m, n) (%n::nat. abs (f n))" by (import real SUM_ABS_LE) -lemma SUM_ZERO: "ALL f N. - (ALL n. N <= n --> f n = 0) --> - (ALL m n. N <= m --> real.sum (m, n) f = 0)" +lemma SUM_ZERO: "ALL (f::nat => real) N::nat. + (ALL n::nat. N <= n --> f n = (0::real)) --> + (ALL (m::nat) n::nat. N <= m --> real.sum (m, n) f = (0::real))" by (import real SUM_ZERO) -lemma SUM_ADD: "ALL f g m n. - real.sum (m, n) (%n. f n + g n) = real.sum (m, n) f + real.sum (m, n) g" +lemma SUM_ADD: "ALL (f::nat => real) (g::nat => real) (m::nat) n::nat. + real.sum (m, n) (%n::nat. f n + g n) = + real.sum (m, n) f + real.sum (m, n) g" by (import real SUM_ADD) -lemma SUM_CMUL: "ALL f c m n. real.sum (m, n) (%n. c * f n) = c * real.sum (m, n) f" +lemma SUM_CMUL: "ALL (f::nat => real) (c::real) (m::nat) n::nat. + real.sum (m, n) (%n::nat. c * f n) = c * real.sum (m, n) f" by (import real SUM_CMUL) -lemma SUM_NEG: "ALL f n d. real.sum (n, d) (%n. - f n) = - real.sum (n, d) f" +lemma SUM_NEG: "ALL (f::nat => real) (n::nat) d::nat. + real.sum (n, d) (%n::nat. - f n) = - real.sum (n, d) f" by (import real SUM_NEG) -lemma SUM_SUB: "ALL f g m n. - real.sum (m, n) (%x. f x - g x) = real.sum (m, n) f - real.sum (m, n) g" +lemma SUM_SUB: "ALL (f::nat => real) (g::nat => real) (m::nat) n::nat. + real.sum (m, n) (%x::nat. f x - g x) = + real.sum (m, n) f - real.sum (m, n) g" by (import real SUM_SUB) -lemma SUM_SUBST: "ALL f g m n. - (ALL p. m <= p & p < m + n --> f p = g p) --> +lemma SUM_SUBST: "ALL (f::nat => real) (g::nat => real) (m::nat) n::nat. + (ALL p::nat. m <= p & p < m + n --> f p = g p) --> real.sum (m, n) f = real.sum (m, n) g" by (import real SUM_SUBST) -lemma SUM_NSUB: "ALL n f c. real.sum (0, n) f - real n * c = real.sum (0, n) (%p. f p - c)" +lemma SUM_NSUB: "ALL (n::nat) (f::nat => real) c::real. + real.sum (0::nat, n) f - real n * c = + real.sum (0::nat, n) (%p::nat. f p - c)" by (import real SUM_NSUB) -lemma SUM_BOUND: "ALL f k m n. - (ALL p. m <= p & p < m + n --> f p <= k) --> +lemma SUM_BOUND: "ALL (f::nat => real) (k::real) (m::nat) n::nat. + (ALL p::nat. m <= p & p < m + n --> f p <= k) --> real.sum (m, n) f <= real n * k" by (import real SUM_BOUND) -lemma SUM_GROUP: "ALL n k f. - real.sum (0, n) (%m. real.sum (m * k, k) f) = real.sum (0, n * k) f" +lemma SUM_GROUP: "ALL (n::nat) (k::nat) f::nat => real. + real.sum (0::nat, n) (%m::nat. real.sum (m * k, k) f) = + real.sum (0::nat, n * k) f" by (import real SUM_GROUP) -lemma SUM_1: "ALL f n. real.sum (n, 1) f = f n" +lemma SUM_1: "ALL (f::nat => real) n::nat. real.sum (n, 1::nat) f = f n" by (import real SUM_1) -lemma SUM_2: "ALL f n. real.sum (n, 2) f = f n + f (n + 1)" +lemma SUM_2: "ALL (f::nat => real) n::nat. real.sum (n, 2::nat) f = f n + f (n + (1::nat))" by (import real SUM_2) -lemma SUM_OFFSET: "ALL f n k. - real.sum (0, n) (%m. f (m + k)) = - real.sum (0, n + k) f - real.sum (0, k) f" +lemma SUM_OFFSET: "ALL (f::nat => real) (n::nat) k::nat. + real.sum (0::nat, n) (%m::nat. f (m + k)) = + real.sum (0::nat, n + k) f - real.sum (0::nat, k) f" by (import real SUM_OFFSET) -lemma SUM_REINDEX: "ALL f m k n. real.sum (m + k, n) f = real.sum (m, n) (%r. f (r + k))" +lemma SUM_REINDEX: "ALL (f::nat => real) (m::nat) (k::nat) n::nat. + real.sum (m + k, n) f = real.sum (m, n) (%r::nat. f (r + k))" by (import real SUM_REINDEX) -lemma SUM_0: "ALL m n. real.sum (m, n) (%r. 0) = 0" +lemma SUM_0: "ALL (m::nat) n::nat. real.sum (m, n) (%r::nat. 0::real) = (0::real)" by (import real SUM_0) lemma SUM_PERMUTE_0: "(All::(nat => bool) => bool) @@ -697,7 +752,8 @@ ((Pair::nat => nat => nat * nat) (0::nat) n) f)))))" by (import real SUM_PERMUTE_0) -lemma SUM_CANCEL: "ALL f n d. real.sum (n, d) (%n. f (Suc n) - f n) = f (n + d) - f n" +lemma SUM_CANCEL: "ALL (f::nat => real) (n::nat) d::nat. + real.sum (n, d) (%n::nat. f (Suc n) - f n) = f (n + d) - f n" by (import real SUM_CANCEL) lemma REAL_EQ_RDIV_EQ: "ALL (x::real) (xa::real) xb::real. @@ -713,267 +769,328 @@ ;setup_theory topology constdefs - re_Union :: "(('a => bool) => bool) => 'a => bool" - "re_Union == %P x. EX s. P s & s x" - -lemma re_Union: "ALL P. re_Union P = (%x. EX s. P s & s x)" + re_Union :: "(('a::type => bool) => bool) => 'a::type => bool" + "re_Union == +%(P::('a::type => bool) => bool) x::'a::type. + EX s::'a::type => bool. P s & s x" + +lemma re_Union: "ALL P::('a::type => bool) => bool. + re_Union P = (%x::'a::type. EX s::'a::type => bool. P s & s x)" by (import topology re_Union) constdefs - re_union :: "('a => bool) => ('a => bool) => 'a => bool" - "re_union == %P Q x. P x | Q x" - -lemma re_union: "ALL P Q. re_union P Q = (%x. P x | Q x)" + re_union :: "('a::type => bool) => ('a::type => bool) => 'a::type => bool" + "re_union == +%(P::'a::type => bool) (Q::'a::type => bool) x::'a::type. P x | Q x" + +lemma re_union: "ALL (P::'a::type => bool) Q::'a::type => bool. + re_union P Q = (%x::'a::type. P x | Q x)" by (import topology re_union) constdefs - re_intersect :: "('a => bool) => ('a => bool) => 'a => bool" - "re_intersect == %P Q x. P x & Q x" - -lemma re_intersect: "ALL P Q. re_intersect P Q = (%x. P x & Q x)" + re_intersect :: "('a::type => bool) => ('a::type => bool) => 'a::type => bool" + "re_intersect == +%(P::'a::type => bool) (Q::'a::type => bool) x::'a::type. P x & Q x" + +lemma re_intersect: "ALL (P::'a::type => bool) Q::'a::type => bool. + re_intersect P Q = (%x::'a::type. P x & Q x)" by (import topology re_intersect) constdefs - re_null :: "'a => bool" - "re_null == %x. False" - -lemma re_null: "re_null = (%x. False)" + re_null :: "'a::type => bool" + "re_null == %x::'a::type. False" + +lemma re_null: "re_null = (%x::'a::type. False)" by (import topology re_null) constdefs - re_universe :: "'a => bool" - "re_universe == %x. True" - -lemma re_universe: "re_universe = (%x. True)" + re_universe :: "'a::type => bool" + "re_universe == %x::'a::type. True" + +lemma re_universe: "re_universe = (%x::'a::type. True)" by (import topology re_universe) constdefs - re_subset :: "('a => bool) => ('a => bool) => bool" - "re_subset == %P Q. ALL x. P x --> Q x" - -lemma re_subset: "ALL P Q. re_subset P Q = (ALL x. P x --> Q x)" + re_subset :: "('a::type => bool) => ('a::type => bool) => bool" + "re_subset == +%(P::'a::type => bool) Q::'a::type => bool. ALL x::'a::type. P x --> Q x" + +lemma re_subset: "ALL (P::'a::type => bool) Q::'a::type => bool. + re_subset P Q = (ALL x::'a::type. P x --> Q x)" by (import topology re_subset) constdefs - re_compl :: "('a => bool) => 'a => bool" - "re_compl == %P x. ~ P x" - -lemma re_compl: "ALL P. re_compl P = (%x. ~ P x)" + re_compl :: "('a::type => bool) => 'a::type => bool" + "re_compl == %(P::'a::type => bool) x::'a::type. ~ P x" + +lemma re_compl: "ALL P::'a::type => bool. re_compl P = (%x::'a::type. ~ P x)" by (import topology re_compl) -lemma SUBSET_REFL: "ALL P. re_subset P P" +lemma SUBSET_REFL: "ALL P::'a::type => bool. re_subset P P" by (import topology SUBSET_REFL) -lemma COMPL_MEM: "ALL P x. P x = (~ re_compl P x)" +lemma COMPL_MEM: "ALL (P::'a::type => bool) x::'a::type. P x = (~ re_compl P x)" by (import topology COMPL_MEM) -lemma SUBSET_ANTISYM: "ALL P Q. (re_subset P Q & re_subset Q P) = (P = Q)" +lemma SUBSET_ANTISYM: "ALL (P::'a::type => bool) Q::'a::type => bool. + (re_subset P Q & re_subset Q P) = (P = Q)" by (import topology SUBSET_ANTISYM) -lemma SUBSET_TRANS: "ALL P Q R. re_subset P Q & re_subset Q R --> re_subset P R" +lemma SUBSET_TRANS: "ALL (P::'a::type => bool) (Q::'a::type => bool) R::'a::type => bool. + re_subset P Q & re_subset Q R --> re_subset P R" by (import topology SUBSET_TRANS) constdefs - istopology :: "(('a => bool) => bool) => bool" + istopology :: "(('a::type => bool) => bool) => bool" "istopology == -%L. L re_null & - L re_universe & - (ALL a b. L a & L b --> L (re_intersect a b)) & - (ALL P. re_subset P L --> L (re_Union P))" - -lemma istopology: "ALL L. +%L::('a::type => bool) => bool. + L re_null & + L re_universe & + (ALL (a::'a::type => bool) b::'a::type => bool. + L a & L b --> L (re_intersect a b)) & + (ALL P::('a::type => bool) => bool. re_subset P L --> L (re_Union P))" + +lemma istopology: "ALL L::('a::type => bool) => bool. istopology L = (L re_null & L re_universe & - (ALL a b. L a & L b --> L (re_intersect a b)) & - (ALL P. re_subset P L --> L (re_Union P)))" + (ALL (a::'a::type => bool) b::'a::type => bool. + L a & L b --> L (re_intersect a b)) & + (ALL P::('a::type => bool) => bool. re_subset P L --> L (re_Union P)))" by (import topology istopology) -typedef (open) ('a) topology = "(Collect::((('a => bool) => bool) => bool) => (('a => bool) => bool) set) - (istopology::(('a => bool) => bool) => bool)" +typedef (open) ('a) topology = "(Collect::((('a::type => bool) => bool) => bool) + => (('a::type => bool) => bool) set) + (istopology::(('a::type => bool) => bool) => bool)" by (rule typedef_helper,import topology topology_TY_DEF) lemmas topology_TY_DEF = typedef_hol2hol4 [OF type_definition_topology] consts - topology :: "(('a => bool) => bool) => 'a topology" - "open" :: "'a topology => ('a => bool) => bool" - -specification ("open" topology) topology_tybij: "(ALL a::'a topology. topology (open a) = a) & -(ALL r::('a => bool) => bool. istopology r = (open (topology r) = r))" + topology :: "(('a::type => bool) => bool) => 'a::type topology" + "open" :: "'a::type topology => ('a::type => bool) => bool" + +specification ("open" topology) topology_tybij: "(ALL a::'a::type topology. topology (open a) = a) & +(ALL r::('a::type => bool) => bool. istopology r = (open (topology r) = r))" by (import topology topology_tybij) -lemma TOPOLOGY: "ALL L. +lemma TOPOLOGY: "ALL L::'a::type topology. open L re_null & open L re_universe & - (ALL a b. open L a & open L b --> open L (re_intersect a b)) & - (ALL P. re_subset P (open L) --> open L (re_Union P))" + (ALL (a::'a::type => bool) b::'a::type => bool. + open L a & open L b --> open L (re_intersect a b)) & + (ALL P::('a::type => bool) => bool. + re_subset P (open L) --> open L (re_Union P))" by (import topology TOPOLOGY) -lemma TOPOLOGY_UNION: "ALL x xa. re_subset xa (open x) --> open x (re_Union xa)" +lemma TOPOLOGY_UNION: "ALL (x::'a::type topology) xa::('a::type => bool) => bool. + re_subset xa (open x) --> open x (re_Union xa)" by (import topology TOPOLOGY_UNION) constdefs - neigh :: "'a topology => ('a => bool) * 'a => bool" - "neigh == %top (N, x). EX P. open top P & re_subset P N & P x" - -lemma neigh: "ALL top N x. neigh top (N, x) = (EX P. open top P & re_subset P N & P x)" + neigh :: "'a::type topology => ('a::type => bool) * 'a::type => bool" + "neigh == +%(top::'a::type topology) (N::'a::type => bool, x::'a::type). + EX P::'a::type => bool. open top P & re_subset P N & P x" + +lemma neigh: "ALL (top::'a::type topology) (N::'a::type => bool) x::'a::type. + neigh top (N, x) = + (EX P::'a::type => bool. open top P & re_subset P N & P x)" by (import topology neigh) -lemma OPEN_OWN_NEIGH: "ALL S' top x. open top S' & S' x --> neigh top (S', x)" +lemma OPEN_OWN_NEIGH: "ALL (S'::'a::type => bool) (top::'a::type topology) x::'a::type. + open top S' & S' x --> neigh top (S', x)" by (import topology OPEN_OWN_NEIGH) -lemma OPEN_UNOPEN: "ALL S' top. open top S' = (re_Union (%P. open top P & re_subset P S') = S')" +lemma OPEN_UNOPEN: "ALL (S'::'a::type => bool) top::'a::type topology. + open top S' = + (re_Union (%P::'a::type => bool. open top P & re_subset P S') = S')" by (import topology OPEN_UNOPEN) -lemma OPEN_SUBOPEN: "ALL S' top. - open top S' = (ALL x. S' x --> (EX P. P x & open top P & re_subset P S'))" +lemma OPEN_SUBOPEN: "ALL (S'::'a::type => bool) top::'a::type topology. + open top S' = + (ALL x::'a::type. + S' x --> (EX P::'a::type => bool. P x & open top P & re_subset P S'))" by (import topology OPEN_SUBOPEN) -lemma OPEN_NEIGH: "ALL S' top. - open top S' = (ALL x. S' x --> (EX N. neigh top (N, x) & re_subset N S'))" +lemma OPEN_NEIGH: "ALL (S'::'a::type => bool) top::'a::type topology. + open top S' = + (ALL x::'a::type. + S' x --> (EX N::'a::type => bool. neigh top (N, x) & re_subset N S'))" by (import topology OPEN_NEIGH) constdefs - closed :: "'a topology => ('a => bool) => bool" - "closed == %L S'. open L (re_compl S')" - -lemma closed: "ALL L S'. closed L S' = open L (re_compl S')" + closed :: "'a::type topology => ('a::type => bool) => bool" + "closed == %(L::'a::type topology) S'::'a::type => bool. open L (re_compl S')" + +lemma closed: "ALL (L::'a::type topology) S'::'a::type => bool. + closed L S' = open L (re_compl S')" by (import topology closed) constdefs - limpt :: "'a topology => 'a => ('a => bool) => bool" - "limpt == %top x S'. ALL N. neigh top (N, x) --> (EX y. x ~= y & S' y & N y)" - -lemma limpt: "ALL top x S'. + limpt :: "'a::type topology => 'a::type => ('a::type => bool) => bool" + "limpt == +%(top::'a::type topology) (x::'a::type) S'::'a::type => bool. + ALL N::'a::type => bool. + neigh top (N, x) --> (EX y::'a::type. x ~= y & S' y & N y)" + +lemma limpt: "ALL (top::'a::type topology) (x::'a::type) S'::'a::type => bool. limpt top x S' = - (ALL N. neigh top (N, x) --> (EX y. x ~= y & S' y & N y))" + (ALL N::'a::type => bool. + neigh top (N, x) --> (EX y::'a::type. x ~= y & S' y & N y))" by (import topology limpt) -lemma CLOSED_LIMPT: "ALL top S'. closed top S' = (ALL x. limpt top x S' --> S' x)" +lemma CLOSED_LIMPT: "ALL (top::'a::type topology) S'::'a::type => bool. + closed top S' = (ALL x::'a::type. limpt top x S' --> S' x)" by (import topology CLOSED_LIMPT) constdefs - ismet :: "('a * 'a => real) => bool" + ismet :: "('a::type * 'a::type => real) => bool" "ismet == -%m. (ALL x y. (m (x, y) = 0) = (x = y)) & - (ALL x y z. m (y, z) <= m (x, y) + m (x, z))" - -lemma ismet: "ALL m. +%m::'a::type * 'a::type => real. + (ALL (x::'a::type) y::'a::type. (m (x, y) = (0::real)) = (x = y)) & + (ALL (x::'a::type) (y::'a::type) z::'a::type. + m (y, z) <= m (x, y) + m (x, z))" + +lemma ismet: "ALL m::'a::type * 'a::type => real. ismet m = - ((ALL x y. (m (x, y) = 0) = (x = y)) & - (ALL x y z. m (y, z) <= m (x, y) + m (x, z)))" + ((ALL (x::'a::type) y::'a::type. (m (x, y) = (0::real)) = (x = y)) & + (ALL (x::'a::type) (y::'a::type) z::'a::type. + m (y, z) <= m (x, y) + m (x, z)))" by (import topology ismet) -typedef (open) ('a) metric = "(Collect::(('a * 'a => real) => bool) => ('a * 'a => real) set) - (ismet::('a * 'a => real) => bool)" +typedef (open) ('a) metric = "(Collect::(('a::type * 'a::type => real) => bool) + => ('a::type * 'a::type => real) set) + (ismet::('a::type * 'a::type => real) => bool)" by (rule typedef_helper,import topology metric_TY_DEF) lemmas metric_TY_DEF = typedef_hol2hol4 [OF type_definition_metric] consts - metric :: "('a * 'a => real) => 'a metric" - dist :: "'a metric => 'a * 'a => real" - -specification (dist metric) metric_tybij: "(ALL a::'a metric. metric (dist a) = a) & -(ALL r::'a * 'a => real. ismet r = (dist (metric r) = r))" + metric :: "('a::type * 'a::type => real) => 'a::type metric" + dist :: "'a::type metric => 'a::type * 'a::type => real" + +specification (dist metric) metric_tybij: "(ALL a::'a::type metric. metric (dist a) = a) & +(ALL r::'a::type * 'a::type => real. ismet r = (dist (metric r) = r))" by (import topology metric_tybij) -lemma METRIC_ISMET: "ALL m. ismet (dist m)" +lemma METRIC_ISMET: "ALL m::'a::type metric. ismet (dist m)" by (import topology METRIC_ISMET) -lemma METRIC_ZERO: "ALL m x y. (dist m (x, y) = 0) = (x = y)" +lemma METRIC_ZERO: "ALL (m::'a::type metric) (x::'a::type) y::'a::type. + (dist m (x, y) = (0::real)) = (x = y)" by (import topology METRIC_ZERO) -lemma METRIC_SAME: "ALL m x. dist m (x, x) = 0" +lemma METRIC_SAME: "ALL (m::'a::type metric) x::'a::type. dist m (x, x) = (0::real)" by (import topology METRIC_SAME) -lemma METRIC_POS: "ALL m x y. 0 <= dist m (x, y)" +lemma METRIC_POS: "ALL (m::'a::type metric) (x::'a::type) y::'a::type. + (0::real) <= dist m (x, y)" by (import topology METRIC_POS) -lemma METRIC_SYM: "ALL m x y. dist m (x, y) = dist m (y, x)" +lemma METRIC_SYM: "ALL (m::'a::type metric) (x::'a::type) y::'a::type. + dist m (x, y) = dist m (y, x)" by (import topology METRIC_SYM) -lemma METRIC_TRIANGLE: "ALL m x y z. dist m (x, z) <= dist m (x, y) + dist m (y, z)" +lemma METRIC_TRIANGLE: "ALL (m::'a::type metric) (x::'a::type) (y::'a::type) z::'a::type. + dist m (x, z) <= dist m (x, y) + dist m (y, z)" by (import topology METRIC_TRIANGLE) -lemma METRIC_NZ: "ALL m x y. x ~= y --> 0 < dist m (x, y)" +lemma METRIC_NZ: "ALL (m::'a::type metric) (x::'a::type) y::'a::type. + x ~= y --> (0::real) < dist m (x, y)" by (import topology METRIC_NZ) constdefs - mtop :: "'a metric => 'a topology" + mtop :: "'a::type metric => 'a::type topology" "mtop == -%m. topology - (%S'. ALL x. S' x --> (EX e>0. ALL y. dist m (x, y) < e --> S' y))" - -lemma mtop: "ALL m. +%m::'a::type metric. + topology + (%S'::'a::type => bool. + ALL x::'a::type. + S' x --> + (EX e>0::real. ALL y::'a::type. dist m (x, y) < e --> S' y))" + +lemma mtop: "ALL m::'a::type metric. mtop m = topology - (%S'. ALL x. S' x --> (EX e>0. ALL y. dist m (x, y) < e --> S' y))" + (%S'::'a::type => bool. + ALL x::'a::type. + S' x --> + (EX e>0::real. ALL y::'a::type. dist m (x, y) < e --> S' y))" by (import topology mtop) -lemma mtop_istopology: "ALL m. +lemma mtop_istopology: "ALL m::'a::type metric. istopology - (%S'. ALL x. S' x --> (EX e>0. ALL y. dist m (x, y) < e --> S' y))" + (%S'::'a::type => bool. + ALL x::'a::type. + S' x --> + (EX e>0::real. ALL y::'a::type. dist m (x, y) < e --> S' y))" by (import topology mtop_istopology) -lemma MTOP_OPEN: "ALL S' x. +lemma MTOP_OPEN: "ALL (S'::'a::type => bool) x::'a::type metric. open (mtop x) S' = - (ALL xa. S' xa --> (EX e>0. ALL y. dist x (xa, y) < e --> S' y))" + (ALL xa::'a::type. + S' xa --> + (EX e>0::real. ALL y::'a::type. dist x (xa, y) < e --> S' y))" by (import topology MTOP_OPEN) constdefs - B :: "'a metric => 'a * real => 'a => bool" - "B == %m (x, e) y. dist m (x, y) < e" - -lemma ball: "ALL m x e. B m (x, e) = (%y. dist m (x, y) < e)" + B :: "'a::type metric => 'a::type * real => 'a::type => bool" + "B == +%(m::'a::type metric) (x::'a::type, e::real) y::'a::type. dist m (x, y) < e" + +lemma ball: "ALL (m::'a::type metric) (x::'a::type) e::real. + B m (x, e) = (%y::'a::type. dist m (x, y) < e)" by (import topology ball) -lemma BALL_OPEN: "ALL m x e. 0 < e --> open (mtop m) (B m (x, e))" +lemma BALL_OPEN: "ALL (m::'a::type metric) (x::'a::type) e::real. + (0::real) < e --> open (mtop m) (B m (x, e))" by (import topology BALL_OPEN) -lemma BALL_NEIGH: "ALL m x e. 0 < e --> neigh (mtop m) (B m (x, e), x)" +lemma BALL_NEIGH: "ALL (m::'a::type metric) (x::'a::type) e::real. + (0::real) < e --> neigh (mtop m) (B m (x, e), x)" by (import topology BALL_NEIGH) -lemma MTOP_LIMPT: "ALL m x S'. - limpt (mtop m) x S' = (ALL e>0. EX y. x ~= y & S' y & dist m (x, y) < e)" +lemma MTOP_LIMPT: "ALL (m::'a::type metric) (x::'a::type) S'::'a::type => bool. + limpt (mtop m) x S' = + (ALL e>0::real. EX y::'a::type. x ~= y & S' y & dist m (x, y) < e)" by (import topology MTOP_LIMPT) -lemma ISMET_R1: "ismet (%(x, y). abs (y - x))" +lemma ISMET_R1: "ismet (%(x::real, y::real). abs (y - x))" by (import topology ISMET_R1) constdefs mr1 :: "real metric" - "mr1 == metric (%(x, y). abs (y - x))" - -lemma mr1: "mr1 = metric (%(x, y). abs (y - x))" + "mr1 == metric (%(x::real, y::real). abs (y - x))" + +lemma mr1: "mr1 = metric (%(x::real, y::real). abs (y - x))" by (import topology mr1) -lemma MR1_DEF: "ALL x y. dist mr1 (x, y) = abs (y - x)" +lemma MR1_DEF: "ALL (x::real) y::real. dist mr1 (x, y) = abs (y - x)" by (import topology MR1_DEF) -lemma MR1_ADD: "ALL x d. dist mr1 (x, x + d) = abs d" +lemma MR1_ADD: "ALL (x::real) d::real. dist mr1 (x, x + d) = abs d" by (import topology MR1_ADD) -lemma MR1_SUB: "ALL x d. dist mr1 (x, x - d) = abs d" +lemma MR1_SUB: "ALL (x::real) d::real. dist mr1 (x, x - d) = abs d" by (import topology MR1_SUB) -lemma MR1_ADD_POS: "ALL x d. 0 <= d --> dist mr1 (x, x + d) = d" +lemma MR1_ADD_POS: "ALL (x::real) d::real. (0::real) <= d --> dist mr1 (x, x + d) = d" by (import topology MR1_ADD_POS) -lemma MR1_SUB_LE: "ALL x d. 0 <= d --> dist mr1 (x, x - d) = d" +lemma MR1_SUB_LE: "ALL (x::real) d::real. (0::real) <= d --> dist mr1 (x, x - d) = d" by (import topology MR1_SUB_LE) -lemma MR1_ADD_LT: "ALL x d. 0 < d --> dist mr1 (x, x + d) = d" +lemma MR1_ADD_LT: "ALL (x::real) d::real. (0::real) < d --> dist mr1 (x, x + d) = d" by (import topology MR1_ADD_LT) -lemma MR1_SUB_LT: "ALL x d. 0 < d --> dist mr1 (x, x - d) = d" +lemma MR1_SUB_LT: "ALL (x::real) d::real. (0::real) < d --> dist mr1 (x, x - d) = d" by (import topology MR1_SUB_LT) -lemma MR1_BETWEEN1: "ALL x y z. x < z & dist mr1 (x, y) < z - x --> y < z" +lemma MR1_BETWEEN1: "ALL (x::real) (y::real) z::real. x < z & dist mr1 (x, y) < z - x --> y < z" by (import topology MR1_BETWEEN1) -lemma MR1_LIMPT: "ALL x. limpt (mtop mr1) x re_universe" +lemma MR1_LIMPT: "ALL x::real. limpt (mtop mr1) x re_universe" by (import topology MR1_LIMPT) ;end_setup @@ -981,189 +1098,220 @@ ;setup_theory nets constdefs - dorder :: "('a => 'a => bool) => bool" + dorder :: "('a::type => 'a::type => bool) => bool" "dorder == -%g. ALL x y. - g x x & g y y --> (EX z. g z z & (ALL w. g w z --> g w x & g w y))" - -lemma dorder: "ALL g. +%g::'a::type => 'a::type => bool. + ALL (x::'a::type) y::'a::type. + g x x & g y y --> + (EX z::'a::type. g z z & (ALL w::'a::type. g w z --> g w x & g w y))" + +lemma dorder: "ALL g::'a::type => 'a::type => bool. dorder g = - (ALL x y. - g x x & g y y --> (EX z. g z z & (ALL w. g w z --> g w x & g w y)))" + (ALL (x::'a::type) y::'a::type. + g x x & g y y --> + (EX z::'a::type. g z z & (ALL w::'a::type. g w z --> g w x & g w y)))" by (import nets dorder) constdefs - tends :: "('b => 'a) => 'a => 'a topology * ('b => 'b => bool) => bool" + tends :: "('b::type => 'a::type) +=> 'a::type => 'a::type topology * ('b::type => 'b::type => bool) => bool" "tends == -%(s::'b => 'a) (l::'a) (top::'a topology, g::'b => 'b => bool). - ALL N::'a => bool. +%(s::'b::type => 'a::type) (l::'a::type) (top::'a::type topology, + g::'b::type => 'b::type => bool). + ALL N::'a::type => bool. neigh top (N, l) --> - (EX n::'b. g n n & (ALL m::'b. g m n --> N (s m)))" - -lemma tends: "ALL (s::'b => 'a) (l::'a) (top::'a topology) g::'b => 'b => bool. + (EX n::'b::type. g n n & (ALL m::'b::type. g m n --> N (s m)))" + +lemma tends: "ALL (s::'b::type => 'a::type) (l::'a::type) (top::'a::type topology) + g::'b::type => 'b::type => bool. tends s l (top, g) = - (ALL N::'a => bool. + (ALL N::'a::type => bool. neigh top (N, l) --> - (EX n::'b. g n n & (ALL m::'b. g m n --> N (s m))))" + (EX n::'b::type. g n n & (ALL m::'b::type. g m n --> N (s m))))" by (import nets tends) constdefs - bounded :: "'a metric * ('b => 'b => bool) => ('b => 'a) => bool" + bounded :: "'a::type metric * ('b::type => 'b::type => bool) +=> ('b::type => 'a::type) => bool" "bounded == -%(m, g) f. EX k x N. g N N & (ALL n. g n N --> dist m (f n, x) < k)" - -lemma bounded: "ALL m g f. +%(m::'a::type metric, g::'b::type => 'b::type => bool) + f::'b::type => 'a::type. + EX (k::real) (x::'a::type) N::'b::type. + g N N & (ALL n::'b::type. g n N --> dist m (f n, x) < k)" + +lemma bounded: "ALL (m::'a::type metric) (g::'b::type => 'b::type => bool) + f::'b::type => 'a::type. bounded (m, g) f = - (EX k x N. g N N & (ALL n. g n N --> dist m (f n, x) < k))" + (EX (k::real) (x::'a::type) N::'b::type. + g N N & (ALL n::'b::type. g n N --> dist m (f n, x) < k))" by (import nets bounded) constdefs - tendsto :: "'a metric * 'a => 'a => 'a => bool" - "tendsto == %(m, x) y z. 0 < dist m (x, y) & dist m (x, y) <= dist m (x, z)" - -lemma tendsto: "ALL m x y z. - tendsto (m, x) y z = (0 < dist m (x, y) & dist m (x, y) <= dist m (x, z))" + tendsto :: "'a::type metric * 'a::type => 'a::type => 'a::type => bool" + "tendsto == +%(m::'a::type metric, x::'a::type) (y::'a::type) z::'a::type. + (0::real) < dist m (x, y) & dist m (x, y) <= dist m (x, z)" + +lemma tendsto: "ALL (m::'a::type metric) (x::'a::type) (y::'a::type) z::'a::type. + tendsto (m, x) y z = + ((0::real) < dist m (x, y) & dist m (x, y) <= dist m (x, z))" by (import nets tendsto) -lemma DORDER_LEMMA: "ALL g. +lemma DORDER_LEMMA: "ALL g::'a::type => 'a::type => bool. dorder g --> - (ALL P Q. - (EX n. g n n & (ALL m. g m n --> P m)) & - (EX n. g n n & (ALL m. g m n --> Q m)) --> - (EX n. g n n & (ALL m. g m n --> P m & Q m)))" + (ALL (P::'a::type => bool) Q::'a::type => bool. + (EX n::'a::type. g n n & (ALL m::'a::type. g m n --> P m)) & + (EX n::'a::type. g n n & (ALL m::'a::type. g m n --> Q m)) --> + (EX n::'a::type. g n n & (ALL m::'a::type. g m n --> P m & Q m)))" by (import nets DORDER_LEMMA) lemma DORDER_NGE: "dorder nat_ge" by (import nets DORDER_NGE) -lemma DORDER_TENDSTO: "ALL m x. dorder (tendsto (m, x))" +lemma DORDER_TENDSTO: "ALL (m::'a::type metric) x::'a::type. dorder (tendsto (m, x))" by (import nets DORDER_TENDSTO) -lemma MTOP_TENDS: "ALL d g x x0. +lemma MTOP_TENDS: "ALL (d::'a::type metric) (g::'b::type => 'b::type => bool) + (x::'b::type => 'a::type) x0::'a::type. tends x x0 (mtop d, g) = - (ALL e>0. EX n. g n n & (ALL m. g m n --> dist d (x m, x0) < e))" + (ALL e>0::real. + EX n::'b::type. + g n n & (ALL m::'b::type. g m n --> dist d (x m, x0) < e))" by (import nets MTOP_TENDS) -lemma MTOP_TENDS_UNIQ: "ALL (g::'b => 'b => bool) d::'a metric. +lemma MTOP_TENDS_UNIQ: "ALL (g::'b::type => 'b::type => bool) d::'a::type metric. dorder g --> - tends (x::'b => 'a) (x0::'a) (mtop d, g) & - tends x (x1::'a) (mtop d, g) --> + tends (x::'b::type => 'a::type) (x0::'a::type) (mtop d, g) & + tends x (x1::'a::type) (mtop d, g) --> x0 = x1" by (import nets MTOP_TENDS_UNIQ) -lemma SEQ_TENDS: "ALL d x x0. +lemma SEQ_TENDS: "ALL (d::'a::type metric) (x::nat => 'a::type) x0::'a::type. tends x x0 (mtop d, nat_ge) = - (ALL xa>0. EX xb. ALL xc. xb <= xc --> dist d (x xc, x0) < xa)" + (ALL xa>0::real. + EX xb::nat. ALL xc::nat. xb <= xc --> dist d (x xc, x0) < xa)" by (import nets SEQ_TENDS) -lemma LIM_TENDS: "ALL m1 m2 f x0 y0. +lemma LIM_TENDS: "ALL (m1::'a::type metric) (m2::'b::type metric) (f::'a::type => 'b::type) + (x0::'a::type) y0::'b::type. limpt (mtop m1) x0 re_universe --> tends f y0 (mtop m2, tendsto (m1, x0)) = - (ALL e>0. - EX d>0. - ALL x. - 0 < dist m1 (x, x0) & dist m1 (x, x0) <= d --> + (ALL e>0::real. + EX d>0::real. + ALL x::'a::type. + (0::real) < dist m1 (x, x0) & dist m1 (x, x0) <= d --> dist m2 (f x, y0) < e)" by (import nets LIM_TENDS) -lemma LIM_TENDS2: "ALL m1 m2 f x0 y0. +lemma LIM_TENDS2: "ALL (m1::'a::type metric) (m2::'b::type metric) (f::'a::type => 'b::type) + (x0::'a::type) y0::'b::type. limpt (mtop m1) x0 re_universe --> tends f y0 (mtop m2, tendsto (m1, x0)) = - (ALL e>0. - EX d>0. - ALL x. - 0 < dist m1 (x, x0) & dist m1 (x, x0) < d --> + (ALL e>0::real. + EX d>0::real. + ALL x::'a::type. + (0::real) < dist m1 (x, x0) & dist m1 (x, x0) < d --> dist m2 (f x, y0) < e)" by (import nets LIM_TENDS2) -lemma MR1_BOUNDED: "ALL g f. - bounded (mr1, g) f = (EX k N. g N N & (ALL n. g n N --> abs (f n) < k))" +lemma MR1_BOUNDED: "ALL (g::'a::type => 'a::type => bool) f::'a::type => real. + bounded (mr1, g) f = + (EX (k::real) N::'a::type. + g N N & (ALL n::'a::type. g n N --> abs (f n) < k))" by (import nets MR1_BOUNDED) -lemma NET_NULL: "ALL g x x0. tends x x0 (mtop mr1, g) = tends (%n. x n - x0) 0 (mtop mr1, g)" +lemma NET_NULL: "ALL (g::'a::type => 'a::type => bool) (x::'a::type => real) x0::real. + tends x x0 (mtop mr1, g) = + tends (%n::'a::type. x n - x0) (0::real) (mtop mr1, g)" by (import nets NET_NULL) -lemma NET_CONV_BOUNDED: "ALL g x x0. tends x x0 (mtop mr1, g) --> bounded (mr1, g) x" +lemma NET_CONV_BOUNDED: "ALL (g::'a::type => 'a::type => bool) (x::'a::type => real) x0::real. + tends x x0 (mtop mr1, g) --> bounded (mr1, g) x" by (import nets NET_CONV_BOUNDED) -lemma NET_CONV_NZ: "ALL g x x0. - tends x x0 (mtop mr1, g) & x0 ~= 0 --> - (EX N. g N N & (ALL n. g n N --> x n ~= 0))" +lemma NET_CONV_NZ: "ALL (g::'a::type => 'a::type => bool) (x::'a::type => real) x0::real. + tends x x0 (mtop mr1, g) & x0 ~= (0::real) --> + (EX N::'a::type. g N N & (ALL n::'a::type. g n N --> x n ~= (0::real)))" by (import nets NET_CONV_NZ) -lemma NET_CONV_IBOUNDED: "ALL g x x0. - tends x x0 (mtop mr1, g) & x0 ~= 0 --> - bounded (mr1, g) (%n. inverse (x n))" +lemma NET_CONV_IBOUNDED: "ALL (g::'a::type => 'a::type => bool) (x::'a::type => real) x0::real. + tends x x0 (mtop mr1, g) & x0 ~= (0::real) --> + bounded (mr1, g) (%n::'a::type. inverse (x n))" by (import nets NET_CONV_IBOUNDED) -lemma NET_NULL_ADD: "ALL g. +lemma NET_NULL_ADD: "ALL g::'a::type => 'a::type => bool. dorder g --> - (ALL x y. - tends x 0 (mtop mr1, g) & tends y 0 (mtop mr1, g) --> - tends (%n. x n + y n) 0 (mtop mr1, g))" + (ALL (x::'a::type => real) y::'a::type => real. + tends x (0::real) (mtop mr1, g) & tends y (0::real) (mtop mr1, g) --> + tends (%n::'a::type. x n + y n) (0::real) (mtop mr1, g))" by (import nets NET_NULL_ADD) -lemma NET_NULL_MUL: "ALL g. +lemma NET_NULL_MUL: "ALL g::'a::type => 'a::type => bool. dorder g --> - (ALL x y. - bounded (mr1, g) x & tends y 0 (mtop mr1, g) --> - tends (%n. x n * y n) 0 (mtop mr1, g))" + (ALL (x::'a::type => real) y::'a::type => real. + bounded (mr1, g) x & tends y (0::real) (mtop mr1, g) --> + tends (%n::'a::type. x n * y n) (0::real) (mtop mr1, g))" by (import nets NET_NULL_MUL) -lemma NET_NULL_CMUL: "ALL g k x. tends x 0 (mtop mr1, g) --> tends (%n. k * x n) 0 (mtop mr1, g)" +lemma NET_NULL_CMUL: "ALL (g::'a::type => 'a::type => bool) (k::real) x::'a::type => real. + tends x (0::real) (mtop mr1, g) --> + tends (%n::'a::type. k * x n) (0::real) (mtop mr1, g)" by (import nets NET_NULL_CMUL) -lemma NET_ADD: "ALL g. +lemma NET_ADD: "ALL g::'a::type => 'a::type => bool. dorder g --> - (ALL x x0 y y0. + (ALL (x::'a::type => real) (x0::real) (y::'a::type => real) y0::real. tends x x0 (mtop mr1, g) & tends y y0 (mtop mr1, g) --> - tends (%n. x n + y n) (x0 + y0) (mtop mr1, g))" + tends (%n::'a::type. x n + y n) (x0 + y0) (mtop mr1, g))" by (import nets NET_ADD) -lemma NET_NEG: "ALL g. +lemma NET_NEG: "ALL g::'a::type => 'a::type => bool. dorder g --> - (ALL x x0. - tends x x0 (mtop mr1, g) = tends (%n. - x n) (- x0) (mtop mr1, g))" + (ALL (x::'a::type => real) x0::real. + tends x x0 (mtop mr1, g) = + tends (%n::'a::type. - x n) (- x0) (mtop mr1, g))" by (import nets NET_NEG) -lemma NET_SUB: "ALL g. +lemma NET_SUB: "ALL g::'a::type => 'a::type => bool. dorder g --> - (ALL x x0 y y0. + (ALL (x::'a::type => real) (x0::real) (y::'a::type => real) y0::real. tends x x0 (mtop mr1, g) & tends y y0 (mtop mr1, g) --> - tends (%xa. x xa - y xa) (x0 - y0) (mtop mr1, g))" + tends (%xa::'a::type. x xa - y xa) (x0 - y0) (mtop mr1, g))" by (import nets NET_SUB) -lemma NET_MUL: "ALL g. +lemma NET_MUL: "ALL g::'a::type => 'a::type => bool. dorder g --> - (ALL x y x0 y0. + (ALL (x::'a::type => real) (y::'a::type => real) (x0::real) y0::real. tends x x0 (mtop mr1, g) & tends y y0 (mtop mr1, g) --> - tends (%n. x n * y n) (x0 * y0) (mtop mr1, g))" + tends (%n::'a::type. x n * y n) (x0 * y0) (mtop mr1, g))" by (import nets NET_MUL) -lemma NET_INV: "ALL g. +lemma NET_INV: "ALL g::'a::type => 'a::type => bool. dorder g --> - (ALL x x0. - tends x x0 (mtop mr1, g) & x0 ~= 0 --> - tends (%n. inverse (x n)) (inverse x0) (mtop mr1, g))" + (ALL (x::'a::type => real) x0::real. + tends x x0 (mtop mr1, g) & x0 ~= (0::real) --> + tends (%n::'a::type. inverse (x n)) (inverse x0) (mtop mr1, g))" by (import nets NET_INV) -lemma NET_DIV: "ALL g. +lemma NET_DIV: "ALL g::'a::type => 'a::type => bool. dorder g --> - (ALL x x0 y y0. - tends x x0 (mtop mr1, g) & tends y y0 (mtop mr1, g) & y0 ~= 0 --> - tends (%xa. x xa / y xa) (x0 / y0) (mtop mr1, g))" + (ALL (x::'a::type => real) (x0::real) (y::'a::type => real) y0::real. + tends x x0 (mtop mr1, g) & + tends y y0 (mtop mr1, g) & y0 ~= (0::real) --> + tends (%xa::'a::type. x xa / y xa) (x0 / y0) (mtop mr1, g))" by (import nets NET_DIV) -lemma NET_ABS: "ALL g x x0. - tends x x0 (mtop mr1, g) --> tends (%n. abs (x n)) (abs x0) (mtop mr1, g)" +lemma NET_ABS: "ALL (g::'a::type => 'a::type => bool) (x::'a::type => real) x0::real. + tends x x0 (mtop mr1, g) --> + tends (%n::'a::type. abs (x n)) (abs x0) (mtop mr1, g)" by (import nets NET_ABS) -lemma NET_LE: "ALL g. +lemma NET_LE: "ALL g::'a::type => 'a::type => bool. dorder g --> - (ALL x x0 y y0. + (ALL (x::'a::type => real) (x0::real) (y::'a::type => real) y0::real. tends x x0 (mtop mr1, g) & tends y y0 (mtop mr1, g) & - (EX N. g N N & (ALL n. g n N --> x n <= y n)) --> + (EX N::'a::type. g N N & (ALL n::'a::type. g n N --> x n <= y n)) --> x0 <= y0)" by (import nets NET_LE) @@ -1175,9 +1323,9 @@ "-->" :: "(nat => real) => real => bool" ("-->") defs - "-->_def": "--> == %x x0. tends x x0 (mtop mr1, nat_ge)" - -lemma tends_num_real: "ALL x x0. --> x x0 = tends x x0 (mtop mr1, nat_ge)" + "-->_def": "--> == %(x::nat => real) x0::real. tends x x0 (mtop mr1, nat_ge)" + +lemma tends_num_real: "ALL (x::nat => real) x0::real. --> x x0 = tends x x0 (mtop mr1, nat_ge)" by (import seq tends_num_real) lemma SEQ: "(All::((nat => real) => bool) => bool) @@ -1202,7 +1350,7 @@ e))))))))" by (import seq SEQ) -lemma SEQ_CONST: "ALL k. --> (%x. k) k" +lemma SEQ_CONST: "ALL k::real. --> (%x::nat. k) k" by (import seq SEQ_CONST) lemma SEQ_ADD: "(All::((nat => real) => bool) => bool) @@ -1239,7 +1387,7 @@ ((op *::real => real => real) x0 y0))))))" by (import seq SEQ_MUL) -lemma SEQ_NEG: "ALL x x0. --> x x0 = --> (%n. - x n) (- x0)" +lemma SEQ_NEG: "ALL (x::nat => real) x0::real. --> x x0 = --> (%n::nat. - x n) (- x0)" by (import seq SEQ_NEG) lemma SEQ_INV: "(All::((nat => real) => bool) => bool) @@ -1308,9 +1456,9 @@ constdefs convergent :: "(nat => real) => bool" - "convergent == %f. Ex (--> f)" - -lemma convergent: "ALL f. convergent f = Ex (--> f)" + "convergent == %f::nat => real. Ex (--> f)" + +lemma convergent: "ALL f::nat => real. convergent f = Ex (--> f)" by (import seq convergent) constdefs @@ -1363,12 +1511,12 @@ constdefs lim :: "(nat => real) => real" - "lim == %f. Eps (--> f)" - -lemma lim: "ALL f. lim f = Eps (--> f)" + "lim == %f::nat => real. Eps (--> f)" + +lemma lim: "ALL f::nat => real. lim f = Eps (--> f)" by (import seq lim) -lemma SEQ_LIM: "ALL f. convergent f = --> f (lim f)" +lemma SEQ_LIM: "ALL f::nat => real. convergent f = --> f (lim f)" by (import seq SEQ_LIM) constdefs @@ -1396,7 +1544,7 @@ ((op <::nat => nat => bool) (f m) (f n))))))" by (import seq subseq) -lemma SUBSEQ_SUC: "ALL f. subseq f = (ALL n. f n < f (Suc n))" +lemma SUBSEQ_SUC: "ALL f::nat => nat. subseq f = (ALL n::nat. f n < f (Suc n))" by (import seq SUBSEQ_SUC) consts @@ -1442,7 +1590,9 @@ ((op <=::real => real => bool) (f n) (f m)))))))" by (import seq mono) -lemma MONO_SUC: "ALL f. seq.mono f = ((ALL x. f x <= f (Suc x)) | (ALL n. f (Suc n) <= f n))" +lemma MONO_SUC: "ALL f::nat => real. + seq.mono f = + ((ALL x::nat. f x <= f (Suc x)) | (ALL n::nat. f (Suc n) <= f n))" by (import seq MONO_SUC) lemma MAX_LEMMA: "(All::((nat => real) => bool) => bool) @@ -1459,7 +1609,8 @@ ((abs::real => real) (s n)) k)))))" by (import seq MAX_LEMMA) -lemma SEQ_BOUNDED: "ALL s. bounded (mr1, nat_ge) s = (EX k. ALL n. abs (s n) < k)" +lemma SEQ_BOUNDED: "ALL s::nat => real. + bounded (mr1, nat_ge) s = (EX k::real. ALL n::nat. abs (s n) < k)" by (import seq SEQ_BOUNDED) lemma SEQ_BOUNDED_2: "(All::((nat => real) => bool) => bool) @@ -1516,10 +1667,11 @@ ((convergent::(nat => real) => bool) f))" by (import seq SEQ_ICONV) -lemma SEQ_NEG_CONV: "ALL f. convergent f = convergent (%n. - f n)" +lemma SEQ_NEG_CONV: "ALL f::nat => real. convergent f = convergent (%n::nat. - f n)" by (import seq SEQ_NEG_CONV) -lemma SEQ_NEG_BOUNDED: "ALL f. bounded (mr1, nat_ge) (%n. - f n) = bounded (mr1, nat_ge) f" +lemma SEQ_NEG_BOUNDED: "ALL f::nat => real. + bounded (mr1, nat_ge) (%n::nat. - f n) = bounded (mr1, nat_ge) f" by (import seq SEQ_NEG_BOUNDED) lemma SEQ_BCONV: "(All::((nat => real) => bool) => bool) @@ -1537,7 +1689,7 @@ ((convergent::(nat => real) => bool) f))" by (import seq SEQ_BCONV) -lemma SEQ_MONOSUB: "ALL s. EX f. subseq f & seq.mono (%n. s (f n))" +lemma SEQ_MONOSUB: "ALL s::nat => real. EX f::nat => nat. subseq f & seq.mono (%n::nat. s (f n))" by (import seq SEQ_MONOSUB) lemma SEQ_SBOUNDED: "(All::((nat => real) => bool) => bool) @@ -1582,7 +1734,7 @@ ((op <=::nat => nat => bool) N2 (f x)))))))" by (import seq SEQ_DIRECT) -lemma SEQ_CAUCHY: "ALL f. cauchy f = convergent f" +lemma SEQ_CAUCHY: "ALL f::nat => real. cauchy f = convergent f" by (import seq SEQ_CAUCHY) lemma SEQ_LE: "(All::((nat => real) => bool) => bool) @@ -1609,10 +1761,10 @@ ((op <=::real => real => bool) l m)))))" by (import seq SEQ_LE) -lemma SEQ_SUC: "ALL f l. --> f l = --> (%n. f (Suc n)) l" +lemma SEQ_SUC: "ALL (f::nat => real) l::real. --> f l = --> (%n::nat. f (Suc n)) l" by (import seq SEQ_SUC) -lemma SEQ_ABS: "ALL f. --> (%n. abs (f n)) 0 = --> f 0" +lemma SEQ_ABS: "ALL f::nat => real. --> (%n::nat. abs (f n)) (0::real) = --> f (0::real)" by (import seq SEQ_ABS) lemma SEQ_ABS_IMP: "(All::((nat => real) => bool) => bool) @@ -1777,23 +1929,24 @@ constdefs sums :: "(nat => real) => real => bool" - "sums == %f. --> (%n. real.sum (0, n) f)" - -lemma sums: "ALL f s. sums f s = --> (%n. real.sum (0, n) f) s" + "sums == %f::nat => real. --> (%n::nat. real.sum (0::nat, n) f)" + +lemma sums: "ALL (f::nat => real) s::real. + sums f s = --> (%n::nat. real.sum (0::nat, n) f) s" by (import seq sums) constdefs summable :: "(nat => real) => bool" - "summable == %f. Ex (sums f)" - -lemma summable: "ALL f. summable f = Ex (sums f)" + "summable == %f::nat => real. Ex (sums f)" + +lemma summable: "ALL f::nat => real. summable f = Ex (sums f)" by (import seq summable) constdefs suminf :: "(nat => real) => real" - "suminf == %f. Eps (sums f)" - -lemma suminf: "ALL f. suminf f = Eps (sums f)" + "suminf == %f::nat => real. Eps (sums f)" + +lemma suminf: "ALL f::nat => real. suminf f = Eps (sums f)" by (import seq suminf) lemma SUM_SUMMABLE: "(All::((nat => real) => bool) => bool) @@ -2219,9 +2372,12 @@ constdefs tends_real_real :: "(real => real) => real => real => bool" - "tends_real_real == %f l x0. tends f l (mtop mr1, tendsto (mr1, x0))" - -lemma tends_real_real: "ALL f l x0. tends_real_real f l x0 = tends f l (mtop mr1, tendsto (mr1, x0))" + "tends_real_real == +%(f::real => real) (l::real) x0::real. + tends f l (mtop mr1, tendsto (mr1, x0))" + +lemma tends_real_real: "ALL (f::real => real) (l::real) x0::real. + tends_real_real f l x0 = tends f l (mtop mr1, tendsto (mr1, x0))" by (import lim tends_real_real) lemma LIM: "(All::((real => real) => bool) => bool) @@ -2253,7 +2409,7 @@ ((abs::real => real) ((op -::real => real => real) (f x) y0)) e))))))))))" by (import lim LIM) -lemma LIM_CONST: "ALL k. All (tends_real_real (%x. k) k)" +lemma LIM_CONST: "ALL k::real. All (tends_real_real (%x::real. k) k)" by (import lim LIM_CONST) lemma LIM_ADD: "(All::((real => real) => bool) => bool) @@ -2306,7 +2462,8 @@ ((op *::real => real => real) l m) x))))))" by (import lim LIM_MUL) -lemma LIM_NEG: "ALL f l x. tends_real_real f l x = tends_real_real (%x. - f x) (- l) x" +lemma LIM_NEG: "ALL (f::real => real) (l::real) x::real. + tends_real_real f l x = tends_real_real (%x::real. - f x) (- l) x" by (import lim LIM_NEG) lemma LIM_INV: "(All::((real => real) => bool) => bool) @@ -2380,10 +2537,11 @@ ((op /::real => real => real) l m) x))))))" by (import lim LIM_DIV) -lemma LIM_NULL: "ALL f l x. tends_real_real f l x = tends_real_real (%x. f x - l) 0 x" +lemma LIM_NULL: "ALL (f::real => real) (l::real) x::real. + tends_real_real f l x = tends_real_real (%x::real. f x - l) (0::real) x" by (import lim LIM_NULL) -lemma LIM_X: "ALL x0. tends_real_real (%x. x) x0 x0" +lemma LIM_X: "ALL x0::real. tends_real_real (%x::real. x) x0 x0" by (import lim LIM_X) lemma LIM_UNIQ: "(All::((real => real) => bool) => bool) @@ -2454,23 +2612,31 @@ constdefs diffl :: "(real => real) => real => real => bool" - "diffl == %f l x. tends_real_real (%h. (f (x + h) - f x) / h) l 0" - -lemma diffl: "ALL f l x. diffl f l x = tends_real_real (%h. (f (x + h) - f x) / h) l 0" + "diffl == +%(f::real => real) (l::real) x::real. + tends_real_real (%h::real. (f (x + h) - f x) / h) l (0::real)" + +lemma diffl: "ALL (f::real => real) (l::real) x::real. + diffl f l x = + tends_real_real (%h::real. (f (x + h) - f x) / h) l (0::real)" by (import lim diffl) constdefs contl :: "(real => real) => real => bool" - "contl == %f x. tends_real_real (%h. f (x + h)) (f x) 0" - -lemma contl: "ALL f x. contl f x = tends_real_real (%h. f (x + h)) (f x) 0" + "contl == +%(f::real => real) x::real. + tends_real_real (%h::real. f (x + h)) (f x) (0::real)" + +lemma contl: "ALL (f::real => real) x::real. + contl f x = tends_real_real (%h::real. f (x + h)) (f x) (0::real)" by (import lim contl) constdefs differentiable :: "(real => real) => real => bool" - "differentiable == %f x. EX l. diffl f l x" - -lemma differentiable: "ALL f x. differentiable f x = (EX l. diffl f l x)" + "differentiable == %(f::real => real) x::real. EX l::real. diffl f l x" + +lemma differentiable: "ALL (f::real => real) x::real. + differentiable f x = (EX l::real. diffl f l x)" by (import lim differentiable) lemma DIFF_UNIQ: "(All::((real => real) => bool) => bool) @@ -2501,15 +2667,16 @@ ((contl::(real => real) => real => bool) f x))))" by (import lim DIFF_CONT) -lemma CONTL_LIM: "ALL f x. contl f x = tends_real_real f (f x) x" +lemma CONTL_LIM: "ALL (f::real => real) x::real. contl f x = tends_real_real f (f x) x" by (import lim CONTL_LIM) -lemma DIFF_CARAT: "ALL f l x. +lemma DIFF_CARAT: "ALL (f::real => real) (l::real) x::real. diffl f l x = - (EX g. (ALL z. f z - f x = g z * (z - x)) & contl g x & g x = l)" + (EX g::real => real. + (ALL z::real. f z - f x = g z * (z - x)) & contl g x & g x = l)" by (import lim DIFF_CARAT) -lemma CONT_CONST: "ALL k. All (contl (%x. k))" +lemma CONT_CONST: "ALL k::real. All (contl (%x::real. k))" by (import lim CONT_CONST) lemma CONT_ADD: "(All::((real => real) => bool) => bool) @@ -2672,7 +2839,7 @@ ((op =::real => real => bool) (f x) y))))))))" by (import lim IVT2) -lemma DIFF_CONST: "ALL k. All (diffl (%x. k) 0)" +lemma DIFF_CONST: "ALL k::real. All (diffl (%x::real. k) (0::real))" by (import lim DIFF_CONST) lemma DIFF_ADD: "(All::((real => real) => bool) => bool) @@ -2793,10 +2960,11 @@ ((op *::real => real => real) l m) x))))))" by (import lim DIFF_CHAIN) -lemma DIFF_X: "All (diffl (%x. x) 1)" +lemma DIFF_X: "All (diffl (%x::real. x) (1::real))" by (import lim DIFF_X) -lemma DIFF_POW: "ALL n x. diffl (%x. x ^ n) (real n * x ^ (n - 1)) x" +lemma DIFF_POW: "ALL (n::nat) x::real. + diffl (%x::real. x ^ n) (real n * x ^ (n - (1::nat))) x" by (import lim DIFF_POW) lemma DIFF_XM1: "(All::(real => bool) => bool) @@ -3708,19 +3876,19 @@ ;setup_theory powser -lemma POWDIFF_LEMMA: "ALL n x y. - real.sum (0, Suc n) (%p. x ^ p * y ^ (Suc n - p)) = - y * real.sum (0, Suc n) (%p. x ^ p * y ^ (n - p))" +lemma POWDIFF_LEMMA: "ALL (n::nat) (x::real) y::real. + real.sum (0::nat, Suc n) (%p::nat. x ^ p * y ^ (Suc n - p)) = + y * real.sum (0::nat, Suc n) (%p::nat. x ^ p * y ^ (n - p))" by (import powser POWDIFF_LEMMA) -lemma POWDIFF: "ALL n x y. +lemma POWDIFF: "ALL (n::nat) (x::real) y::real. x ^ Suc n - y ^ Suc n = - (x - y) * real.sum (0, Suc n) (%p. x ^ p * y ^ (n - p))" + (x - y) * real.sum (0::nat, Suc n) (%p::nat. x ^ p * y ^ (n - p))" by (import powser POWDIFF) -lemma POWREV: "ALL n x y. - real.sum (0, Suc n) (%xa. x ^ xa * y ^ (n - xa)) = - real.sum (0, Suc n) (%xa. x ^ (n - xa) * y ^ xa)" +lemma POWREV: "ALL (n::nat) (x::real) y::real. + real.sum (0::nat, Suc n) (%xa::nat. x ^ xa * y ^ (n - xa)) = + real.sum (0::nat, Suc n) (%xa::nat. x ^ (n - xa) * y ^ xa)" by (import powser POWREV) lemma POWSER_INSIDEA: "(All::((nat => real) => bool) => bool) @@ -3766,23 +3934,24 @@ constdefs diffs :: "(nat => real) => nat => real" - "diffs == %c n. real (Suc n) * c (Suc n)" - -lemma diffs: "ALL c. diffs c = (%n. real (Suc n) * c (Suc n))" + "diffs == %(c::nat => real) n::nat. real (Suc n) * c (Suc n)" + +lemma diffs: "ALL c::nat => real. diffs c = (%n::nat. real (Suc n) * c (Suc n))" by (import powser diffs) -lemma DIFFS_NEG: "ALL c. diffs (%n. - c n) = (%x. - diffs c x)" +lemma DIFFS_NEG: "ALL c::nat => real. diffs (%n::nat. - c n) = (%x::nat. - diffs c x)" by (import powser DIFFS_NEG) -lemma DIFFS_LEMMA: "ALL n c x. - real.sum (0, n) (%n. diffs c n * x ^ n) = - real.sum (0, n) (%n. real n * (c n * x ^ (n - 1))) + - real n * (c n * x ^ (n - 1))" +lemma DIFFS_LEMMA: "ALL (n::nat) (c::nat => real) x::real. + real.sum (0::nat, n) (%n::nat. diffs c n * x ^ n) = + real.sum (0::nat, n) (%n::nat. real n * (c n * x ^ (n - (1::nat)))) + + real n * (c n * x ^ (n - (1::nat)))" by (import powser DIFFS_LEMMA) -lemma DIFFS_LEMMA2: "ALL n c x. - real.sum (0, n) (%n. real n * (c n * x ^ (n - 1))) = - real.sum (0, n) (%n. diffs c n * x ^ n) - real n * (c n * x ^ (n - 1))" +lemma DIFFS_LEMMA2: "ALL (n::nat) (c::nat => real) x::real. + real.sum (0::nat, n) (%n::nat. real n * (c n * x ^ (n - (1::nat)))) = + real.sum (0::nat, n) (%n::nat. diffs c n * x ^ n) - + real n * (c n * x ^ (n - (1::nat)))" by (import powser DIFFS_LEMMA2) lemma DIFFS_EQUIV: "(All::((nat => real) => bool) => bool) @@ -3808,9 +3977,9 @@ ((op ^::real => nat => real) x n))))))" by (import powser DIFFS_EQUIV) -lemma TERMDIFF_LEMMA1: "ALL m z h. - real.sum (0, m) (%p. (z + h) ^ (m - p) * z ^ p - z ^ m) = - real.sum (0, m) (%p. z ^ p * ((z + h) ^ (m - p) - z ^ (m - p)))" +lemma TERMDIFF_LEMMA1: "ALL (m::nat) (z::real) h::real. + real.sum (0::nat, m) (%p::nat. (z + h) ^ (m - p) * z ^ p - z ^ m) = + real.sum (0::nat, m) (%p::nat. z ^ p * ((z + h) ^ (m - p) - z ^ (m - p)))" by (import powser TERMDIFF_LEMMA1) lemma TERMDIFF_LEMMA2: "(All::(real => bool) => bool) @@ -4009,79 +4178,119 @@ constdefs exp :: "real => real" - "exp == %x. suminf (%n. inverse (real (FACT n)) * x ^ n)" - -lemma exp: "ALL x. exp x = suminf (%n. inverse (real (FACT n)) * x ^ n)" + "exp == %x::real. suminf (%n::nat. inverse (real (FACT n)) * x ^ n)" + +lemma exp: "ALL x::real. exp x = suminf (%n::nat. inverse (real (FACT n)) * x ^ n)" by (import transc exp) constdefs cos :: "real => real" - "cos == -%x. suminf - (%n. (if EVEN n then (- 1) ^ (n div 2) / real (FACT n) else 0) * x ^ n)" - -lemma cos: "ALL x. - cos x = - suminf - (%n. (if EVEN n then (- 1) ^ (n div 2) / real (FACT n) else 0) * x ^ n)" + "(cos == + (%(x::real). + (suminf + (%(n::nat). + ((if (EVEN n) + then (((- (1::real)) ^ (n div (2::nat))) / (real (FACT n))) + else (0::real)) * + (x ^ n))))))" + +lemma cos: "(ALL (x::real). + ((cos x) = + (suminf + (%(n::nat). + ((if (EVEN n) + then (((- (1::real)) ^ (n div (2::nat))) / (real (FACT n))) + else (0::real)) * + (x ^ n))))))" by (import transc cos) constdefs sin :: "real => real" "sin == -%x. suminf - (%n. (if EVEN n then 0 else (- 1) ^ ((n - 1) div 2) / real (FACT n)) * - x ^ n)" - -lemma sin: "ALL x. +%x::real. + suminf + (%n::nat. + (if EVEN n then 0::real + else (- (1::real)) ^ ((n - (1::nat)) div (2::nat)) / + real (FACT n)) * + x ^ n)" + +lemma sin: "ALL x::real. sin x = suminf - (%n. (if EVEN n then 0 else (- 1) ^ ((n - 1) div 2) / real (FACT n)) * - x ^ n)" + (%n::nat. + (if EVEN n then 0::real + else (- (1::real)) ^ ((n - (1::nat)) div (2::nat)) / + real (FACT n)) * + x ^ n)" by (import transc sin) -lemma EXP_CONVERGES: "ALL x. sums (%n. inverse (real (FACT n)) * x ^ n) (exp x)" +lemma EXP_CONVERGES: "ALL x::real. sums (%n::nat. inverse (real (FACT n)) * x ^ n) (exp x)" by (import transc EXP_CONVERGES) -lemma SIN_CONVERGES: "ALL x. +lemma SIN_CONVERGES: "ALL x::real. sums - (%n. (if EVEN n then 0 else (- 1) ^ ((n - 1) div 2) / real (FACT n)) * - x ^ n) + (%n::nat. + (if EVEN n then 0::real + else (- (1::real)) ^ ((n - (1::nat)) div (2::nat)) / + real (FACT n)) * + x ^ n) (sin x)" by (import transc SIN_CONVERGES) -lemma COS_CONVERGES: "ALL x. - sums - (%n. (if EVEN n then (- 1) ^ (n div 2) / real (FACT n) else 0) * x ^ n) - (cos x)" +lemma COS_CONVERGES: "(ALL (x::real). + (sums + (%(n::nat). + ((if (EVEN n) + then (((- (1::real)) ^ (n div (2::nat))) / (real (FACT n))) + else (0::real)) * + (x ^ n))) + (cos x)))" by (import transc COS_CONVERGES) -lemma EXP_FDIFF: "diffs (%n. inverse (real (FACT n))) = (%n. inverse (real (FACT n)))" +lemma EXP_FDIFF: "diffs (%n::nat. inverse (real (FACT n))) = +(%n::nat. inverse (real (FACT n)))" by (import transc EXP_FDIFF) -lemma SIN_FDIFF: "diffs (%n. if EVEN n then 0 else (- 1) ^ ((n - 1) div 2) / real (FACT n)) = -(%n. if EVEN n then (- 1) ^ (n div 2) / real (FACT n) else 0)" +lemma SIN_FDIFF: "((diffs + (%(n::nat). + (if (EVEN n) then (0::real) + else (((- (1::real)) ^ ((n - (1::nat)) div (2::nat))) / + (real (FACT n)))))) = + (%(n::nat). + (if (EVEN n) + then (((- (1::real)) ^ (n div (2::nat))) / (real (FACT n))) + else (0::real))))" by (import transc SIN_FDIFF) -lemma COS_FDIFF: "diffs (%n. if EVEN n then (- 1) ^ (n div 2) / real (FACT n) else 0) = -(%n. - (if EVEN n then 0 else (- 1) ^ ((n - 1) div 2) / real (FACT n)))" +lemma COS_FDIFF: "((diffs + (%(n::nat). + (if (EVEN n) + then (((- (1::real)) ^ (n div (2::nat))) / (real (FACT n))) + else (0::real)))) = + (%(n::nat). + (- (if (EVEN n) then (0::real) + else (((- (1::real)) ^ ((n - (1::nat)) div (2::nat))) / + (real (FACT n)))))))" by (import transc COS_FDIFF) -lemma SIN_NEGLEMMA: "ALL x. +lemma SIN_NEGLEMMA: "ALL x::real. - sin x = suminf - (%n. - ((if EVEN n then 0 - else (- 1) ^ ((n - 1) div 2) / real (FACT n)) * - x ^ n))" + (%n::nat. + - ((if EVEN n then 0::real + else (- (1::real)) ^ ((n - (1::nat)) div (2::nat)) / + real (FACT n)) * + x ^ n))" by (import transc SIN_NEGLEMMA) -lemma DIFF_EXP: "ALL x. diffl exp (exp x) x" +lemma DIFF_EXP: "ALL x::real. diffl exp (exp x) x" by (import transc DIFF_EXP) -lemma DIFF_SIN: "ALL x. diffl sin (cos x) x" +lemma DIFF_SIN: "ALL x::real. diffl sin (cos x) x" by (import transc DIFF_SIN) -lemma DIFF_COS: "ALL x. diffl cos (- sin x) x" +lemma DIFF_COS: "ALL x::real. diffl cos (- sin x) x" by (import transc DIFF_COS) lemma DIFF_COMPOSITE: "(op &::bool => bool => bool) @@ -4193,43 +4402,43 @@ x))))))))))" by (import transc DIFF_COMPOSITE) -lemma EXP_0: "exp 0 = 1" +lemma EXP_0: "exp (0::real) = (1::real)" by (import transc EXP_0) -lemma EXP_LE_X: "ALL x>=0. 1 + x <= exp x" +lemma EXP_LE_X: "ALL x>=0::real. (1::real) + x <= exp x" by (import transc EXP_LE_X) -lemma EXP_LT_1: "ALL x>0. 1 < exp x" +lemma EXP_LT_1: "ALL x>0::real. (1::real) < exp x" by (import transc EXP_LT_1) -lemma EXP_ADD_MUL: "ALL x y. exp (x + y) * exp (- x) = exp y" +lemma EXP_ADD_MUL: "ALL (x::real) y::real. exp (x + y) * exp (- x) = exp y" by (import transc EXP_ADD_MUL) -lemma EXP_NEG_MUL: "ALL x. exp x * exp (- x) = 1" +lemma EXP_NEG_MUL: "ALL x::real. exp x * exp (- x) = (1::real)" by (import transc EXP_NEG_MUL) -lemma EXP_NEG_MUL2: "ALL x. exp (- x) * exp x = 1" +lemma EXP_NEG_MUL2: "ALL x::real. exp (- x) * exp x = (1::real)" by (import transc EXP_NEG_MUL2) -lemma EXP_NEG: "ALL x. exp (- x) = inverse (exp x)" +lemma EXP_NEG: "ALL x::real. exp (- x) = inverse (exp x)" by (import transc EXP_NEG) -lemma EXP_ADD: "ALL x y. exp (x + y) = exp x * exp y" +lemma EXP_ADD: "ALL (x::real) y::real. exp (x + y) = exp x * exp y" by (import transc EXP_ADD) -lemma EXP_POS_LE: "ALL x. 0 <= exp x" +lemma EXP_POS_LE: "ALL x::real. (0::real) <= exp x" by (import transc EXP_POS_LE) -lemma EXP_NZ: "ALL x. exp x ~= 0" +lemma EXP_NZ: "ALL x::real. exp x ~= (0::real)" by (import transc EXP_NZ) -lemma EXP_POS_LT: "ALL x. 0 < exp x" +lemma EXP_POS_LT: "ALL x::real. (0::real) < exp x" by (import transc EXP_POS_LT) -lemma EXP_N: "ALL n x. exp (real n * x) = exp x ^ n" +lemma EXP_N: "ALL (n::nat) x::real. exp (real n * x) = exp x ^ n" by (import transc EXP_N) -lemma EXP_SUB: "ALL x y. exp (x - y) = exp x / exp y" +lemma EXP_SUB: "ALL (x::real) y::real. exp (x - y) = exp x / exp y" by (import transc EXP_SUB) lemma EXP_MONO_IMP: "(All::(real => bool) => bool) @@ -4241,32 +4450,32 @@ ((exp::real => real) y))))" by (import transc EXP_MONO_IMP) -lemma EXP_MONO_LT: "ALL x y. (exp x < exp y) = (x < y)" +lemma EXP_MONO_LT: "ALL (x::real) y::real. (exp x < exp y) = (x < y)" by (import transc EXP_MONO_LT) -lemma EXP_MONO_LE: "ALL x y. (exp x <= exp y) = (x <= y)" +lemma EXP_MONO_LE: "ALL (x::real) y::real. (exp x <= exp y) = (x <= y)" by (import transc EXP_MONO_LE) -lemma EXP_INJ: "ALL x y. (exp x = exp y) = (x = y)" +lemma EXP_INJ: "ALL (x::real) y::real. (exp x = exp y) = (x = y)" by (import transc EXP_INJ) -lemma EXP_TOTAL_LEMMA: "ALL y>=1. EX x>=0. x <= y - 1 & exp x = y" +lemma EXP_TOTAL_LEMMA: "ALL y>=1::real. EX x>=0::real. x <= y - (1::real) & exp x = y" by (import transc EXP_TOTAL_LEMMA) -lemma EXP_TOTAL: "ALL y>0. EX x. exp x = y" +lemma EXP_TOTAL: "ALL y>0::real. EX x::real. exp x = y" by (import transc EXP_TOTAL) constdefs ln :: "real => real" - "ln == %x. SOME u. exp u = x" - -lemma ln: "ALL x. ln x = (SOME u. exp u = x)" + "ln == %x::real. SOME u::real. exp u = x" + +lemma ln: "ALL x::real. ln x = (SOME u::real. exp u = x)" by (import transc ln) -lemma LN_EXP: "ALL x. ln (exp x) = x" +lemma LN_EXP: "ALL x::real. ln (exp x) = x" by (import transc LN_EXP) -lemma EXP_LN: "ALL x. (exp (ln x) = x) = (0 < x)" +lemma EXP_LN: "ALL x::real. (exp (ln x) = x) = ((0::real) < x)" by (import transc EXP_LN) lemma LN_MUL: "(All::(real => bool) => bool) @@ -4297,10 +4506,10 @@ ((op =::real => real => bool) x y))))" by (import transc LN_INJ) -lemma LN_1: "ln 1 = 0" +lemma LN_1: "ln (1::real) = (0::real)" by (import transc LN_1) -lemma LN_INV: "ALL x>0. ln (inverse x) = - ln x" +lemma LN_INV: "ALL x>0::real. ln (inverse x) = - ln x" by (import transc LN_INV) lemma LN_DIV: "(All::(real => bool) => bool) @@ -4357,13 +4566,13 @@ ((ln::real => real) x)))))" by (import transc LN_POW) -lemma LN_LE: "ALL x>=0. ln (1 + x) <= x" +lemma LN_LE: "ALL x>=0::real. ln ((1::real) + x) <= x" by (import transc LN_LE) -lemma LN_LT_X: "ALL x>0. ln x < x" +lemma LN_LT_X: "ALL x>0::real. ln x < x" by (import transc LN_LT_X) -lemma LN_POS: "ALL x>=1. 0 <= ln x" +lemma LN_POS: "ALL x>=1::real. (0::real) <= ln x" by (import transc LN_POS) constdefs @@ -4397,9 +4606,9 @@ constdefs sqrt :: "real => real" - "sqrt == root 2" - -lemma sqrt: "ALL x. sqrt x = root 2 x" + "sqrt == root (2::nat)" + +lemma sqrt: "ALL x::real. sqrt x = root (2::nat) x" by (import transc sqrt) lemma ROOT_LT_LEMMA: "(All::(nat => bool) => bool) @@ -4430,10 +4639,10 @@ ((real::nat => real) ((Suc::nat => nat) n)))))))" by (import transc ROOT_LN) -lemma ROOT_0: "ALL n. root (Suc n) 0 = 0" +lemma ROOT_0: "ALL n::nat. root (Suc n) (0::real) = (0::real)" by (import transc ROOT_0) -lemma ROOT_1: "ALL n. root (Suc n) 1 = 1" +lemma ROOT_1: "ALL n::nat. root (Suc n) (1::real) = (1::real)" by (import transc ROOT_1) lemma ROOT_POS_LT: "(All::(nat => bool) => bool) @@ -4564,22 +4773,22 @@ ((root::nat => real => real) ((Suc::nat => nat) n) y))))" by (import transc ROOT_MONO_LE) -lemma SQRT_0: "sqrt 0 = 0" +lemma SQRT_0: "sqrt (0::real) = (0::real)" by (import transc SQRT_0) -lemma SQRT_1: "sqrt 1 = 1" +lemma SQRT_1: "sqrt (1::real) = (1::real)" by (import transc SQRT_1) -lemma SQRT_POS_LT: "ALL x>0. 0 < sqrt x" +lemma SQRT_POS_LT: "ALL x>0::real. (0::real) < sqrt x" by (import transc SQRT_POS_LT) -lemma SQRT_POS_LE: "ALL x>=0. 0 <= sqrt x" +lemma SQRT_POS_LE: "ALL x>=0::real. (0::real) <= sqrt x" by (import transc SQRT_POS_LE) -lemma SQRT_POW2: "ALL x. (sqrt x ^ 2 = x) = (0 <= x)" +lemma SQRT_POW2: "ALL x::real. (sqrt x ^ 2 = x) = ((0::real) <= x)" by (import transc SQRT_POW2) -lemma SQRT_POW_2: "ALL x>=0. sqrt x ^ 2 = x" +lemma SQRT_POW_2: "ALL x>=0::real. sqrt x ^ 2 = x" by (import transc SQRT_POW_2) lemma POW_2_SQRT: "(op -->::bool => bool => bool) @@ -4628,7 +4837,7 @@ ((sqrt::real => real) xa)))))" by (import transc SQRT_MUL) -lemma SQRT_INV: "ALL x>=0. sqrt (inverse x) = inverse (sqrt x)" +lemma SQRT_INV: "ALL x>=0::real. sqrt (inverse x) = inverse (sqrt x)" by (import transc SQRT_INV) lemma SQRT_DIV: "(All::(real => bool) => bool) @@ -4682,7 +4891,7 @@ (bit.B0::bit)))))))" by (import transc SQRT_EVEN_POW2) -lemma REAL_DIV_SQRT: "ALL x>=0. x / sqrt x = sqrt x" +lemma REAL_DIV_SQRT: "ALL x>=0::real. x / sqrt x = sqrt x" by (import transc REAL_DIV_SQRT) lemma SQRT_EQ: "(All::(real => bool) => bool) @@ -4703,56 +4912,60 @@ ((op =::real => real => bool) x ((sqrt::real => real) y))))" by (import transc SQRT_EQ) -lemma SIN_0: "sin 0 = 0" +lemma SIN_0: "sin (0::real) = (0::real)" by (import transc SIN_0) -lemma COS_0: "cos 0 = 1" +lemma COS_0: "cos (0::real) = (1::real)" by (import transc COS_0) -lemma SIN_CIRCLE: "ALL x. sin x ^ 2 + cos x ^ 2 = 1" +lemma SIN_CIRCLE: "ALL x::real. sin x ^ 2 + cos x ^ 2 = (1::real)" by (import transc SIN_CIRCLE) -lemma SIN_BOUND: "ALL x. abs (sin x) <= 1" +lemma SIN_BOUND: "ALL x::real. abs (sin x) <= (1::real)" by (import transc SIN_BOUND) -lemma SIN_BOUNDS: "ALL x. - 1 <= sin x & sin x <= 1" +lemma SIN_BOUNDS: "ALL x::real. - (1::real) <= sin x & sin x <= (1::real)" by (import transc SIN_BOUNDS) -lemma COS_BOUND: "ALL x. abs (cos x) <= 1" +lemma COS_BOUND: "ALL x::real. abs (cos x) <= (1::real)" by (import transc COS_BOUND) -lemma COS_BOUNDS: "ALL x. - 1 <= cos x & cos x <= 1" +lemma COS_BOUNDS: "ALL x::real. - (1::real) <= cos x & cos x <= (1::real)" by (import transc COS_BOUNDS) -lemma SIN_COS_ADD: "ALL x y. +lemma SIN_COS_ADD: "ALL (x::real) y::real. (sin (x + y) - (sin x * cos y + cos x * sin y)) ^ 2 + (cos (x + y) - (cos x * cos y - sin x * sin y)) ^ 2 = - 0" + (0::real)" by (import transc SIN_COS_ADD) -lemma SIN_COS_NEG: "ALL x. (sin (- x) + sin x) ^ 2 + (cos (- x) - cos x) ^ 2 = 0" +lemma SIN_COS_NEG: "ALL x::real. (sin (- x) + sin x) ^ 2 + (cos (- x) - cos x) ^ 2 = (0::real)" by (import transc SIN_COS_NEG) -lemma SIN_ADD: "ALL x y. sin (x + y) = sin x * cos y + cos x * sin y" +lemma SIN_ADD: "ALL (x::real) y::real. sin (x + y) = sin x * cos y + cos x * sin y" by (import transc SIN_ADD) -lemma COS_ADD: "ALL x y. cos (x + y) = cos x * cos y - sin x * sin y" +lemma COS_ADD: "ALL (x::real) y::real. cos (x + y) = cos x * cos y - sin x * sin y" by (import transc COS_ADD) -lemma SIN_NEG: "ALL x. sin (- x) = - sin x" +lemma SIN_NEG: "ALL x::real. sin (- x) = - sin x" by (import transc SIN_NEG) -lemma COS_NEG: "ALL x. cos (- x) = cos x" +lemma COS_NEG: "ALL x::real. cos (- x) = cos x" by (import transc COS_NEG) -lemma SIN_DOUBLE: "ALL x. sin (2 * x) = 2 * (sin x * cos x)" +lemma SIN_DOUBLE: "ALL x::real. sin ((2::real) * x) = (2::real) * (sin x * cos x)" by (import transc SIN_DOUBLE) -lemma COS_DOUBLE: "ALL x. cos (2 * x) = cos x ^ 2 - sin x ^ 2" +lemma COS_DOUBLE: "ALL x::real. cos ((2::real) * x) = cos x ^ 2 - sin x ^ 2" by (import transc COS_DOUBLE) -lemma SIN_PAIRED: "ALL x. - sums (%n. (- 1) ^ n / real (FACT (2 * n + 1)) * x ^ (2 * n + 1)) (sin x)" +lemma SIN_PAIRED: "ALL x::real. + sums + (%n::nat. + (- (1::real)) ^ n / real (FACT ((2::nat) * n + (1::nat))) * + x ^ ((2::nat) * n + (1::nat))) + (sin x)" by (import transc SIN_PAIRED) lemma SIN_POS: "(All::(real => bool) => bool) @@ -4768,62 +4981,71 @@ ((op <::real => real => bool) (0::real) ((sin::real => real) x)))" by (import transc SIN_POS) -lemma COS_PAIRED: "ALL x. sums (%n. (- 1) ^ n / real (FACT (2 * n)) * x ^ (2 * n)) (cos x)" +lemma COS_PAIRED: "ALL x::real. + sums + (%n::nat. + (- (1::real)) ^ n / real (FACT ((2::nat) * n)) * x ^ ((2::nat) * n)) + (cos x)" by (import transc COS_PAIRED) -lemma COS_2: "cos 2 < 0" +lemma COS_2: "cos (2::real) < (0::real)" by (import transc COS_2) -lemma COS_ISZERO: "EX! x. 0 <= x & x <= 2 & cos x = 0" +lemma COS_ISZERO: "EX! x::real. (0::real) <= x & x <= (2::real) & cos x = (0::real)" by (import transc COS_ISZERO) constdefs pi :: "real" - "pi == 2 * (SOME x. 0 <= x & x <= 2 & cos x = 0)" - -lemma pi: "pi = 2 * (SOME x. 0 <= x & x <= 2 & cos x = 0)" + "pi == +(2::real) * +(SOME x::real. (0::real) <= x & x <= (2::real) & cos x = (0::real))" + +lemma pi: "pi = +(2::real) * +(SOME x::real. (0::real) <= x & x <= (2::real) & cos x = (0::real))" by (import transc pi) -lemma PI2: "pi / 2 = (SOME x. 0 <= x & x <= 2 & cos x = 0)" +lemma PI2: "pi / (2::real) = +(SOME x::real. (0::real) <= x & x <= (2::real) & cos x = (0::real))" by (import transc PI2) -lemma COS_PI2: "cos (pi / 2) = 0" +lemma COS_PI2: "cos (pi / (2::real)) = (0::real)" by (import transc COS_PI2) -lemma PI2_BOUNDS: "0 < pi / 2 & pi / 2 < 2" +lemma PI2_BOUNDS: "(0::real) < pi / (2::real) & pi / (2::real) < (2::real)" by (import transc PI2_BOUNDS) -lemma PI_POS: "0 < pi" +lemma PI_POS: "(0::real) < pi" by (import transc PI_POS) -lemma SIN_PI2: "sin (pi / 2) = 1" +lemma SIN_PI2: "sin (pi / (2::real)) = (1::real)" by (import transc SIN_PI2) -lemma COS_PI: "cos pi = - 1" +lemma COS_PI: "cos pi = - (1::real)" by (import transc COS_PI) -lemma SIN_PI: "sin pi = 0" +lemma SIN_PI: "sin pi = (0::real)" by (import transc SIN_PI) -lemma SIN_COS: "ALL x. sin x = cos (pi / 2 - x)" +lemma SIN_COS: "ALL x::real. sin x = cos (pi / (2::real) - x)" by (import transc SIN_COS) -lemma COS_SIN: "ALL x. cos x = sin (pi / 2 - x)" +lemma COS_SIN: "ALL x::real. cos x = sin (pi / (2::real) - x)" by (import transc COS_SIN) -lemma SIN_PERIODIC_PI: "ALL x. sin (x + pi) = - sin x" +lemma SIN_PERIODIC_PI: "ALL x::real. sin (x + pi) = - sin x" by (import transc SIN_PERIODIC_PI) -lemma COS_PERIODIC_PI: "ALL x. cos (x + pi) = - cos x" +lemma COS_PERIODIC_PI: "ALL x::real. cos (x + pi) = - cos x" by (import transc COS_PERIODIC_PI) -lemma SIN_PERIODIC: "ALL x. sin (x + 2 * pi) = sin x" +lemma SIN_PERIODIC: "ALL x::real. sin (x + (2::real) * pi) = sin x" by (import transc SIN_PERIODIC) -lemma COS_PERIODIC: "ALL x. cos (x + 2 * pi) = cos x" +lemma COS_PERIODIC: "ALL x::real. cos (x + (2::real) * pi) = cos x" by (import transc COS_PERIODIC) -lemma COS_NPI: "ALL n. cos (real n * pi) = (- 1) ^ n" +lemma COS_NPI: "ALL n::nat. cos (real n * pi) = (- (1::real)) ^ n" by (import transc COS_NPI) lemma SIN_NPI: "ALL n::nat. sin (real n * pi) = (0::real)" @@ -5036,38 +5258,38 @@ (bit.B0::bit)))))))))" by (import transc SIN_ZERO_LEMMA) -lemma COS_ZERO: "ALL x. - (cos x = 0) = - ((EX n. ~ EVEN n & x = real n * (pi / 2)) | - (EX n. ~ EVEN n & x = - (real n * (pi / 2))))" +lemma COS_ZERO: "ALL x::real. + (cos x = (0::real)) = + ((EX n::nat. ~ EVEN n & x = real n * (pi / (2::real))) | + (EX n::nat. ~ EVEN n & x = - (real n * (pi / (2::real)))))" by (import transc COS_ZERO) -lemma SIN_ZERO: "ALL x. - (sin x = 0) = - ((EX n. EVEN n & x = real n * (pi / 2)) | - (EX n. EVEN n & x = - (real n * (pi / 2))))" +lemma SIN_ZERO: "ALL x::real. + (sin x = (0::real)) = + ((EX n::nat. EVEN n & x = real n * (pi / (2::real))) | + (EX n::nat. EVEN n & x = - (real n * (pi / (2::real)))))" by (import transc SIN_ZERO) constdefs tan :: "real => real" - "tan == %x. sin x / cos x" - -lemma tan: "ALL x. tan x = sin x / cos x" + "tan == %x::real. sin x / cos x" + +lemma tan: "ALL x::real. tan x = sin x / cos x" by (import transc tan) -lemma TAN_0: "tan 0 = 0" +lemma TAN_0: "tan (0::real) = (0::real)" by (import transc TAN_0) -lemma TAN_PI: "tan pi = 0" +lemma TAN_PI: "tan pi = (0::real)" by (import transc TAN_PI) lemma TAN_NPI: "ALL n::nat. tan (real n * pi) = (0::real)" by (import transc TAN_NPI) -lemma TAN_NEG: "ALL x. tan (- x) = - tan x" +lemma TAN_NEG: "ALL x::real. tan (- x) = - tan x" by (import transc TAN_NEG) -lemma TAN_PERIODIC: "ALL x. tan (x + 2 * pi) = tan x" +lemma TAN_PERIODIC: "ALL x::real. tan (x + (2::real) * pi) = tan x" by (import transc TAN_PERIODIC) lemma TAN_ADD: "(All::(real => bool) => bool) @@ -5171,34 +5393,43 @@ x))" by (import transc DIFF_TAN) -lemma TAN_TOTAL_LEMMA: "ALL y>0. EX x>0. x < pi / 2 & y < tan x" +lemma TAN_TOTAL_LEMMA: "ALL y>0::real. EX x>0::real. x < pi / (2::real) & y < tan x" by (import transc TAN_TOTAL_LEMMA) -lemma TAN_TOTAL_POS: "ALL y>=0. EX x>=0. x < pi / 2 & tan x = y" +lemma TAN_TOTAL_POS: "ALL y>=0::real. EX x>=0::real. x < pi / (2::real) & tan x = y" by (import transc TAN_TOTAL_POS) -lemma TAN_TOTAL: "ALL y. EX! x. - (pi / 2) < x & x < pi / 2 & tan x = y" +lemma TAN_TOTAL: "ALL y::real. + EX! x::real. - (pi / (2::real)) < x & x < pi / (2::real) & tan x = y" by (import transc TAN_TOTAL) constdefs asn :: "real => real" - "asn == %y. SOME x. - (pi / 2) <= x & x <= pi / 2 & sin x = y" - -lemma asn: "ALL y. asn y = (SOME x. - (pi / 2) <= x & x <= pi / 2 & sin x = y)" + "asn == +%y::real. + SOME x::real. - (pi / (2::real)) <= x & x <= pi / (2::real) & sin x = y" + +lemma asn: "ALL y::real. + asn y = + (SOME x::real. - (pi / (2::real)) <= x & x <= pi / (2::real) & sin x = y)" by (import transc asn) constdefs acs :: "real => real" - "acs == %y. SOME x. 0 <= x & x <= pi & cos x = y" - -lemma acs: "ALL y. acs y = (SOME x. 0 <= x & x <= pi & cos x = y)" + "acs == %y::real. SOME x::real. (0::real) <= x & x <= pi & cos x = y" + +lemma acs: "ALL y::real. acs y = (SOME x::real. (0::real) <= x & x <= pi & cos x = y)" by (import transc acs) constdefs atn :: "real => real" - "atn == %y. SOME x. - (pi / 2) < x & x < pi / 2 & tan x = y" - -lemma atn: "ALL y. atn y = (SOME x. - (pi / 2) < x & x < pi / 2 & tan x = y)" + "atn == +%y::real. + SOME x::real. - (pi / (2::real)) < x & x < pi / (2::real) & tan x = y" + +lemma atn: "ALL y::real. + atn y = + (SOME x::real. - (pi / (2::real)) < x & x < pi / (2::real) & tan x = y)" by (import transc atn) lemma ASN: "(All::(real => bool) => bool) @@ -5369,13 +5600,14 @@ ((acs::real => real) ((cos::real => real) x)) x))" by (import transc COS_ACS) -lemma ATN: "ALL y. - (pi / 2) < atn y & atn y < pi / 2 & tan (atn y) = y" +lemma ATN: "ALL y::real. + - (pi / (2::real)) < atn y & atn y < pi / (2::real) & tan (atn y) = y" by (import transc ATN) -lemma ATN_TAN: "ALL x. tan (atn x) = x" +lemma ATN_TAN: "ALL x::real. tan (atn x) = x" by (import transc ATN_TAN) -lemma ATN_BOUNDS: "ALL x. - (pi / 2) < atn x & atn x < pi / 2" +lemma ATN_BOUNDS: "ALL x::real. - (pi / (2::real)) < atn x & atn x < pi / (2::real)" by (import transc ATN_BOUNDS) lemma TAN_ATN: "(All::(real => bool) => bool) @@ -5471,7 +5703,7 @@ (bit.B0::bit))))))))" by (import transc COS_SIN_SQ) -lemma COS_ATN_NZ: "ALL x. cos (atn x) ~= 0" +lemma COS_ATN_NZ: "ALL x::real. cos (atn x) ~= (0::real)" by (import transc COS_ATN_NZ) lemma COS_ASN_NZ: "(All::(real => bool) => bool) @@ -5526,7 +5758,7 @@ (bit.B0::bit))))))))" by (import transc SIN_COS_SQRT) -lemma DIFF_LN: "ALL x>0. diffl ln (inverse x) x" +lemma DIFF_LN: "ALL x>0::real. diffl ln (inverse x) x" by (import transc DIFF_LN) lemma DIFF_ASN_LEMMA: "(All::(real => bool) => bool) @@ -5593,7 +5825,7 @@ x))" by (import transc DIFF_ACS) -lemma DIFF_ATN: "ALL x. diffl atn (inverse (1 + x ^ 2)) x" +lemma DIFF_ATN: "ALL x::real. diffl atn (inverse ((1::real) + x ^ 2)) x" by (import transc DIFF_ATN) constdefs @@ -5691,11 +5923,12 @@ constdefs tdiv :: "real * real => (nat => real) * (nat => real) => bool" "tdiv == -%(a, b) (D, p). division (a, b) D & (ALL n. D n <= p n & p n <= D (Suc n))" - -lemma tdiv: "ALL a b D p. +%(a::real, b::real) (D::nat => real, p::nat => real). + division (a, b) D & (ALL n::nat. D n <= p n & p n <= D (Suc n))" + +lemma tdiv: "ALL (a::real) (b::real) (D::nat => real) p::nat => real. tdiv (a, b) (D, p) = - (division (a, b) D & (ALL n. D n <= p n & p n <= D (Suc n)))" + (division (a, b) D & (ALL n::nat. D n <= p n & p n <= D (Suc n)))" by (import transc tdiv) constdefs @@ -5767,10 +6000,13 @@ constdefs rsum :: "(nat => real) * (nat => real) => (real => real) => real" - "rsum == %(D, p) f. real.sum (0, dsize D) (%n. f (p n) * (D (Suc n) - D n))" - -lemma rsum: "ALL D p f. - rsum (D, p) f = real.sum (0, dsize D) (%n. f (p n) * (D (Suc n) - D n))" + "rsum == +%(D::nat => real, p::nat => real) f::real => real. + real.sum (0::nat, dsize D) (%n::nat. f (p n) * (D (Suc n) - D n))" + +lemma rsum: "ALL (D::nat => real) (p::nat => real) f::real => real. + rsum (D, p) f = + real.sum (0::nat, dsize D) (%n::nat. f (p n) * (D (Suc n) - D n))" by (import transc rsum) constdefs @@ -6283,7 +6519,7 @@ ((op =::real => real => bool) k1 k2))))))" by (import transc DINT_UNIQ) -lemma INTEGRAL_NULL: "ALL f a. Dint (a, a) f 0" +lemma INTEGRAL_NULL: "ALL (f::real => real) a::real. Dint (a, a) f (0::real)" by (import transc INTEGRAL_NULL) lemma FTC1: "(All::((real => real) => bool) => bool) @@ -6555,11 +6791,11 @@ ((op ^::real => nat => real) x n)))))))))" by (import transc MCLAURIN_EXP_LT) -lemma MCLAURIN_EXP_LE: "ALL x n. - EX xa. +lemma MCLAURIN_EXP_LE: "ALL (x::real) n::nat. + EX xa::real. abs xa <= abs x & exp x = - real.sum (0, n) (%m. x ^ m / real (FACT m)) + + real.sum (0::nat, n) (%m::nat. x ^ m / real (FACT m)) + exp xa / real (FACT n) * x ^ n" by (import transc MCLAURIN_EXP_LE) @@ -6587,14 +6823,15 @@ consts poly :: "real list => real => real" -specification (poly_primdef: poly) poly_def: "(ALL x. poly [] x = 0) & (ALL h t x. poly (h # t) x = h + x * poly t x)" +specification (poly_primdef: poly) poly_def: "(ALL x::real. poly [] x = (0::real)) & +(ALL (h::real) (t::real list) x::real. poly (h # t) x = h + x * poly t x)" by (import poly poly_def) consts poly_add :: "real list => real list => real list" -specification (poly_add_primdef: poly_add) poly_add_def: "(ALL l2. poly_add [] l2 = l2) & -(ALL h t l2. +specification (poly_add_primdef: poly_add) poly_add_def: "(ALL l2::real list. poly_add [] l2 = l2) & +(ALL (h::real) (t::real list) l2::real list. poly_add (h # t) l2 = (if l2 = [] then h # t else (h + hd l2) # poly_add t (tl l2)))" by (import poly poly_add_def) @@ -6602,94 +6839,104 @@ consts "##" :: "real => real list => real list" ("##") -specification ("##") poly_cmul_def: "(ALL c. ## c [] = []) & (ALL c h t. ## c (h # t) = c * h # ## c t)" +specification ("##") poly_cmul_def: "(ALL c::real. ## c [] = []) & +(ALL (c::real) (h::real) t::real list. ## c (h # t) = c * h # ## c t)" by (import poly poly_cmul_def) consts poly_neg :: "real list => real list" defs - poly_neg_primdef: "poly_neg == ## (- 1)" - -lemma poly_neg_def: "poly_neg = ## (- 1)" + poly_neg_primdef: "poly_neg == ## (- (1::real))" + +lemma poly_neg_def: "poly_neg = ## (- (1::real))" by (import poly poly_neg_def) consts poly_mul :: "real list => real list => real list" -specification (poly_mul_primdef: poly_mul) poly_mul_def: "(ALL l2. poly_mul [] l2 = []) & -(ALL h t l2. +specification (poly_mul_primdef: poly_mul) poly_mul_def: "(ALL l2::real list. poly_mul [] l2 = []) & +(ALL (h::real) (t::real list) l2::real list. poly_mul (h # t) l2 = - (if t = [] then ## h l2 else poly_add (## h l2) (0 # poly_mul t l2)))" + (if t = [] then ## h l2 + else poly_add (## h l2) ((0::real) # poly_mul t l2)))" by (import poly poly_mul_def) consts poly_exp :: "real list => nat => real list" -specification (poly_exp_primdef: poly_exp) poly_exp_def: "(ALL p. poly_exp p 0 = [1]) & -(ALL p n. poly_exp p (Suc n) = poly_mul p (poly_exp p n))" +specification (poly_exp_primdef: poly_exp) poly_exp_def: "(ALL p::real list. poly_exp p (0::nat) = [1::real]) & +(ALL (p::real list) n::nat. poly_exp p (Suc n) = poly_mul p (poly_exp p n))" by (import poly poly_exp_def) consts poly_diff_aux :: "nat => real list => real list" -specification (poly_diff_aux_primdef: poly_diff_aux) poly_diff_aux_def: "(ALL n. poly_diff_aux n [] = []) & -(ALL n h t. poly_diff_aux n (h # t) = real n * h # poly_diff_aux (Suc n) t)" +specification (poly_diff_aux_primdef: poly_diff_aux) poly_diff_aux_def: "(ALL n::nat. poly_diff_aux n [] = []) & +(ALL (n::nat) (h::real) t::real list. + poly_diff_aux n (h # t) = real n * h # poly_diff_aux (Suc n) t)" by (import poly poly_diff_aux_def) constdefs diff :: "real list => real list" - "diff == %l. if l = [] then [] else poly_diff_aux 1 (tl l)" - -lemma poly_diff_def: "ALL l. diff l = (if l = [] then [] else poly_diff_aux 1 (tl l))" + "diff == %l::real list. if l = [] then [] else poly_diff_aux (1::nat) (tl l)" + +lemma poly_diff_def: "ALL l::real list. + diff l = (if l = [] then [] else poly_diff_aux (1::nat) (tl l))" by (import poly poly_diff_def) -lemma POLY_ADD_CLAUSES: "poly_add [] p2 = p2 & -poly_add p1 [] = p1 & -poly_add (h1 # t1) (h2 # t2) = (h1 + h2) # poly_add t1 t2" +lemma POLY_ADD_CLAUSES: "poly_add [] (p2::real list) = p2 & +poly_add (p1::real list) [] = p1 & +poly_add ((h1::real) # (t1::real list)) ((h2::real) # (t2::real list)) = +(h1 + h2) # poly_add t1 t2" by (import poly POLY_ADD_CLAUSES) -lemma POLY_CMUL_CLAUSES: "## c [] = [] & ## c (h # t) = c * h # ## c t" +lemma POLY_CMUL_CLAUSES: "## (c::real) [] = [] & ## c ((h::real) # (t::real list)) = c * h # ## c t" by (import poly POLY_CMUL_CLAUSES) -lemma POLY_NEG_CLAUSES: "poly_neg [] = [] & poly_neg (h # t) = - h # poly_neg t" +lemma POLY_NEG_CLAUSES: "poly_neg [] = [] & poly_neg ((h::real) # (t::real list)) = - h # poly_neg t" by (import poly POLY_NEG_CLAUSES) -lemma POLY_MUL_CLAUSES: "poly_mul [] p2 = [] & -poly_mul [h1] p2 = ## h1 p2 & -poly_mul (h1 # k1 # t1) p2 = poly_add (## h1 p2) (0 # poly_mul (k1 # t1) p2)" +lemma POLY_MUL_CLAUSES: "poly_mul [] (p2::real list) = [] & +poly_mul [h1::real] p2 = ## h1 p2 & +poly_mul (h1 # (k1::real) # (t1::real list)) p2 = +poly_add (## h1 p2) ((0::real) # poly_mul (k1 # t1) p2)" by (import poly POLY_MUL_CLAUSES) -lemma POLY_DIFF_CLAUSES: "diff [] = [] & diff [c] = [] & diff (h # t) = poly_diff_aux 1 t" +lemma POLY_DIFF_CLAUSES: "diff [] = [] & +diff [c::real] = [] & +diff ((h::real) # (t::real list)) = poly_diff_aux (1::nat) t" by (import poly POLY_DIFF_CLAUSES) -lemma POLY_ADD: "ALL t p2 x. poly (poly_add t p2) x = poly t x + poly p2 x" +lemma POLY_ADD: "ALL (t::real list) (p2::real list) x::real. + poly (poly_add t p2) x = poly t x + poly p2 x" by (import poly POLY_ADD) -lemma POLY_CMUL: "ALL t c x. poly (## c t) x = c * poly t x" +lemma POLY_CMUL: "ALL (t::real list) (c::real) x::real. poly (## c t) x = c * poly t x" by (import poly POLY_CMUL) -lemma POLY_NEG: "ALL x xa. poly (poly_neg x) xa = - poly x xa" +lemma POLY_NEG: "ALL (x::real list) xa::real. poly (poly_neg x) xa = - poly x xa" by (import poly POLY_NEG) -lemma POLY_MUL: "ALL x t p2. poly (poly_mul t p2) x = poly t x * poly p2 x" +lemma POLY_MUL: "ALL (x::real) (t::real list) p2::real list. + poly (poly_mul t p2) x = poly t x * poly p2 x" by (import poly POLY_MUL) -lemma POLY_EXP: "ALL p n x. poly (poly_exp p n) x = poly p x ^ n" +lemma POLY_EXP: "ALL (p::real list) (n::nat) x::real. poly (poly_exp p n) x = poly p x ^ n" by (import poly POLY_EXP) -lemma POLY_DIFF_LEMMA: "ALL t n x. - diffl (%x. x ^ Suc n * poly t x) +lemma POLY_DIFF_LEMMA: "ALL (t::real list) (n::nat) x::real. + diffl (%x::real. x ^ Suc n * poly t x) (x ^ n * poly (poly_diff_aux (Suc n) t) x) x" by (import poly POLY_DIFF_LEMMA) -lemma POLY_DIFF: "ALL t x. diffl (poly t) (poly (diff t) x) x" +lemma POLY_DIFF: "ALL (t::real list) x::real. diffl (poly t) (poly (diff t) x) x" by (import poly POLY_DIFF) -lemma POLY_DIFFERENTIABLE: "ALL l. All (differentiable (poly l))" +lemma POLY_DIFFERENTIABLE: "ALL l::real list. All (differentiable (poly l))" by (import poly POLY_DIFFERENTIABLE) -lemma POLY_CONT: "ALL l. All (contl (poly l))" +lemma POLY_CONT: "ALL l::real list. All (contl (poly l))" by (import poly POLY_CONT) lemma POLY_IVT_POS: "(All::(real list => bool) => bool) @@ -6766,69 +7013,77 @@ ((diff::real list => real list) p) x)))))))))" by (import poly POLY_MVT) -lemma POLY_ADD_RZERO: "ALL x. poly (poly_add x []) = poly x" +lemma POLY_ADD_RZERO: "ALL x::real list. poly (poly_add x []) = poly x" by (import poly POLY_ADD_RZERO) -lemma POLY_MUL_ASSOC: "ALL x xa xb. +lemma POLY_MUL_ASSOC: "ALL (x::real list) (xa::real list) xb::real list. poly (poly_mul x (poly_mul xa xb)) = poly (poly_mul (poly_mul x xa) xb)" by (import poly POLY_MUL_ASSOC) -lemma POLY_EXP_ADD: "ALL x xa xb. +lemma POLY_EXP_ADD: "ALL (x::nat) (xa::nat) xb::real list. poly (poly_exp xb (xa + x)) = poly (poly_mul (poly_exp xb xa) (poly_exp xb x))" by (import poly POLY_EXP_ADD) -lemma POLY_DIFF_AUX_ADD: "ALL t p2 n. +lemma POLY_DIFF_AUX_ADD: "ALL (t::real list) (p2::real list) n::nat. poly (poly_diff_aux n (poly_add t p2)) = poly (poly_add (poly_diff_aux n t) (poly_diff_aux n p2))" by (import poly POLY_DIFF_AUX_ADD) -lemma POLY_DIFF_AUX_CMUL: "ALL t c n. poly (poly_diff_aux n (## c t)) = poly (## c (poly_diff_aux n t))" +lemma POLY_DIFF_AUX_CMUL: "ALL (t::real list) (c::real) n::nat. + poly (poly_diff_aux n (## c t)) = poly (## c (poly_diff_aux n t))" by (import poly POLY_DIFF_AUX_CMUL) -lemma POLY_DIFF_AUX_NEG: "ALL x xa. +lemma POLY_DIFF_AUX_NEG: "ALL (x::real list) xa::nat. poly (poly_diff_aux xa (poly_neg x)) = poly (poly_neg (poly_diff_aux xa x))" by (import poly POLY_DIFF_AUX_NEG) -lemma POLY_DIFF_AUX_MUL_LEMMA: "ALL t n. +lemma POLY_DIFF_AUX_MUL_LEMMA: "ALL (t::real list) n::nat. poly (poly_diff_aux (Suc n) t) = poly (poly_add (poly_diff_aux n t) t)" by (import poly POLY_DIFF_AUX_MUL_LEMMA) -lemma POLY_DIFF_ADD: "ALL t p2. poly (diff (poly_add t p2)) = poly (poly_add (diff t) (diff p2))" +lemma POLY_DIFF_ADD: "ALL (t::real list) p2::real list. + poly (diff (poly_add t p2)) = poly (poly_add (diff t) (diff p2))" by (import poly POLY_DIFF_ADD) -lemma POLY_DIFF_CMUL: "ALL t c. poly (diff (## c t)) = poly (## c (diff t))" +lemma POLY_DIFF_CMUL: "ALL (t::real list) c::real. poly (diff (## c t)) = poly (## c (diff t))" by (import poly POLY_DIFF_CMUL) -lemma POLY_DIFF_NEG: "ALL x. poly (diff (poly_neg x)) = poly (poly_neg (diff x))" +lemma POLY_DIFF_NEG: "ALL x::real list. poly (diff (poly_neg x)) = poly (poly_neg (diff x))" by (import poly POLY_DIFF_NEG) -lemma POLY_DIFF_MUL_LEMMA: "ALL x xa. poly (diff (xa # x)) = poly (poly_add (0 # diff x) x)" +lemma POLY_DIFF_MUL_LEMMA: "ALL (x::real list) xa::real. + poly (diff (xa # x)) = poly (poly_add ((0::real) # diff x) x)" by (import poly POLY_DIFF_MUL_LEMMA) -lemma POLY_DIFF_MUL: "ALL t p2. +lemma POLY_DIFF_MUL: "ALL (t::real list) p2::real list. poly (diff (poly_mul t p2)) = poly (poly_add (poly_mul t (diff p2)) (poly_mul (diff t) p2))" by (import poly POLY_DIFF_MUL) -lemma POLY_DIFF_EXP: "ALL p n. +lemma POLY_DIFF_EXP: "ALL (p::real list) n::nat. poly (diff (poly_exp p (Suc n))) = poly (poly_mul (## (real (Suc n)) (poly_exp p n)) (diff p))" by (import poly POLY_DIFF_EXP) -lemma POLY_DIFF_EXP_PRIME: "ALL n a. - poly (diff (poly_exp [- a, 1] (Suc n))) = - poly (## (real (Suc n)) (poly_exp [- a, 1] n))" +lemma POLY_DIFF_EXP_PRIME: "ALL (n::nat) a::real. + poly (diff (poly_exp [- a, 1::real] (Suc n))) = + poly (## (real (Suc n)) (poly_exp [- a, 1::real] n))" by (import poly POLY_DIFF_EXP_PRIME) -lemma POLY_LINEAR_REM: "ALL t h. EX q r. h # t = poly_add [r] (poly_mul [- a, 1] q)" +lemma POLY_LINEAR_REM: "ALL (t::real list) h::real. + EX (q::real list) r::real. + h # t = poly_add [r] (poly_mul [- (a::real), 1::real] q)" by (import poly POLY_LINEAR_REM) -lemma POLY_LINEAR_DIVIDES: "ALL a t. (poly t a = 0) = (t = [] | (EX q. t = poly_mul [- a, 1] q))" +lemma POLY_LINEAR_DIVIDES: "ALL (a::real) t::real list. + (poly t a = (0::real)) = + (t = [] | (EX q::real list. t = poly_mul [- a, 1::real] q))" by (import poly POLY_LINEAR_DIVIDES) -lemma POLY_LENGTH_MUL: "ALL x. length (poly_mul [- a, 1] x) = Suc (length x)" +lemma POLY_LENGTH_MUL: "ALL x::real list. + length (poly_mul [- (a::real), 1::real] x) = Suc (length x)" by (import poly POLY_LENGTH_MUL) lemma POLY_ROOTS_INDEX_LEMMA: "(All::(nat => bool) => bool) @@ -6968,22 +7223,23 @@ ((poly::real list => real => real) ([]::real list))))))" by (import poly POLY_ENTIRE_LEMMA) -lemma POLY_ENTIRE: "ALL p q. +lemma POLY_ENTIRE: "ALL (p::real list) q::real list. (poly (poly_mul p q) = poly []) = (poly p = poly [] | poly q = poly [])" by (import poly POLY_ENTIRE) -lemma POLY_MUL_LCANCEL: "ALL x xa xb. +lemma POLY_MUL_LCANCEL: "ALL (x::real list) (xa::real list) xb::real list. (poly (poly_mul x xa) = poly (poly_mul x xb)) = (poly x = poly [] | poly xa = poly xb)" by (import poly POLY_MUL_LCANCEL) -lemma POLY_EXP_EQ_0: "ALL p n. (poly (poly_exp p n) = poly []) = (poly p = poly [] & n ~= 0)" +lemma POLY_EXP_EQ_0: "ALL (p::real list) n::nat. + (poly (poly_exp p n) = poly []) = (poly p = poly [] & n ~= (0::nat))" by (import poly POLY_EXP_EQ_0) -lemma POLY_PRIME_EQ_0: "ALL a. poly [a, 1] ~= poly []" +lemma POLY_PRIME_EQ_0: "ALL a::real. poly [a, 1::real] ~= poly []" by (import poly POLY_PRIME_EQ_0) -lemma POLY_EXP_PRIME_EQ_0: "ALL a n. poly (poly_exp [a, 1] n) ~= poly []" +lemma POLY_EXP_PRIME_EQ_0: "ALL (a::real) n::nat. poly (poly_exp [a, 1::real] n) ~= poly []" by (import poly POLY_EXP_PRIME_EQ_0) lemma POLY_ZERO_LEMMA: "(All::(real => bool) => bool) @@ -7002,11 +7258,12 @@ ((poly::real list => real => real) ([]::real list))))))" by (import poly POLY_ZERO_LEMMA) -lemma POLY_ZERO: "ALL t. (poly t = poly []) = list_all (%c. c = 0) t" +lemma POLY_ZERO: "ALL t::real list. (poly t = poly []) = list_all (%c::real. c = (0::real)) t" by (import poly POLY_ZERO) -lemma POLY_DIFF_AUX_ISZERO: "ALL t n. - list_all (%c. c = 0) (poly_diff_aux (Suc n) t) = list_all (%c. c = 0) t" +lemma POLY_DIFF_AUX_ISZERO: "ALL (t::real list) n::nat. + list_all (%c::real. c = (0::real)) (poly_diff_aux (Suc n) t) = + list_all (%c::real. c = (0::real)) t" by (import poly POLY_DIFF_AUX_ISZERO) lemma POLY_DIFF_ISZERO: "(All::(real list => bool) => bool) @@ -7054,17 +7311,20 @@ constdefs poly_divides :: "real list => real list => bool" - "poly_divides == %p1 p2. EX q. poly p2 = poly (poly_mul p1 q)" - -lemma poly_divides: "ALL p1 p2. poly_divides p1 p2 = (EX q. poly p2 = poly (poly_mul p1 q))" + "poly_divides == +%(p1::real list) p2::real list. + EX q::real list. poly p2 = poly (poly_mul p1 q)" + +lemma poly_divides: "ALL (p1::real list) p2::real list. + poly_divides p1 p2 = (EX q::real list. poly p2 = poly (poly_mul p1 q))" by (import poly poly_divides) -lemma POLY_PRIMES: "ALL a p q. - poly_divides [a, 1] (poly_mul p q) = - (poly_divides [a, 1] p | poly_divides [a, 1] q)" +lemma POLY_PRIMES: "ALL (a::real) (p::real list) q::real list. + poly_divides [a, 1::real] (poly_mul p q) = + (poly_divides [a, 1::real] p | poly_divides [a, 1::real] q)" by (import poly POLY_PRIMES) -lemma POLY_DIVIDES_REFL: "ALL p. poly_divides p p" +lemma POLY_DIVIDES_REFL: "ALL p::real list. poly_divides p p" by (import poly POLY_DIVIDES_REFL) lemma POLY_DIVIDES_TRANS: "(All::(real list => bool) => bool) @@ -7234,20 +7494,21 @@ constdefs poly_order :: "real => real list => nat" "poly_order == -%a p. SOME n. - poly_divides (poly_exp [- a, 1] n) p & - ~ poly_divides (poly_exp [- a, 1] (Suc n)) p" - -lemma poly_order: "ALL a p. +%(a::real) p::real list. + SOME n::nat. + poly_divides (poly_exp [- a, 1::real] n) p & + ~ poly_divides (poly_exp [- a, 1::real] (Suc n)) p" + +lemma poly_order: "ALL (a::real) p::real list. poly_order a p = - (SOME n. - poly_divides (poly_exp [- a, 1] n) p & - ~ poly_divides (poly_exp [- a, 1] (Suc n)) p)" + (SOME n::nat. + poly_divides (poly_exp [- a, 1::real] n) p & + ~ poly_divides (poly_exp [- a, 1::real] (Suc n)) p)" by (import poly poly_order) -lemma ORDER: "ALL p a n. - (poly_divides (poly_exp [- a, 1] n) p & - ~ poly_divides (poly_exp [- a, 1] (Suc n)) p) = +lemma ORDER: "ALL (p::real list) (a::real) n::nat. + (poly_divides (poly_exp [- a, 1::real] n) p & + ~ poly_divides (poly_exp [- a, 1::real] (Suc n)) p) = (n = poly_order a p & poly p ~= poly [])" by (import poly ORDER) @@ -7330,11 +7591,12 @@ ((poly_order::real => real list => nat) a q)))))" by (import poly ORDER_POLY) -lemma ORDER_ROOT: "ALL p a. (poly p a = 0) = (poly p = poly [] | poly_order a p ~= 0)" +lemma ORDER_ROOT: "ALL (p::real list) a::real. + (poly p a = (0::real)) = (poly p = poly [] | poly_order a p ~= (0::nat))" by (import poly ORDER_ROOT) -lemma ORDER_DIVIDES: "ALL p a n. - poly_divides (poly_exp [- a, 1] n) p = +lemma ORDER_DIVIDES: "ALL (p::real list) (a::real) n::nat. + poly_divides (poly_exp [- a, 1::real] n) p = (poly p = poly [] | n <= poly_order a p)" by (import poly ORDER_DIVIDES) @@ -7465,14 +7727,19 @@ constdefs rsquarefree :: "real list => bool" "rsquarefree == -%p. poly p ~= poly [] & (ALL a. poly_order a p = 0 | poly_order a p = 1)" - -lemma rsquarefree: "ALL p. +%p::real list. + poly p ~= poly [] & + (ALL a::real. poly_order a p = (0::nat) | poly_order a p = (1::nat))" + +lemma rsquarefree: "ALL p::real list. rsquarefree p = - (poly p ~= poly [] & (ALL a. poly_order a p = 0 | poly_order a p = 1))" + (poly p ~= poly [] & + (ALL a::real. poly_order a p = (0::nat) | poly_order a p = (1::nat)))" by (import poly rsquarefree) -lemma RSQUAREFREE_ROOTS: "ALL p. rsquarefree p = (ALL a. ~ (poly p a = 0 & poly (diff p) a = 0))" +lemma RSQUAREFREE_ROOTS: "ALL p::real list. + rsquarefree p = + (ALL a::real. ~ (poly p a = (0::real) & poly (diff p) a = (0::real)))" by (import poly RSQUAREFREE_ROOTS) lemma RSQUAREFREE_DECOMP: "(All::(real list => bool) => bool) @@ -7558,20 +7825,20 @@ normalize :: "real list => real list" specification (normalize) normalize: "normalize [] = [] & -(ALL h t. +(ALL (h::real) t::real list. normalize (h # t) = - (if normalize t = [] then if h = 0 then [] else [h] + (if normalize t = [] then if h = (0::real) then [] else [h] else h # normalize t))" by (import poly normalize) -lemma POLY_NORMALIZE: "ALL t. poly (normalize t) = poly t" +lemma POLY_NORMALIZE: "ALL t::real list. poly (normalize t) = poly t" by (import poly POLY_NORMALIZE) constdefs degree :: "real list => nat" - "degree == %p. PRE (length (normalize p))" - -lemma degree: "ALL p. degree p = PRE (length (normalize p))" + "degree == %p::real list. PRE (length (normalize p))" + +lemma degree: "ALL p::real list. degree p = PRE (length (normalize p))" by (import poly degree) lemma DEGREE_ZERO: "(All::(real list => bool) => bool) diff -r 7e417a7cbf9f -r bd59bfd4bf37 src/HOL/Import/HOL/HOL4Vec.thy --- a/src/HOL/Import/HOL/HOL4Vec.thy Mon Sep 26 02:06:44 2005 +0200 +++ b/src/HOL/Import/HOL/HOL4Vec.thy Mon Sep 26 02:27:14 2005 +0200 @@ -4,113 +4,137 @@ ;setup_theory res_quan -lemma RES_FORALL_CONJ_DIST: "ALL P Q R. RES_FORALL P (%i. Q i & R i) = (RES_FORALL P Q & RES_FORALL P R)" +lemma RES_FORALL_CONJ_DIST: "ALL (P::'a::type => bool) (Q::'a::type => bool) R::'a::type => bool. + RES_FORALL P (%i::'a::type. Q i & R i) = + (RES_FORALL P Q & RES_FORALL P R)" by (import res_quan RES_FORALL_CONJ_DIST) -lemma RES_FORALL_DISJ_DIST: "ALL P Q R. RES_FORALL (%j. P j | Q j) R = (RES_FORALL P R & RES_FORALL Q R)" +lemma RES_FORALL_DISJ_DIST: "ALL (P::'a::type => bool) (Q::'a::type => bool) R::'a::type => bool. + RES_FORALL (%j::'a::type. P j | Q j) R = + (RES_FORALL P R & RES_FORALL Q R)" by (import res_quan RES_FORALL_DISJ_DIST) -lemma RES_FORALL_UNIQUE: "ALL x xa. RES_FORALL (op = xa) x = x xa" +lemma RES_FORALL_UNIQUE: "ALL (x::'a::type => bool) xa::'a::type. RES_FORALL (op = xa) x = x xa" by (import res_quan RES_FORALL_UNIQUE) -lemma RES_FORALL_FORALL: "ALL (P::'a => bool) (R::'a => 'b => bool) x::'b. - (ALL x::'b. RES_FORALL P (%i::'a. R i x)) = - RES_FORALL P (%i::'a. All (R i))" +lemma RES_FORALL_FORALL: "ALL (P::'a::type => bool) (R::'a::type => 'b::type => bool) x::'b::type. + (ALL x::'b::type. RES_FORALL P (%i::'a::type. R i x)) = + RES_FORALL P (%i::'a::type. All (R i))" by (import res_quan RES_FORALL_FORALL) -lemma RES_FORALL_REORDER: "ALL P Q R. - RES_FORALL P (%i. RES_FORALL Q (R i)) = - RES_FORALL Q (%j. RES_FORALL P (%i. R i j))" +lemma RES_FORALL_REORDER: "ALL (P::'a::type => bool) (Q::'b::type => bool) + R::'a::type => 'b::type => bool. + RES_FORALL P (%i::'a::type. RES_FORALL Q (R i)) = + RES_FORALL Q (%j::'b::type. RES_FORALL P (%i::'a::type. R i j))" by (import res_quan RES_FORALL_REORDER) lemma RES_FORALL_EMPTY: "All (RES_FORALL EMPTY)" by (import res_quan RES_FORALL_EMPTY) -lemma RES_FORALL_UNIV: "ALL p. RES_FORALL pred_set.UNIV p = All p" +lemma RES_FORALL_UNIV: "ALL p::'a::type => bool. RES_FORALL pred_set.UNIV p = All p" by (import res_quan RES_FORALL_UNIV) -lemma RES_FORALL_NULL: "ALL p m. RES_FORALL p (%x. m) = (p = EMPTY | m)" +lemma RES_FORALL_NULL: "ALL (p::'a::type => bool) m::bool. + RES_FORALL p (%x::'a::type. m) = (p = EMPTY | m)" by (import res_quan RES_FORALL_NULL) -lemma RES_EXISTS_DISJ_DIST: "ALL P Q R. RES_EXISTS P (%i. Q i | R i) = (RES_EXISTS P Q | RES_EXISTS P R)" +lemma RES_EXISTS_DISJ_DIST: "ALL (P::'a::type => bool) (Q::'a::type => bool) R::'a::type => bool. + RES_EXISTS P (%i::'a::type. Q i | R i) = + (RES_EXISTS P Q | RES_EXISTS P R)" by (import res_quan RES_EXISTS_DISJ_DIST) -lemma RES_DISJ_EXISTS_DIST: "ALL P Q R. RES_EXISTS (%i. P i | Q i) R = (RES_EXISTS P R | RES_EXISTS Q R)" +lemma RES_DISJ_EXISTS_DIST: "ALL (P::'a::type => bool) (Q::'a::type => bool) R::'a::type => bool. + RES_EXISTS (%i::'a::type. P i | Q i) R = + (RES_EXISTS P R | RES_EXISTS Q R)" by (import res_quan RES_DISJ_EXISTS_DIST) -lemma RES_EXISTS_EQUAL: "ALL x xa. RES_EXISTS (op = xa) x = x xa" +lemma RES_EXISTS_EQUAL: "ALL (x::'a::type => bool) xa::'a::type. RES_EXISTS (op = xa) x = x xa" by (import res_quan RES_EXISTS_EQUAL) -lemma RES_EXISTS_REORDER: "ALL P Q R. - RES_EXISTS P (%i. RES_EXISTS Q (R i)) = - RES_EXISTS Q (%j. RES_EXISTS P (%i. R i j))" +lemma RES_EXISTS_REORDER: "ALL (P::'a::type => bool) (Q::'b::type => bool) + R::'a::type => 'b::type => bool. + RES_EXISTS P (%i::'a::type. RES_EXISTS Q (R i)) = + RES_EXISTS Q (%j::'b::type. RES_EXISTS P (%i::'a::type. R i j))" by (import res_quan RES_EXISTS_REORDER) -lemma RES_EXISTS_EMPTY: "ALL p. ~ RES_EXISTS EMPTY p" +lemma RES_EXISTS_EMPTY: "ALL p::'a::type => bool. ~ RES_EXISTS EMPTY p" by (import res_quan RES_EXISTS_EMPTY) -lemma RES_EXISTS_UNIV: "ALL p. RES_EXISTS pred_set.UNIV p = Ex p" +lemma RES_EXISTS_UNIV: "ALL p::'a::type => bool. RES_EXISTS pred_set.UNIV p = Ex p" by (import res_quan RES_EXISTS_UNIV) -lemma RES_EXISTS_NULL: "ALL p m. RES_EXISTS p (%x. m) = (p ~= EMPTY & m)" +lemma RES_EXISTS_NULL: "ALL (p::'a::type => bool) m::bool. + RES_EXISTS p (%x::'a::type. m) = (p ~= EMPTY & m)" by (import res_quan RES_EXISTS_NULL) -lemma RES_EXISTS_ALT: "ALL p m. RES_EXISTS p m = (IN (RES_SELECT p m) p & m (RES_SELECT p m))" +lemma RES_EXISTS_ALT: "ALL (p::'a::type => bool) m::'a::type => bool. + RES_EXISTS p m = (IN (RES_SELECT p m) p & m (RES_SELECT p m))" by (import res_quan RES_EXISTS_ALT) -lemma RES_EXISTS_UNIQUE_EMPTY: "ALL p. ~ RES_EXISTS_UNIQUE EMPTY p" +lemma RES_EXISTS_UNIQUE_EMPTY: "ALL p::'a::type => bool. ~ RES_EXISTS_UNIQUE EMPTY p" by (import res_quan RES_EXISTS_UNIQUE_EMPTY) -lemma RES_EXISTS_UNIQUE_UNIV: "ALL p. RES_EXISTS_UNIQUE pred_set.UNIV p = Ex1 p" +lemma RES_EXISTS_UNIQUE_UNIV: "ALL p::'a::type => bool. RES_EXISTS_UNIQUE pred_set.UNIV p = Ex1 p" by (import res_quan RES_EXISTS_UNIQUE_UNIV) -lemma RES_EXISTS_UNIQUE_NULL: "ALL p m. RES_EXISTS_UNIQUE p (%x. m) = ((EX x. p = INSERT x EMPTY) & m)" +lemma RES_EXISTS_UNIQUE_NULL: "ALL (p::'a::type => bool) m::bool. + RES_EXISTS_UNIQUE p (%x::'a::type. m) = + ((EX x::'a::type. p = INSERT x EMPTY) & m)" by (import res_quan RES_EXISTS_UNIQUE_NULL) -lemma RES_EXISTS_UNIQUE_ALT: "ALL p m. +lemma RES_EXISTS_UNIQUE_ALT: "ALL (p::'a::type => bool) m::'a::type => bool. RES_EXISTS_UNIQUE p m = - RES_EXISTS p (%x. m x & RES_FORALL p (%y. m y --> y = x))" + RES_EXISTS p + (%x::'a::type. m x & RES_FORALL p (%y::'a::type. m y --> y = x))" by (import res_quan RES_EXISTS_UNIQUE_ALT) -lemma RES_SELECT_EMPTY: "ALL p. RES_SELECT EMPTY p = (SOME x. False)" +lemma RES_SELECT_EMPTY: "ALL p::'a::type => bool. RES_SELECT EMPTY p = (SOME x::'a::type. False)" by (import res_quan RES_SELECT_EMPTY) -lemma RES_SELECT_UNIV: "ALL p. RES_SELECT pred_set.UNIV p = Eps p" +lemma RES_SELECT_UNIV: "ALL p::'a::type => bool. RES_SELECT pred_set.UNIV p = Eps p" by (import res_quan RES_SELECT_UNIV) -lemma RES_ABSTRACT: "ALL p m x. IN x p --> RES_ABSTRACT p m x = m x" +lemma RES_ABSTRACT: "ALL (p::'a::type => bool) (m::'a::type => 'b::type) x::'a::type. + IN x p --> RES_ABSTRACT p m x = m x" by (import res_quan RES_ABSTRACT) -lemma RES_ABSTRACT_EQUAL: "ALL p m1 m2. - (ALL x. IN x p --> m1 x = m2 x) --> RES_ABSTRACT p m1 = RES_ABSTRACT p m2" +lemma RES_ABSTRACT_EQUAL: "ALL (p::'a::type => bool) (m1::'a::type => 'b::type) + m2::'a::type => 'b::type. + (ALL x::'a::type. IN x p --> m1 x = m2 x) --> + RES_ABSTRACT p m1 = RES_ABSTRACT p m2" by (import res_quan RES_ABSTRACT_EQUAL) -lemma RES_ABSTRACT_IDEMPOT: "ALL p m. RES_ABSTRACT p (RES_ABSTRACT p m) = RES_ABSTRACT p m" +lemma RES_ABSTRACT_IDEMPOT: "ALL (p::'a::type => bool) m::'a::type => 'b::type. + RES_ABSTRACT p (RES_ABSTRACT p m) = RES_ABSTRACT p m" by (import res_quan RES_ABSTRACT_IDEMPOT) -lemma RES_ABSTRACT_EQUAL_EQ: "ALL p m1 m2. - (RES_ABSTRACT p m1 = RES_ABSTRACT p m2) = (ALL x. IN x p --> m1 x = m2 x)" +lemma RES_ABSTRACT_EQUAL_EQ: "ALL (p::'a::type => bool) (m1::'a::type => 'b::type) + m2::'a::type => 'b::type. + (RES_ABSTRACT p m1 = RES_ABSTRACT p m2) = + (ALL x::'a::type. IN x p --> m1 x = m2 x)" by (import res_quan RES_ABSTRACT_EQUAL_EQ) ;end_setup ;setup_theory word_base -typedef (open) ('a) word = "(Collect::('a list recspace => bool) => 'a list recspace set) - (%x::'a list recspace. - (All::(('a list recspace => bool) => bool) => bool) - (%word::'a list recspace => bool. +typedef (open) ('a) word = "(Collect::('a::type list recspace => bool) => 'a::type list recspace set) + (%x::'a::type list recspace. + (All::(('a::type list recspace => bool) => bool) => bool) + (%word::'a::type list recspace => bool. (op -->::bool => bool => bool) - ((All::('a list recspace => bool) => bool) - (%a0::'a list recspace. + ((All::('a::type list recspace => bool) => bool) + (%a0::'a::type list recspace. (op -->::bool => bool => bool) - ((Ex::('a list => bool) => bool) - (%a::'a list. - (op =::'a list recspace => 'a list recspace => bool) + ((Ex::('a::type list => bool) => bool) + (%a::'a::type list. + (op =::'a::type list recspace + => 'a::type list recspace => bool) a0 ((CONSTR::nat -=> 'a list => (nat => 'a list recspace) => 'a list recspace) +=> 'a::type list + => (nat => 'a::type list recspace) => 'a::type list recspace) (0::nat) a - (%n::nat. BOTTOM::'a list recspace)))) + (%n::nat. BOTTOM::'a::type list recspace)))) (word a0))) (word x)))" by (rule typedef_helper,import word_base word_TY_DEF) @@ -118,449 +142,543 @@ lemmas word_TY_DEF = typedef_hol2hol4 [OF type_definition_word] consts - mk_word :: "'a list recspace => 'a word" - dest_word :: "'a word => 'a list recspace" + mk_word :: "'a::type list recspace => 'a::type word" + dest_word :: "'a::type word => 'a::type list recspace" -specification (dest_word mk_word) word_repfns: "(ALL a::'a word. mk_word (dest_word a) = a) & -(ALL r::'a list recspace. - (ALL word::'a list recspace => bool. - (ALL a0::'a list recspace. - (EX a::'a list. a0 = CONSTR (0::nat) a (%n::nat. BOTTOM)) --> +specification (dest_word mk_word) word_repfns: "(ALL a::'a::type word. mk_word (dest_word a) = a) & +(ALL r::'a::type list recspace. + (ALL word::'a::type list recspace => bool. + (ALL a0::'a::type list recspace. + (EX a::'a::type list. + a0 = CONSTR (0::nat) a (%n::nat. BOTTOM)) --> word a0) --> word r) = (dest_word (mk_word r) = r))" by (import word_base word_repfns) consts - word_base0 :: "'a list => 'a word" + word_base0 :: "'a::type list => 'a::type word" defs - word_base0_primdef: "word_base0 == %a. mk_word (CONSTR 0 a (%n. BOTTOM))" + word_base0_primdef: "word_base0 == +%a::'a::type list. mk_word (CONSTR (0::nat) a (%n::nat. BOTTOM))" -lemma word_base0_def: "word_base0 = (%a. mk_word (CONSTR 0 a (%n. BOTTOM)))" +lemma word_base0_def: "word_base0 = +(%a::'a::type list. mk_word (CONSTR (0::nat) a (%n::nat. BOTTOM)))" by (import word_base word_base0_def) constdefs - WORD :: "'a list => 'a word" + WORD :: "'a::type list => 'a::type word" "WORD == word_base0" lemma WORD: "WORD = word_base0" by (import word_base WORD) consts - word_case :: "('a list => 'b) => 'a word => 'b" + word_case :: "('a::type list => 'b::type) => 'a::type word => 'b::type" -specification (word_case_primdef: word_case) word_case_def: "ALL f a. word_case f (WORD a) = f a" +specification (word_case_primdef: word_case) word_case_def: "ALL (f::'a::type list => 'b::type) a::'a::type list. + word_case f (WORD a) = f a" by (import word_base word_case_def) consts - word_size :: "('a => nat) => 'a word => nat" + word_size :: "('a::type => nat) => 'a::type word => nat" -specification (word_size_primdef: word_size) word_size_def: "ALL f a. word_size f (WORD a) = 1 + list_size f a" +specification (word_size_primdef: word_size) word_size_def: "ALL (f::'a::type => nat) a::'a::type list. + word_size f (WORD a) = (1::nat) + list_size f a" by (import word_base word_size_def) -lemma word_11: "ALL a a'. (WORD a = WORD a') = (a = a')" +lemma word_11: "ALL (a::'a::type list) a'::'a::type list. (WORD a = WORD a') = (a = a')" by (import word_base word_11) -lemma word_case_cong: "ALL M M' f. - M = M' & (ALL a. M' = WORD a --> f a = f' a) --> +lemma word_case_cong: "ALL (M::'a::type word) (M'::'a::type word) f::'a::type list => 'b::type. + M = M' & + (ALL a::'a::type list. + M' = WORD a --> f a = (f'::'a::type list => 'b::type) a) --> word_case f M = word_case f' M'" by (import word_base word_case_cong) -lemma word_nchotomy: "ALL x. EX l. x = WORD l" +lemma word_nchotomy: "ALL x::'a::type word. EX l::'a::type list. x = WORD l" by (import word_base word_nchotomy) -lemma word_Axiom: "ALL f. EX fn. ALL a. fn (WORD a) = f a" +lemma word_Axiom: "ALL f::'a::type list => 'b::type. + EX fn::'a::type word => 'b::type. ALL a::'a::type list. fn (WORD a) = f a" by (import word_base word_Axiom) -lemma word_induction: "ALL P. (ALL a. P (WORD a)) --> All P" +lemma word_induction: "ALL P::'a::type word => bool. (ALL a::'a::type list. P (WORD a)) --> All P" by (import word_base word_induction) -lemma word_Ax: "ALL f. EX fn. ALL a. fn (WORD a) = f a" +lemma word_Ax: "ALL f::'a::type list => 'b::type. + EX fn::'a::type word => 'b::type. ALL a::'a::type list. fn (WORD a) = f a" by (import word_base word_Ax) -lemma WORD_11: "ALL x xa. (WORD x = WORD xa) = (x = xa)" +lemma WORD_11: "ALL (x::'a::type list) xa::'a::type list. (WORD x = WORD xa) = (x = xa)" by (import word_base WORD_11) -lemma word_induct: "ALL x. (ALL l. x (WORD l)) --> All x" +lemma word_induct: "ALL x::'a::type word => bool. (ALL l::'a::type list. x (WORD l)) --> All x" by (import word_base word_induct) -lemma word_cases: "ALL x. EX l. x = WORD l" +lemma word_cases: "ALL x::'a::type word. EX l::'a::type list. x = WORD l" by (import word_base word_cases) consts - WORDLEN :: "'a word => nat" + WORDLEN :: "'a::type word => nat" -specification (WORDLEN) WORDLEN_DEF: "ALL l. WORDLEN (WORD l) = length l" +specification (WORDLEN) WORDLEN_DEF: "ALL l::'a::type list. WORDLEN (WORD l) = length l" by (import word_base WORDLEN_DEF) consts - PWORDLEN :: "nat => 'a word => bool" + PWORDLEN :: "nat => 'a::type word => bool" defs - PWORDLEN_primdef: "PWORDLEN == %n. GSPEC (%w. (w, WORDLEN w = n))" + PWORDLEN_primdef: "PWORDLEN == %n::nat. GSPEC (%w::'a::type word. (w, WORDLEN w = n))" -lemma PWORDLEN_def: "ALL n. PWORDLEN n = GSPEC (%w. (w, WORDLEN w = n))" +lemma PWORDLEN_def: "ALL n::nat. PWORDLEN n = GSPEC (%w::'a::type word. (w, WORDLEN w = n))" by (import word_base PWORDLEN_def) -lemma IN_PWORDLEN: "ALL n l. IN (WORD l) (PWORDLEN n) = (length l = n)" +lemma IN_PWORDLEN: "ALL (n::nat) l::'a::type list. IN (WORD l) (PWORDLEN n) = (length l = n)" by (import word_base IN_PWORDLEN) -lemma PWORDLEN: "ALL n w. IN w (PWORDLEN n) = (WORDLEN w = n)" +lemma PWORDLEN: "ALL (n::nat) w::'a::type word. IN w (PWORDLEN n) = (WORDLEN w = n)" by (import word_base PWORDLEN) -lemma PWORDLEN0: "ALL w. IN w (PWORDLEN 0) --> w = WORD []" +lemma PWORDLEN0: "ALL w::'a::type word. IN w (PWORDLEN (0::nat)) --> w = WORD []" by (import word_base PWORDLEN0) -lemma PWORDLEN1: "ALL x. IN (WORD [x]) (PWORDLEN 1)" +lemma PWORDLEN1: "ALL x::'a::type. IN (WORD [x]) (PWORDLEN (1::nat))" by (import word_base PWORDLEN1) consts - WSEG :: "nat => nat => 'a word => 'a word" + WSEG :: "nat => nat => 'a::type word => 'a::type word" -specification (WSEG) WSEG_DEF: "ALL m k l. WSEG m k (WORD l) = WORD (LASTN m (BUTLASTN k l))" +specification (WSEG) WSEG_DEF: "ALL (m::nat) (k::nat) l::'a::type list. + WSEG m k (WORD l) = WORD (LASTN m (BUTLASTN k l))" by (import word_base WSEG_DEF) -lemma WSEG0: "ALL k w. WSEG 0 k w = WORD []" +lemma WSEG0: "ALL (k::nat) w::'a::type word. WSEG (0::nat) k w = WORD []" by (import word_base WSEG0) -lemma WSEG_PWORDLEN: "ALL n. +lemma WSEG_PWORDLEN: "ALL n::nat. RES_FORALL (PWORDLEN n) - (%w. ALL m k. m + k <= n --> IN (WSEG m k w) (PWORDLEN m))" + (%w::'a::type word. + ALL (m::nat) k::nat. m + k <= n --> IN (WSEG m k w) (PWORDLEN m))" by (import word_base WSEG_PWORDLEN) -lemma WSEG_WORDLEN: "ALL x. +lemma WSEG_WORDLEN: "ALL x::nat. RES_FORALL (PWORDLEN x) - (%xa. ALL xb xc. xb + xc <= x --> WORDLEN (WSEG xb xc xa) = xb)" + (%xa::'a::type word. + ALL (xb::nat) xc::nat. + xb + xc <= x --> WORDLEN (WSEG xb xc xa) = xb)" by (import word_base WSEG_WORDLEN) -lemma WSEG_WORD_LENGTH: "ALL n. RES_FORALL (PWORDLEN n) (%w. WSEG n 0 w = w)" +lemma WSEG_WORD_LENGTH: "ALL n::nat. + RES_FORALL (PWORDLEN n) (%w::'a::type word. WSEG n (0::nat) w = w)" by (import word_base WSEG_WORD_LENGTH) consts - bit :: "nat => 'a word => 'a" + bit :: "nat => 'a::type word => 'a::type" -specification (bit) BIT_DEF: "ALL k l. bit k (WORD l) = ELL k l" +specification (bit) BIT_DEF: "ALL (k::nat) l::'a::type list. bit k (WORD l) = ELL k l" by (import word_base BIT_DEF) -lemma BIT0: "ALL x. bit 0 (WORD [x]) = x" +lemma BIT0: "ALL x::'a::type. bit (0::nat) (WORD [x]) = x" by (import word_base BIT0) lemma WSEG_BIT: "(All::(nat => bool) => bool) (%n::nat. - (RES_FORALL::('a word => bool) => ('a word => bool) => bool) - ((PWORDLEN::nat => 'a word => bool) n) - (%w::'a word. + (RES_FORALL::('a::type word => bool) + => ('a::type word => bool) => bool) + ((PWORDLEN::nat => 'a::type word => bool) n) + (%w::'a::type word. (All::(nat => bool) => bool) (%k::nat. (op -->::bool => bool => bool) ((op <::nat => nat => bool) k n) - ((op =::'a word => 'a word => bool) - ((WSEG::nat => nat => 'a word => 'a word) (1::nat) k w) - ((WORD::'a list => 'a word) - ((op #::'a => 'a list => 'a list) - ((bit::nat => 'a word => 'a) k w) ([]::'a list)))))))" + ((op =::'a::type word => 'a::type word => bool) + ((WSEG::nat => nat => 'a::type word => 'a::type word) + (1::nat) k w) + ((WORD::'a::type list => 'a::type word) + ((op #::'a::type => 'a::type list => 'a::type list) + ((bit::nat => 'a::type word => 'a::type) k w) + ([]::'a::type list)))))))" by (import word_base WSEG_BIT) -lemma BIT_WSEG: "ALL n. +lemma BIT_WSEG: "ALL n::nat. RES_FORALL (PWORDLEN n) - (%w. ALL m k j. - m + k <= n --> j < m --> bit j (WSEG m k w) = bit (j + k) w)" + (%w::'a::type word. + ALL (m::nat) (k::nat) j::nat. + m + k <= n --> j < m --> bit j (WSEG m k w) = bit (j + k) w)" by (import word_base BIT_WSEG) consts - MSB :: "'a word => 'a" + MSB :: "'a::type word => 'a::type" -specification (MSB) MSB_DEF: "ALL l. MSB (WORD l) = hd l" +specification (MSB) MSB_DEF: "ALL l::'a::type list. MSB (WORD l) = hd l" by (import word_base MSB_DEF) -lemma MSB: "ALL n. RES_FORALL (PWORDLEN n) (%w. 0 < n --> MSB w = bit (PRE n) w)" +lemma MSB: "ALL n::nat. + RES_FORALL (PWORDLEN n) + (%w::'a::type word. (0::nat) < n --> MSB w = bit (PRE n) w)" by (import word_base MSB) consts - LSB :: "'a word => 'a" + LSB :: "'a::type word => 'a::type" -specification (LSB) LSB_DEF: "ALL l. LSB (WORD l) = last l" +specification (LSB) LSB_DEF: "ALL l::'a::type list. LSB (WORD l) = last l" by (import word_base LSB_DEF) -lemma LSB: "ALL n. RES_FORALL (PWORDLEN n) (%w. 0 < n --> LSB w = bit 0 w)" +lemma LSB: "ALL n::nat. + RES_FORALL (PWORDLEN n) + (%w::'a::type word. (0::nat) < n --> LSB w = bit (0::nat) w)" by (import word_base LSB) consts - WSPLIT :: "nat => 'a word => 'a word * 'a word" + WSPLIT :: "nat => 'a::type word => 'a::type word * 'a::type word" -specification (WSPLIT) WSPLIT_DEF: "ALL m l. WSPLIT m (WORD l) = (WORD (BUTLASTN m l), WORD (LASTN m l))" +specification (WSPLIT) WSPLIT_DEF: "ALL (m::nat) l::'a::type list. + WSPLIT m (WORD l) = (WORD (BUTLASTN m l), WORD (LASTN m l))" by (import word_base WSPLIT_DEF) consts - WCAT :: "'a word * 'a word => 'a word" + WCAT :: "'a::type word * 'a::type word => 'a::type word" -specification (WCAT) WCAT_DEF: "ALL l1 l2. WCAT (WORD l1, WORD l2) = WORD (l1 @ l2)" +specification (WCAT) WCAT_DEF: "ALL (l1::'a::type list) l2::'a::type list. + WCAT (WORD l1, WORD l2) = WORD (l1 @ l2)" by (import word_base WCAT_DEF) lemma WORD_PARTITION: "(op &::bool => bool => bool) ((All::(nat => bool) => bool) (%n::nat. - (RES_FORALL::('a word => bool) => ('a word => bool) => bool) - ((PWORDLEN::nat => 'a word => bool) n) - (%w::'a word. + (RES_FORALL::('a::type word => bool) + => ('a::type word => bool) => bool) + ((PWORDLEN::nat => 'a::type word => bool) n) + (%w::'a::type word. (All::(nat => bool) => bool) (%m::nat. (op -->::bool => bool => bool) ((op <=::nat => nat => bool) m n) - ((op =::'a word => 'a word => bool) - ((WCAT::'a word * 'a word => 'a word) - ((WSPLIT::nat => 'a word => 'a word * 'a word) m w)) + ((op =::'a::type word => 'a::type word => bool) + ((WCAT::'a::type word * 'a::type word => 'a::type word) + ((WSPLIT::nat + => 'a::type word + => 'a::type word * 'a::type word) + m w)) w))))) ((All::(nat => bool) => bool) (%n::nat. (All::(nat => bool) => bool) (%m::nat. - (RES_FORALL::('a word => bool) => ('a word => bool) => bool) - ((PWORDLEN::nat => 'a word => bool) n) - (%w1::'a word. - (RES_FORALL::('a word => bool) - => ('a word => bool) => bool) - ((PWORDLEN::nat => 'a word => bool) m) - (%w2::'a word. - (op =::'a word * 'a word => 'a word * 'a word => bool) - ((WSPLIT::nat => 'a word => 'a word * 'a word) m - ((WCAT::'a word * 'a word => 'a word) - ((Pair::'a word => 'a word => 'a word * 'a word) - w1 w2))) - ((Pair::'a word => 'a word => 'a word * 'a word) w1 - w2))))))" + (RES_FORALL::('a::type word => bool) + => ('a::type word => bool) => bool) + ((PWORDLEN::nat => 'a::type word => bool) n) + (%w1::'a::type word. + (RES_FORALL::('a::type word => bool) + => ('a::type word => bool) => bool) + ((PWORDLEN::nat => 'a::type word => bool) m) + (%w2::'a::type word. + (op =::'a::type word * 'a::type word + => 'a::type word * 'a::type word => bool) + ((WSPLIT::nat + => 'a::type word + => 'a::type word * 'a::type word) + m ((WCAT::'a::type word * 'a::type word + => 'a::type word) + ((Pair::'a::type word + => 'a::type word + => 'a::type word * 'a::type word) + w1 w2))) + ((Pair::'a::type word + => 'a::type word + => 'a::type word * 'a::type word) + w1 w2))))))" by (import word_base WORD_PARTITION) -lemma WCAT_ASSOC: "ALL w1 w2 w3. WCAT (w1, WCAT (w2, w3)) = WCAT (WCAT (w1, w2), w3)" +lemma WCAT_ASSOC: "ALL (w1::'a::type word) (w2::'a::type word) w3::'a::type word. + WCAT (w1, WCAT (w2, w3)) = WCAT (WCAT (w1, w2), w3)" by (import word_base WCAT_ASSOC) -lemma WCAT0: "ALL w. WCAT (WORD [], w) = w & WCAT (w, WORD []) = w" +lemma WCAT0: "ALL w::'a::type word. WCAT (WORD [], w) = w & WCAT (w, WORD []) = w" by (import word_base WCAT0) -lemma WCAT_11: "ALL m n. +lemma WCAT_11: "ALL (m::nat) n::nat. RES_FORALL (PWORDLEN m) - (%wm1. RES_FORALL (PWORDLEN m) - (%wm2. RES_FORALL (PWORDLEN n) - (%wn1. RES_FORALL (PWORDLEN n) - (%wn2. (WCAT (wm1, wn1) = WCAT (wm2, wn2)) = - (wm1 = wm2 & wn1 = wn2)))))" + (%wm1::'a::type word. + RES_FORALL (PWORDLEN m) + (%wm2::'a::type word. + RES_FORALL (PWORDLEN n) + (%wn1::'a::type word. + RES_FORALL (PWORDLEN n) + (%wn2::'a::type word. + (WCAT (wm1, wn1) = WCAT (wm2, wn2)) = + (wm1 = wm2 & wn1 = wn2)))))" by (import word_base WCAT_11) lemma WSPLIT_PWORDLEN: "(All::(nat => bool) => bool) (%n::nat. - (RES_FORALL::('a word => bool) => ('a word => bool) => bool) - ((PWORDLEN::nat => 'a word => bool) n) - (%w::'a word. + (RES_FORALL::('a::type word => bool) + => ('a::type word => bool) => bool) + ((PWORDLEN::nat => 'a::type word => bool) n) + (%w::'a::type word. (All::(nat => bool) => bool) (%m::nat. (op -->::bool => bool => bool) ((op <=::nat => nat => bool) m n) ((op &::bool => bool => bool) - ((IN::'a word => ('a word => bool) => bool) - ((fst::'a word * 'a word => 'a word) - ((WSPLIT::nat => 'a word => 'a word * 'a word) m w)) - ((PWORDLEN::nat => 'a word => bool) + ((IN::'a::type word => ('a::type word => bool) => bool) + ((fst::'a::type word * 'a::type word => 'a::type word) + ((WSPLIT::nat + => 'a::type word + => 'a::type word * 'a::type word) + m w)) + ((PWORDLEN::nat => 'a::type word => bool) ((op -::nat => nat => nat) n m))) - ((IN::'a word => ('a word => bool) => bool) - ((snd::'a word * 'a word => 'a word) - ((WSPLIT::nat => 'a word => 'a word * 'a word) m w)) - ((PWORDLEN::nat => 'a word => bool) m))))))" + ((IN::'a::type word => ('a::type word => bool) => bool) + ((snd::'a::type word * 'a::type word => 'a::type word) + ((WSPLIT::nat + => 'a::type word + => 'a::type word * 'a::type word) + m w)) + ((PWORDLEN::nat => 'a::type word => bool) m))))))" by (import word_base WSPLIT_PWORDLEN) -lemma WCAT_PWORDLEN: "ALL n1. +lemma WCAT_PWORDLEN: "ALL n1::nat. RES_FORALL (PWORDLEN n1) - (%w1. ALL n2. - RES_FORALL (PWORDLEN n2) - (%w2. IN (WCAT (w1, w2)) (PWORDLEN (n1 + n2))))" + (%w1::'a::type word. + ALL n2::nat. + RES_FORALL (PWORDLEN n2) + (%w2::'a::type word. IN (WCAT (w1, w2)) (PWORDLEN (n1 + n2))))" by (import word_base WCAT_PWORDLEN) -lemma WORDLEN_SUC_WCAT: "ALL n w. +lemma WORDLEN_SUC_WCAT: "ALL (n::nat) w::'a::type word. IN w (PWORDLEN (Suc n)) --> - RES_EXISTS (PWORDLEN 1) - (%b. RES_EXISTS (PWORDLEN n) (%w'. w = WCAT (b, w')))" + RES_EXISTS (PWORDLEN (1::nat)) + (%b::'a::type word. + RES_EXISTS (PWORDLEN n) (%w'::'a::type word. w = WCAT (b, w')))" by (import word_base WORDLEN_SUC_WCAT) -lemma WSEG_WSEG: "ALL n. +lemma WSEG_WSEG: "ALL n::nat. RES_FORALL (PWORDLEN n) - (%w. ALL m1 k1 m2 k2. - m1 + k1 <= n & m2 + k2 <= m1 --> - WSEG m2 k2 (WSEG m1 k1 w) = WSEG m2 (k1 + k2) w)" + (%w::'a::type word. + ALL (m1::nat) (k1::nat) (m2::nat) k2::nat. + m1 + k1 <= n & m2 + k2 <= m1 --> + WSEG m2 k2 (WSEG m1 k1 w) = WSEG m2 (k1 + k2) w)" by (import word_base WSEG_WSEG) lemma WSPLIT_WSEG: "(All::(nat => bool) => bool) (%n::nat. - (RES_FORALL::('a word => bool) => ('a word => bool) => bool) - ((PWORDLEN::nat => 'a word => bool) n) - (%w::'a word. + (RES_FORALL::('a::type word => bool) + => ('a::type word => bool) => bool) + ((PWORDLEN::nat => 'a::type word => bool) n) + (%w::'a::type word. (All::(nat => bool) => bool) (%k::nat. (op -->::bool => bool => bool) ((op <=::nat => nat => bool) k n) - ((op =::'a word * 'a word => 'a word * 'a word => bool) - ((WSPLIT::nat => 'a word => 'a word * 'a word) k w) - ((Pair::'a word => 'a word => 'a word * 'a word) - ((WSEG::nat => nat => 'a word => 'a word) + ((op =::'a::type word * 'a::type word + => 'a::type word * 'a::type word => bool) + ((WSPLIT::nat + => 'a::type word + => 'a::type word * 'a::type word) + k w) + ((Pair::'a::type word + => 'a::type word => 'a::type word * 'a::type word) + ((WSEG::nat => nat => 'a::type word => 'a::type word) ((op -::nat => nat => nat) n k) k w) - ((WSEG::nat => nat => 'a word => 'a word) k (0::nat) - w))))))" + ((WSEG::nat => nat => 'a::type word => 'a::type word) k + (0::nat) w))))))" by (import word_base WSPLIT_WSEG) lemma WSPLIT_WSEG1: "(All::(nat => bool) => bool) (%n::nat. - (RES_FORALL::('a word => bool) => ('a word => bool) => bool) - ((PWORDLEN::nat => 'a word => bool) n) - (%w::'a word. + (RES_FORALL::('a::type word => bool) + => ('a::type word => bool) => bool) + ((PWORDLEN::nat => 'a::type word => bool) n) + (%w::'a::type word. (All::(nat => bool) => bool) (%k::nat. (op -->::bool => bool => bool) ((op <=::nat => nat => bool) k n) - ((op =::'a word => 'a word => bool) - ((fst::'a word * 'a word => 'a word) - ((WSPLIT::nat => 'a word => 'a word * 'a word) k w)) - ((WSEG::nat => nat => 'a word => 'a word) + ((op =::'a::type word => 'a::type word => bool) + ((fst::'a::type word * 'a::type word => 'a::type word) + ((WSPLIT::nat + => 'a::type word + => 'a::type word * 'a::type word) + k w)) + ((WSEG::nat => nat => 'a::type word => 'a::type word) ((op -::nat => nat => nat) n k) k w)))))" by (import word_base WSPLIT_WSEG1) lemma WSPLIT_WSEG2: "(All::(nat => bool) => bool) (%n::nat. - (RES_FORALL::('a word => bool) => ('a word => bool) => bool) - ((PWORDLEN::nat => 'a word => bool) n) - (%w::'a word. + (RES_FORALL::('a::type word => bool) + => ('a::type word => bool) => bool) + ((PWORDLEN::nat => 'a::type word => bool) n) + (%w::'a::type word. (All::(nat => bool) => bool) (%k::nat. (op -->::bool => bool => bool) ((op <=::nat => nat => bool) k n) - ((op =::'a word => 'a word => bool) - ((snd::'a word * 'a word => 'a word) - ((WSPLIT::nat => 'a word => 'a word * 'a word) k w)) - ((WSEG::nat => nat => 'a word => 'a word) k (0::nat) - w)))))" + ((op =::'a::type word => 'a::type word => bool) + ((snd::'a::type word * 'a::type word => 'a::type word) + ((WSPLIT::nat + => 'a::type word + => 'a::type word * 'a::type word) + k w)) + ((WSEG::nat => nat => 'a::type word => 'a::type word) k + (0::nat) w)))))" by (import word_base WSPLIT_WSEG2) -lemma WCAT_WSEG_WSEG: "ALL n. +lemma WCAT_WSEG_WSEG: "ALL n::nat. RES_FORALL (PWORDLEN n) - (%w. ALL m1 m2 k. - m1 + (m2 + k) <= n --> - WCAT (WSEG m2 (m1 + k) w, WSEG m1 k w) = WSEG (m1 + m2) k w)" + (%w::'a::type word. + ALL (m1::nat) (m2::nat) k::nat. + m1 + (m2 + k) <= n --> + WCAT (WSEG m2 (m1 + k) w, WSEG m1 k w) = WSEG (m1 + m2) k w)" by (import word_base WCAT_WSEG_WSEG) -lemma WORD_SPLIT: "ALL x xa. - RES_FORALL (PWORDLEN (x + xa)) (%w. w = WCAT (WSEG x xa w, WSEG xa 0 w))" +lemma WORD_SPLIT: "ALL (x::nat) xa::nat. + RES_FORALL (PWORDLEN (x + xa)) + (%w::'a::type word. w = WCAT (WSEG x xa w, WSEG xa (0::nat) w))" by (import word_base WORD_SPLIT) -lemma WORDLEN_SUC_WCAT_WSEG_WSEG: "RES_FORALL (PWORDLEN (Suc n)) (%w. w = WCAT (WSEG 1 n w, WSEG n 0 w))" +lemma WORDLEN_SUC_WCAT_WSEG_WSEG: "RES_FORALL (PWORDLEN (Suc (n::nat))) + (%w::'a::type word. w = WCAT (WSEG (1::nat) n w, WSEG n (0::nat) w))" by (import word_base WORDLEN_SUC_WCAT_WSEG_WSEG) -lemma WORDLEN_SUC_WCAT_WSEG_WSEG_RIGHT: "RES_FORALL (PWORDLEN (Suc n)) (%w. w = WCAT (WSEG n 1 w, WSEG 1 0 w))" +lemma WORDLEN_SUC_WCAT_WSEG_WSEG_RIGHT: "RES_FORALL (PWORDLEN (Suc (n::nat))) + (%w::'a::type word. w = WCAT (WSEG n (1::nat) w, WSEG (1::nat) (0::nat) w))" by (import word_base WORDLEN_SUC_WCAT_WSEG_WSEG_RIGHT) -lemma WORDLEN_SUC_WCAT_BIT_WSEG: "ALL n. - RES_FORALL (PWORDLEN (Suc n)) (%w. w = WCAT (WORD [bit n w], WSEG n 0 w))" +lemma WORDLEN_SUC_WCAT_BIT_WSEG: "ALL n::nat. + RES_FORALL (PWORDLEN (Suc n)) + (%w::'a::type word. w = WCAT (WORD [bit n w], WSEG n (0::nat) w))" by (import word_base WORDLEN_SUC_WCAT_BIT_WSEG) -lemma WORDLEN_SUC_WCAT_BIT_WSEG_RIGHT: "ALL n. - RES_FORALL (PWORDLEN (Suc n)) (%w. w = WCAT (WSEG n 1 w, WORD [bit 0 w]))" +lemma WORDLEN_SUC_WCAT_BIT_WSEG_RIGHT: "ALL n::nat. + RES_FORALL (PWORDLEN (Suc n)) + (%w::'a::type word. w = WCAT (WSEG n (1::nat) w, WORD [bit (0::nat) w]))" by (import word_base WORDLEN_SUC_WCAT_BIT_WSEG_RIGHT) -lemma WSEG_WCAT1: "ALL n1 n2. +lemma WSEG_WCAT1: "ALL (n1::nat) n2::nat. RES_FORALL (PWORDLEN n1) - (%w1. RES_FORALL (PWORDLEN n2) (%w2. WSEG n1 n2 (WCAT (w1, w2)) = w1))" + (%w1::'a::type word. + RES_FORALL (PWORDLEN n2) + (%w2::'a::type word. WSEG n1 n2 (WCAT (w1, w2)) = w1))" by (import word_base WSEG_WCAT1) -lemma WSEG_WCAT2: "ALL n1 n2. +lemma WSEG_WCAT2: "ALL (n1::nat) n2::nat. RES_FORALL (PWORDLEN n1) - (%w1. RES_FORALL (PWORDLEN n2) (%w2. WSEG n2 0 (WCAT (w1, w2)) = w2))" + (%w1::'a::type word. + RES_FORALL (PWORDLEN n2) + (%w2::'a::type word. WSEG n2 (0::nat) (WCAT (w1, w2)) = w2))" by (import word_base WSEG_WCAT2) -lemma WSEG_SUC: "ALL n. +lemma WSEG_SUC: "ALL n::nat. RES_FORALL (PWORDLEN n) - (%w. ALL k m1. - k + Suc m1 < n --> - WSEG (Suc m1) k w = WCAT (WSEG 1 (k + m1) w, WSEG m1 k w))" + (%w::'a::type word. + ALL (k::nat) m1::nat. + k + Suc m1 < n --> + WSEG (Suc m1) k w = WCAT (WSEG (1::nat) (k + m1) w, WSEG m1 k w))" by (import word_base WSEG_SUC) -lemma WORD_CONS_WCAT: "ALL x l. WORD (x # l) = WCAT (WORD [x], WORD l)" +lemma WORD_CONS_WCAT: "ALL (x::'a::type) l::'a::type list. WORD (x # l) = WCAT (WORD [x], WORD l)" by (import word_base WORD_CONS_WCAT) -lemma WORD_SNOC_WCAT: "ALL l x. WORD (SNOC x l) = WCAT (WORD l, WORD [x])" +lemma WORD_SNOC_WCAT: "ALL (l::'a::type list) x::'a::type. + WORD (SNOC x l) = WCAT (WORD l, WORD [x])" by (import word_base WORD_SNOC_WCAT) -lemma BIT_WCAT_FST: "ALL n1 n2. +lemma BIT_WCAT_FST: "ALL (n1::nat) n2::nat. RES_FORALL (PWORDLEN n1) - (%w1. RES_FORALL (PWORDLEN n2) - (%w2. ALL k. - n2 <= k & k < n1 + n2 --> - bit k (WCAT (w1, w2)) = bit (k - n2) w1))" + (%w1::'a::type word. + RES_FORALL (PWORDLEN n2) + (%w2::'a::type word. + ALL k::nat. + n2 <= k & k < n1 + n2 --> + bit k (WCAT (w1, w2)) = bit (k - n2) w1))" by (import word_base BIT_WCAT_FST) lemma BIT_WCAT_SND: "(All::(nat => bool) => bool) (%n1::nat. (All::(nat => bool) => bool) (%n2::nat. - (RES_FORALL::('a word => bool) => ('a word => bool) => bool) - ((PWORDLEN::nat => 'a word => bool) n1) - (%w1::'a word. - (RES_FORALL::('a word => bool) => ('a word => bool) => bool) - ((PWORDLEN::nat => 'a word => bool) n2) - (%w2::'a word. + (RES_FORALL::('a::type word => bool) + => ('a::type word => bool) => bool) + ((PWORDLEN::nat => 'a::type word => bool) n1) + (%w1::'a::type word. + (RES_FORALL::('a::type word => bool) + => ('a::type word => bool) => bool) + ((PWORDLEN::nat => 'a::type word => bool) n2) + (%w2::'a::type word. (All::(nat => bool) => bool) (%k::nat. (op -->::bool => bool => bool) ((op <::nat => nat => bool) k n2) - ((op =::'a => 'a => bool) - ((bit::nat => 'a word => 'a) k - ((WCAT::'a word * 'a word => 'a word) - ((Pair::'a word - => 'a word => 'a word * 'a word) + ((op =::'a::type => 'a::type => bool) + ((bit::nat => 'a::type word => 'a::type) k + ((WCAT::'a::type word * 'a::type word +=> 'a::type word) + ((Pair::'a::type word + => 'a::type word => 'a::type word * 'a::type word) w1 w2))) - ((bit::nat => 'a word => 'a) k w2)))))))" + ((bit::nat => 'a::type word => 'a::type) k + w2)))))))" by (import word_base BIT_WCAT_SND) -lemma BIT_WCAT1: "ALL n. RES_FORALL (PWORDLEN n) (%w. ALL b. bit n (WCAT (WORD [b], w)) = b)" +lemma BIT_WCAT1: "ALL n::nat. + RES_FORALL (PWORDLEN n) + (%w::'a::type word. ALL b::'a::type. bit n (WCAT (WORD [b], w)) = b)" by (import word_base BIT_WCAT1) -lemma WSEG_WCAT_WSEG1: "ALL n1 n2. +lemma WSEG_WCAT_WSEG1: "ALL (n1::nat) n2::nat. RES_FORALL (PWORDLEN n1) - (%w1. RES_FORALL (PWORDLEN n2) - (%w2. ALL m k. - m <= n1 & n2 <= k --> - WSEG m k (WCAT (w1, w2)) = WSEG m (k - n2) w1))" + (%w1::'a::type word. + RES_FORALL (PWORDLEN n2) + (%w2::'a::type word. + ALL (m::nat) k::nat. + m <= n1 & n2 <= k --> + WSEG m k (WCAT (w1, w2)) = WSEG m (k - n2) w1))" by (import word_base WSEG_WCAT_WSEG1) -lemma WSEG_WCAT_WSEG2: "ALL n1 n2. +lemma WSEG_WCAT_WSEG2: "ALL (n1::nat) n2::nat. RES_FORALL (PWORDLEN n1) - (%w1. RES_FORALL (PWORDLEN n2) - (%w2. ALL m k. - m + k <= n2 --> WSEG m k (WCAT (w1, w2)) = WSEG m k w2))" + (%w1::'a::type word. + RES_FORALL (PWORDLEN n2) + (%w2::'a::type word. + ALL (m::nat) k::nat. + m + k <= n2 --> WSEG m k (WCAT (w1, w2)) = WSEG m k w2))" by (import word_base WSEG_WCAT_WSEG2) -lemma WSEG_WCAT_WSEG: "ALL n1 n2. +lemma WSEG_WCAT_WSEG: "ALL (n1::nat) n2::nat. RES_FORALL (PWORDLEN n1) - (%w1. RES_FORALL (PWORDLEN n2) - (%w2. ALL m k. - m + k <= n1 + n2 & k < n2 & n2 <= m + k --> - WSEG m k (WCAT (w1, w2)) = - WCAT (WSEG (m + k - n2) 0 w1, WSEG (n2 - k) k w2)))" + (%w1::'a::type word. + RES_FORALL (PWORDLEN n2) + (%w2::'a::type word. + ALL (m::nat) k::nat. + m + k <= n1 + n2 & k < n2 & n2 <= m + k --> + WSEG m k (WCAT (w1, w2)) = + WCAT (WSEG (m + k - n2) (0::nat) w1, WSEG (n2 - k) k w2)))" by (import word_base WSEG_WCAT_WSEG) lemma BIT_EQ_IMP_WORD_EQ: "(All::(nat => bool) => bool) (%n::nat. - (RES_FORALL::('a word => bool) => ('a word => bool) => bool) - ((PWORDLEN::nat => 'a word => bool) n) - (%w1::'a word. - (RES_FORALL::('a word => bool) => ('a word => bool) => bool) - ((PWORDLEN::nat => 'a word => bool) n) - (%w2::'a word. + (RES_FORALL::('a::type word => bool) + => ('a::type word => bool) => bool) + ((PWORDLEN::nat => 'a::type word => bool) n) + (%w1::'a::type word. + (RES_FORALL::('a::type word => bool) + => ('a::type word => bool) => bool) + ((PWORDLEN::nat => 'a::type word => bool) n) + (%w2::'a::type word. (op -->::bool => bool => bool) ((All::(nat => bool) => bool) (%k::nat. (op -->::bool => bool => bool) ((op <::nat => nat => bool) k n) - ((op =::'a => 'a => bool) - ((bit::nat => 'a word => 'a) k w1) - ((bit::nat => 'a word => 'a) k w2)))) - ((op =::'a word => 'a word => bool) w1 w2))))" + ((op =::'a::type => 'a::type => bool) + ((bit::nat => 'a::type word => 'a::type) k w1) + ((bit::nat => 'a::type word => 'a::type) k w2)))) + ((op =::'a::type word => 'a::type word => bool) w1 w2))))" by (import word_base BIT_EQ_IMP_WORD_EQ) ;end_setup @@ -568,80 +686,99 @@ ;setup_theory word_num constdefs - LVAL :: "('a => nat) => nat => 'a list => nat" - "LVAL == %f b. foldl (%e x. b * e + f x) 0" + LVAL :: "('a::type => nat) => nat => 'a::type list => nat" + "LVAL == +%(f::'a::type => nat) b::nat. + foldl (%(e::nat) x::'a::type. b * e + f x) (0::nat)" -lemma LVAL_DEF: "ALL f b l. LVAL f b l = foldl (%e x. b * e + f x) 0 l" +lemma LVAL_DEF: "ALL (f::'a::type => nat) (b::nat) l::'a::type list. + LVAL f b l = foldl (%(e::nat) x::'a::type. b * e + f x) (0::nat) l" by (import word_num LVAL_DEF) consts - NVAL :: "('a => nat) => nat => 'a word => nat" + NVAL :: "('a::type => nat) => nat => 'a::type word => nat" -specification (NVAL) NVAL_DEF: "ALL f b l. NVAL f b (WORD l) = LVAL f b l" +specification (NVAL) NVAL_DEF: "ALL (f::'a::type => nat) (b::nat) l::'a::type list. + NVAL f b (WORD l) = LVAL f b l" by (import word_num NVAL_DEF) -lemma LVAL: "(ALL (x::'a => nat) xa::nat. LVAL x xa [] = (0::nat)) & -(ALL (x::'a list) (xa::'a => nat) (xb::nat) xc::'a. +lemma LVAL: "(ALL (x::'a::type => nat) xa::nat. LVAL x xa [] = (0::nat)) & +(ALL (x::'a::type list) (xa::'a::type => nat) (xb::nat) xc::'a::type. LVAL xa xb (xc # x) = xa xc * xb ^ length x + LVAL xa xb x)" by (import word_num LVAL) -lemma LVAL_SNOC: "ALL l h f b. LVAL f b (SNOC h l) = LVAL f b l * b + f h" +lemma LVAL_SNOC: "ALL (l::'a::type list) (h::'a::type) (f::'a::type => nat) b::nat. + LVAL f b (SNOC h l) = LVAL f b l * b + f h" by (import word_num LVAL_SNOC) -lemma LVAL_MAX: "ALL l f b. (ALL x. f x < b) --> LVAL f b l < b ^ length l" +lemma LVAL_MAX: "ALL (l::'a::type list) (f::'a::type => nat) b::nat. + (ALL x::'a::type. f x < b) --> LVAL f b l < b ^ length l" by (import word_num LVAL_MAX) -lemma NVAL_MAX: "ALL f b. - (ALL x. f x < b) --> - (ALL n. RES_FORALL (PWORDLEN n) (%w. NVAL f b w < b ^ n))" +lemma NVAL_MAX: "ALL (f::'a::type => nat) b::nat. + (ALL x::'a::type. f x < b) --> + (ALL n::nat. + RES_FORALL (PWORDLEN n) (%w::'a::type word. NVAL f b w < b ^ n))" by (import word_num NVAL_MAX) -lemma NVAL0: "ALL x xa. NVAL x xa (WORD []) = 0" +lemma NVAL0: "ALL (x::'a::type => nat) xa::nat. NVAL x xa (WORD []) = (0::nat)" by (import word_num NVAL0) -lemma NVAL1: "ALL x xa xb. NVAL x xa (WORD [xb]) = x xb" +lemma NVAL1: "ALL (x::'a::type => nat) (xa::nat) xb::'a::type. + NVAL x xa (WORD [xb]) = x xb" by (import word_num NVAL1) -lemma NVAL_WORDLEN_0: "RES_FORALL (PWORDLEN 0) (%w. ALL fv r. NVAL fv r w = 0)" +lemma NVAL_WORDLEN_0: "RES_FORALL (PWORDLEN (0::nat)) + (%w::'a::type word. + ALL (fv::'a::type => nat) r::nat. NVAL fv r w = (0::nat))" by (import word_num NVAL_WORDLEN_0) -lemma NVAL_WCAT1: "ALL w f b x. NVAL f b (WCAT (w, WORD [x])) = NVAL f b w * b + f x" +lemma NVAL_WCAT1: "ALL (w::'a::type word) (f::'a::type => nat) (b::nat) x::'a::type. + NVAL f b (WCAT (w, WORD [x])) = NVAL f b w * b + f x" by (import word_num NVAL_WCAT1) -lemma NVAL_WCAT2: "ALL n. +lemma NVAL_WCAT2: "ALL n::nat. RES_FORALL (PWORDLEN n) - (%w. ALL f b x. - NVAL f b (WCAT (WORD [x], w)) = f x * b ^ n + NVAL f b w)" + (%w::'a::type word. + ALL (f::'a::type => nat) (b::nat) x::'a::type. + NVAL f b (WCAT (WORD [x], w)) = f x * b ^ n + NVAL f b w)" by (import word_num NVAL_WCAT2) -lemma NVAL_WCAT: "ALL n m. +lemma NVAL_WCAT: "ALL (n::nat) m::nat. RES_FORALL (PWORDLEN n) - (%w1. RES_FORALL (PWORDLEN m) - (%w2. ALL f b. - NVAL f b (WCAT (w1, w2)) = - NVAL f b w1 * b ^ m + NVAL f b w2))" + (%w1::'a::type word. + RES_FORALL (PWORDLEN m) + (%w2::'a::type word. + ALL (f::'a::type => nat) b::nat. + NVAL f b (WCAT (w1, w2)) = + NVAL f b w1 * b ^ m + NVAL f b w2))" by (import word_num NVAL_WCAT) consts - NLIST :: "nat => (nat => 'a) => nat => nat => 'a list" + NLIST :: "nat => (nat => 'a::type) => nat => nat => 'a::type list" -specification (NLIST) NLIST_DEF: "(ALL (frep::nat => 'a) (b::nat) m::nat. NLIST (0::nat) frep b m = []) & -(ALL (n::nat) (frep::nat => 'a) (b::nat) m::nat. +specification (NLIST) NLIST_DEF: "(ALL (frep::nat => 'a::type) (b::nat) m::nat. + NLIST (0::nat) frep b m = []) & +(ALL (n::nat) (frep::nat => 'a::type) (b::nat) m::nat. NLIST (Suc n) frep b m = SNOC (frep (m mod b)) (NLIST n frep b (m div b)))" by (import word_num NLIST_DEF) constdefs - NWORD :: "nat => (nat => 'a) => nat => nat => 'a word" - "NWORD == %n frep b m. WORD (NLIST n frep b m)" + NWORD :: "nat => (nat => 'a::type) => nat => nat => 'a::type word" + "NWORD == +%(n::nat) (frep::nat => 'a::type) (b::nat) m::nat. WORD (NLIST n frep b m)" -lemma NWORD_DEF: "ALL n frep b m. NWORD n frep b m = WORD (NLIST n frep b m)" +lemma NWORD_DEF: "ALL (n::nat) (frep::nat => 'a::type) (b::nat) m::nat. + NWORD n frep b m = WORD (NLIST n frep b m)" by (import word_num NWORD_DEF) -lemma NWORD_LENGTH: "ALL x xa xb xc. WORDLEN (NWORD x xa xb xc) = x" +lemma NWORD_LENGTH: "ALL (x::nat) (xa::nat => 'a::type) (xb::nat) xc::nat. + WORDLEN (NWORD x xa xb xc) = x" by (import word_num NWORD_LENGTH) -lemma NWORD_PWORDLEN: "ALL x xa xb xc. IN (NWORD x xa xb xc) (PWORDLEN x)" +lemma NWORD_PWORDLEN: "ALL (x::nat) (xa::nat => 'a::type) (xb::nat) xc::nat. + IN (NWORD x xa xb xc) (PWORDLEN x)" by (import word_num NWORD_PWORDLEN) ;end_setup @@ -649,362 +786,428 @@ ;setup_theory word_bitop consts - PBITOP :: "('a word => 'b word) => bool" + PBITOP :: "('a::type word => 'b::type word) => bool" defs PBITOP_primdef: "PBITOP == GSPEC - (%oper. + (%oper::'a::type word => 'b::type word. (oper, - ALL n. + ALL n::nat. RES_FORALL (PWORDLEN n) - (%w. IN (oper w) (PWORDLEN n) & - (ALL m k. - m + k <= n --> oper (WSEG m k w) = WSEG m k (oper w)))))" + (%w::'a::type word. + IN (oper w) (PWORDLEN n) & + (ALL (m::nat) k::nat. + m + k <= n --> oper (WSEG m k w) = WSEG m k (oper w)))))" lemma PBITOP_def: "PBITOP = GSPEC - (%oper. + (%oper::'a::type word => 'b::type word. (oper, - ALL n. + ALL n::nat. RES_FORALL (PWORDLEN n) - (%w. IN (oper w) (PWORDLEN n) & - (ALL m k. - m + k <= n --> oper (WSEG m k w) = WSEG m k (oper w)))))" + (%w::'a::type word. + IN (oper w) (PWORDLEN n) & + (ALL (m::nat) k::nat. + m + k <= n --> oper (WSEG m k w) = WSEG m k (oper w)))))" by (import word_bitop PBITOP_def) -lemma IN_PBITOP: "ALL oper. +lemma IN_PBITOP: "ALL oper::'a::type word => 'b::type word. IN oper PBITOP = - (ALL n. + (ALL n::nat. RES_FORALL (PWORDLEN n) - (%w. IN (oper w) (PWORDLEN n) & - (ALL m k. - m + k <= n --> oper (WSEG m k w) = WSEG m k (oper w))))" + (%w::'a::type word. + IN (oper w) (PWORDLEN n) & + (ALL (m::nat) k::nat. + m + k <= n --> oper (WSEG m k w) = WSEG m k (oper w))))" by (import word_bitop IN_PBITOP) lemma PBITOP_PWORDLEN: "RES_FORALL PBITOP - (%oper. ALL n. RES_FORALL (PWORDLEN n) (%w. IN (oper w) (PWORDLEN n)))" + (%oper::'a::type word => 'b::type word. + ALL n::nat. + RES_FORALL (PWORDLEN n) + (%w::'a::type word. IN (oper w) (PWORDLEN n)))" by (import word_bitop PBITOP_PWORDLEN) lemma PBITOP_WSEG: "RES_FORALL PBITOP - (%oper. - ALL n. + (%oper::'a::type word => 'b::type word. + ALL n::nat. RES_FORALL (PWORDLEN n) - (%w. ALL m k. - m + k <= n --> oper (WSEG m k w) = WSEG m k (oper w)))" + (%w::'a::type word. + ALL (m::nat) k::nat. + m + k <= n --> oper (WSEG m k w) = WSEG m k (oper w)))" by (import word_bitop PBITOP_WSEG) -lemma PBITOP_BIT: "(RES_FORALL::(('a word => 'b word) => bool) - => (('a word => 'b word) => bool) => bool) - (PBITOP::('a word => 'b word) => bool) - (%oper::'a word => 'b word. +lemma PBITOP_BIT: "(RES_FORALL::(('a::type word => 'b::type word) => bool) + => (('a::type word => 'b::type word) => bool) => bool) + (PBITOP::('a::type word => 'b::type word) => bool) + (%oper::'a::type word => 'b::type word. (All::(nat => bool) => bool) (%n::nat. - (RES_FORALL::('a word => bool) => ('a word => bool) => bool) - ((PWORDLEN::nat => 'a word => bool) n) - (%w::'a word. + (RES_FORALL::('a::type word => bool) + => ('a::type word => bool) => bool) + ((PWORDLEN::nat => 'a::type word => bool) n) + (%w::'a::type word. (All::(nat => bool) => bool) (%k::nat. (op -->::bool => bool => bool) ((op <::nat => nat => bool) k n) - ((op =::'b word => 'b word => bool) + ((op =::'b::type word => 'b::type word => bool) (oper - ((WORD::'a list => 'a word) - ((op #::'a => 'a list => 'a list) - ((bit::nat => 'a word => 'a) k w) - ([]::'a list)))) - ((WORD::'b list => 'b word) - ((op #::'b => 'b list => 'b list) - ((bit::nat => 'b word => 'b) k (oper w)) - ([]::'b list))))))))" + ((WORD::'a::type list => 'a::type word) + ((op #::'a::type + => 'a::type list => 'a::type list) + ((bit::nat => 'a::type word => 'a::type) k w) + ([]::'a::type list)))) + ((WORD::'b::type list => 'b::type word) + ((op #::'b::type => 'b::type list => 'b::type list) + ((bit::nat => 'b::type word => 'b::type) k + (oper w)) + ([]::'b::type list))))))))" by (import word_bitop PBITOP_BIT) consts - PBITBOP :: "('a word => 'b word => 'c word) => bool" + PBITBOP :: "('a::type word => 'b::type word => 'c::type word) => bool" defs PBITBOP_primdef: "PBITBOP == GSPEC - (%oper. + (%oper::'a::type word => 'b::type word => 'c::type word. (oper, - ALL n. + ALL n::nat. RES_FORALL (PWORDLEN n) - (%w1. RES_FORALL (PWORDLEN n) - (%w2. IN (oper w1 w2) (PWORDLEN n) & - (ALL m k. - m + k <= n --> - oper (WSEG m k w1) (WSEG m k w2) = - WSEG m k (oper w1 w2))))))" + (%w1::'a::type word. + RES_FORALL (PWORDLEN n) + (%w2::'b::type word. + IN (oper w1 w2) (PWORDLEN n) & + (ALL (m::nat) k::nat. + m + k <= n --> + oper (WSEG m k w1) (WSEG m k w2) = + WSEG m k (oper w1 w2))))))" lemma PBITBOP_def: "PBITBOP = GSPEC - (%oper. + (%oper::'a::type word => 'b::type word => 'c::type word. (oper, - ALL n. + ALL n::nat. RES_FORALL (PWORDLEN n) - (%w1. RES_FORALL (PWORDLEN n) - (%w2. IN (oper w1 w2) (PWORDLEN n) & - (ALL m k. - m + k <= n --> - oper (WSEG m k w1) (WSEG m k w2) = - WSEG m k (oper w1 w2))))))" + (%w1::'a::type word. + RES_FORALL (PWORDLEN n) + (%w2::'b::type word. + IN (oper w1 w2) (PWORDLEN n) & + (ALL (m::nat) k::nat. + m + k <= n --> + oper (WSEG m k w1) (WSEG m k w2) = + WSEG m k (oper w1 w2))))))" by (import word_bitop PBITBOP_def) -lemma IN_PBITBOP: "ALL oper. +lemma IN_PBITBOP: "ALL oper::'a::type word => 'b::type word => 'c::type word. IN oper PBITBOP = - (ALL n. + (ALL n::nat. RES_FORALL (PWORDLEN n) - (%w1. RES_FORALL (PWORDLEN n) - (%w2. IN (oper w1 w2) (PWORDLEN n) & - (ALL m k. - m + k <= n --> - oper (WSEG m k w1) (WSEG m k w2) = - WSEG m k (oper w1 w2)))))" + (%w1::'a::type word. + RES_FORALL (PWORDLEN n) + (%w2::'b::type word. + IN (oper w1 w2) (PWORDLEN n) & + (ALL (m::nat) k::nat. + m + k <= n --> + oper (WSEG m k w1) (WSEG m k w2) = + WSEG m k (oper w1 w2)))))" by (import word_bitop IN_PBITBOP) lemma PBITBOP_PWORDLEN: "RES_FORALL PBITBOP - (%oper. - ALL n. + (%oper::'a::type word => 'b::type word => 'c::type word. + ALL n::nat. RES_FORALL (PWORDLEN n) - (%w1. RES_FORALL (PWORDLEN n) (%w2. IN (oper w1 w2) (PWORDLEN n))))" + (%w1::'a::type word. + RES_FORALL (PWORDLEN n) + (%w2::'b::type word. IN (oper w1 w2) (PWORDLEN n))))" by (import word_bitop PBITBOP_PWORDLEN) lemma PBITBOP_WSEG: "RES_FORALL PBITBOP - (%oper. - ALL n. + (%oper::'a::type word => 'b::type word => 'c::type word. + ALL n::nat. RES_FORALL (PWORDLEN n) - (%w1. RES_FORALL (PWORDLEN n) - (%w2. ALL m k. - m + k <= n --> - oper (WSEG m k w1) (WSEG m k w2) = - WSEG m k (oper w1 w2))))" + (%w1::'a::type word. + RES_FORALL (PWORDLEN n) + (%w2::'b::type word. + ALL (m::nat) k::nat. + m + k <= n --> + oper (WSEG m k w1) (WSEG m k w2) = + WSEG m k (oper w1 w2))))" by (import word_bitop PBITBOP_WSEG) -lemma PBITBOP_EXISTS: "ALL f. EX x. ALL l1 l2. x (WORD l1) (WORD l2) = WORD (map2 f l1 l2)" +lemma PBITBOP_EXISTS: "ALL f::'a::type => 'b::type => 'c::type. + EX x::'a::type word => 'b::type word => 'c::type word. + ALL (l1::'a::type list) l2::'b::type list. + x (WORD l1) (WORD l2) = WORD (map2 f l1 l2)" by (import word_bitop PBITBOP_EXISTS) consts - WMAP :: "('a => 'b) => 'a word => 'b word" + WMAP :: "('a::type => 'b::type) => 'a::type word => 'b::type word" -specification (WMAP) WMAP_DEF: "ALL f l. WMAP f (WORD l) = WORD (map f l)" +specification (WMAP) WMAP_DEF: "ALL (f::'a::type => 'b::type) l::'a::type list. + WMAP f (WORD l) = WORD (map f l)" by (import word_bitop WMAP_DEF) -lemma WMAP_PWORDLEN: "RES_FORALL (PWORDLEN n) (%w. ALL f. IN (WMAP f w) (PWORDLEN n))" +lemma WMAP_PWORDLEN: "RES_FORALL (PWORDLEN (n::nat)) + (%w::'a::type word. + ALL f::'a::type => 'b::type. IN (WMAP f w) (PWORDLEN n))" by (import word_bitop WMAP_PWORDLEN) -lemma WMAP_0: "ALL x. WMAP x (WORD []) = WORD []" +lemma WMAP_0: "ALL x::'a::type => 'b::type. WMAP x (WORD []) = WORD []" by (import word_bitop WMAP_0) lemma WMAP_BIT: "(All::(nat => bool) => bool) (%n::nat. - (RES_FORALL::('a word => bool) => ('a word => bool) => bool) - ((PWORDLEN::nat => 'a word => bool) n) - (%w::'a word. + (RES_FORALL::('a::type word => bool) + => ('a::type word => bool) => bool) + ((PWORDLEN::nat => 'a::type word => bool) n) + (%w::'a::type word. (All::(nat => bool) => bool) (%k::nat. (op -->::bool => bool => bool) ((op <::nat => nat => bool) k n) - ((All::(('a => 'b) => bool) => bool) - (%f::'a => 'b. - (op =::'b => 'b => bool) - ((bit::nat => 'b word => 'b) k - ((WMAP::('a => 'b) => 'a word => 'b word) f w)) - (f ((bit::nat => 'a word => 'a) k w)))))))" + ((All::(('a::type => 'b::type) => bool) => bool) + (%f::'a::type => 'b::type. + (op =::'b::type => 'b::type => bool) + ((bit::nat => 'b::type word => 'b::type) k + ((WMAP::('a::type => 'b::type) + => 'a::type word => 'b::type word) + f w)) + (f ((bit::nat => 'a::type word => 'a::type) k + w)))))))" by (import word_bitop WMAP_BIT) -lemma WMAP_WSEG: "ALL n. +lemma WMAP_WSEG: "ALL n::nat. RES_FORALL (PWORDLEN n) - (%w. ALL m k. - m + k <= n --> - (ALL f. WMAP f (WSEG m k w) = WSEG m k (WMAP f w)))" + (%w::'a::type word. + ALL (m::nat) k::nat. + m + k <= n --> + (ALL f::'a::type => 'b::type. + WMAP f (WSEG m k w) = WSEG m k (WMAP f w)))" by (import word_bitop WMAP_WSEG) -lemma WMAP_PBITOP: "ALL f. IN (WMAP f) PBITOP" +lemma WMAP_PBITOP: "ALL f::'a::type => 'b::type. IN (WMAP f) PBITOP" by (import word_bitop WMAP_PBITOP) -lemma WMAP_WCAT: "ALL w1 w2 f. WMAP f (WCAT (w1, w2)) = WCAT (WMAP f w1, WMAP f w2)" +lemma WMAP_WCAT: "ALL (w1::'a::type word) (w2::'a::type word) f::'a::type => 'b::type. + WMAP f (WCAT (w1, w2)) = WCAT (WMAP f w1, WMAP f w2)" by (import word_bitop WMAP_WCAT) -lemma WMAP_o: "ALL w f g. WMAP g (WMAP f w) = WMAP (g o f) w" +lemma WMAP_o: "ALL (w::'a::type word) (f::'a::type => 'b::type) g::'b::type => 'c::type. + WMAP g (WMAP f w) = WMAP (g o f) w" by (import word_bitop WMAP_o) consts - FORALLBITS :: "('a => bool) => 'a word => bool" + FORALLBITS :: "('a::type => bool) => 'a::type word => bool" -specification (FORALLBITS) FORALLBITS_DEF: "ALL P l. FORALLBITS P (WORD l) = list_all P l" +specification (FORALLBITS) FORALLBITS_DEF: "ALL (P::'a::type => bool) l::'a::type list. + FORALLBITS P (WORD l) = list_all P l" by (import word_bitop FORALLBITS_DEF) lemma FORALLBITS: "(All::(nat => bool) => bool) (%n::nat. - (RES_FORALL::('a word => bool) => ('a word => bool) => bool) - ((PWORDLEN::nat => 'a word => bool) n) - (%w::'a word. - (All::(('a => bool) => bool) => bool) - (%P::'a => bool. + (RES_FORALL::('a::type word => bool) + => ('a::type word => bool) => bool) + ((PWORDLEN::nat => 'a::type word => bool) n) + (%w::'a::type word. + (All::(('a::type => bool) => bool) => bool) + (%P::'a::type => bool. (op =::bool => bool => bool) - ((FORALLBITS::('a => bool) => 'a word => bool) P w) + ((FORALLBITS::('a::type => bool) => 'a::type word => bool) P + w) ((All::(nat => bool) => bool) (%k::nat. (op -->::bool => bool => bool) ((op <::nat => nat => bool) k n) - (P ((bit::nat => 'a word => 'a) k w)))))))" + (P ((bit::nat => 'a::type word => 'a::type) k + w)))))))" by (import word_bitop FORALLBITS) -lemma FORALLBITS_WSEG: "ALL n. +lemma FORALLBITS_WSEG: "ALL n::nat. RES_FORALL (PWORDLEN n) - (%w. ALL P. - FORALLBITS P w --> - (ALL m k. m + k <= n --> FORALLBITS P (WSEG m k w)))" + (%w::'a::type word. + ALL P::'a::type => bool. + FORALLBITS P w --> + (ALL (m::nat) k::nat. m + k <= n --> FORALLBITS P (WSEG m k w)))" by (import word_bitop FORALLBITS_WSEG) -lemma FORALLBITS_WCAT: "ALL w1 w2 P. +lemma FORALLBITS_WCAT: "ALL (w1::'a::type word) (w2::'a::type word) P::'a::type => bool. FORALLBITS P (WCAT (w1, w2)) = (FORALLBITS P w1 & FORALLBITS P w2)" by (import word_bitop FORALLBITS_WCAT) consts - EXISTSABIT :: "('a => bool) => 'a word => bool" + EXISTSABIT :: "('a::type => bool) => 'a::type word => bool" -specification (EXISTSABIT) EXISTSABIT_DEF: "ALL P l. EXISTSABIT P (WORD l) = list_exists P l" +specification (EXISTSABIT) EXISTSABIT_DEF: "ALL (P::'a::type => bool) l::'a::type list. + EXISTSABIT P (WORD l) = list_exists P l" by (import word_bitop EXISTSABIT_DEF) -lemma NOT_EXISTSABIT: "ALL P w. (~ EXISTSABIT P w) = FORALLBITS (Not o P) w" +lemma NOT_EXISTSABIT: "ALL (P::'a::type => bool) w::'a::type word. + (~ EXISTSABIT P w) = FORALLBITS (Not o P) w" by (import word_bitop NOT_EXISTSABIT) -lemma NOT_FORALLBITS: "ALL P w. (~ FORALLBITS P w) = EXISTSABIT (Not o P) w" +lemma NOT_FORALLBITS: "ALL (P::'a::type => bool) w::'a::type word. + (~ FORALLBITS P w) = EXISTSABIT (Not o P) w" by (import word_bitop NOT_FORALLBITS) lemma EXISTSABIT: "(All::(nat => bool) => bool) (%n::nat. - (RES_FORALL::('a word => bool) => ('a word => bool) => bool) - ((PWORDLEN::nat => 'a word => bool) n) - (%w::'a word. - (All::(('a => bool) => bool) => bool) - (%P::'a => bool. + (RES_FORALL::('a::type word => bool) + => ('a::type word => bool) => bool) + ((PWORDLEN::nat => 'a::type word => bool) n) + (%w::'a::type word. + (All::(('a::type => bool) => bool) => bool) + (%P::'a::type => bool. (op =::bool => bool => bool) - ((EXISTSABIT::('a => bool) => 'a word => bool) P w) + ((EXISTSABIT::('a::type => bool) => 'a::type word => bool) P + w) ((Ex::(nat => bool) => bool) (%k::nat. (op &::bool => bool => bool) ((op <::nat => nat => bool) k n) - (P ((bit::nat => 'a word => 'a) k w)))))))" + (P ((bit::nat => 'a::type word => 'a::type) k + w)))))))" by (import word_bitop EXISTSABIT) -lemma EXISTSABIT_WSEG: "ALL n. +lemma EXISTSABIT_WSEG: "ALL n::nat. RES_FORALL (PWORDLEN n) - (%w. ALL m k. - m + k <= n --> - (ALL P. EXISTSABIT P (WSEG m k w) --> EXISTSABIT P w))" + (%w::'a::type word. + ALL (m::nat) k::nat. + m + k <= n --> + (ALL P::'a::type => bool. + EXISTSABIT P (WSEG m k w) --> EXISTSABIT P w))" by (import word_bitop EXISTSABIT_WSEG) -lemma EXISTSABIT_WCAT: "ALL w1 w2 P. +lemma EXISTSABIT_WCAT: "ALL (w1::'a::type word) (w2::'a::type word) P::'a::type => bool. EXISTSABIT P (WCAT (w1, w2)) = (EXISTSABIT P w1 | EXISTSABIT P w2)" by (import word_bitop EXISTSABIT_WCAT) constdefs - SHR :: "bool => 'a => 'a word => 'a word * 'a" + SHR :: "bool => 'a::type => 'a::type word => 'a::type word * 'a::type" "SHR == -%f b w. +%(f::bool) (b::'a::type) w::'a::type word. (WCAT - (if f then WSEG 1 (PRE (WORDLEN w)) w else WORD [b], - WSEG (PRE (WORDLEN w)) 1 w), - bit 0 w)" + (if f then WSEG (1::nat) (PRE (WORDLEN w)) w else WORD [b], + WSEG (PRE (WORDLEN w)) (1::nat) w), + bit (0::nat) w)" -lemma SHR_DEF: "ALL f b w. +lemma SHR_DEF: "ALL (f::bool) (b::'a::type) w::'a::type word. SHR f b w = (WCAT - (if f then WSEG 1 (PRE (WORDLEN w)) w else WORD [b], - WSEG (PRE (WORDLEN w)) 1 w), - bit 0 w)" + (if f then WSEG (1::nat) (PRE (WORDLEN w)) w else WORD [b], + WSEG (PRE (WORDLEN w)) (1::nat) w), + bit (0::nat) w)" by (import word_bitop SHR_DEF) constdefs - SHL :: "bool => 'a word => 'a => 'a * 'a word" + SHL :: "bool => 'a::type word => 'a::type => 'a::type * 'a::type word" "SHL == -%f w b. +%(f::bool) (w::'a::type word) b::'a::type. (bit (PRE (WORDLEN w)) w, - WCAT (WSEG (PRE (WORDLEN w)) 0 w, if f then WSEG 1 0 w else WORD [b]))" + WCAT + (WSEG (PRE (WORDLEN w)) (0::nat) w, + if f then WSEG (1::nat) (0::nat) w else WORD [b]))" -lemma SHL_DEF: "ALL f w b. +lemma SHL_DEF: "ALL (f::bool) (w::'a::type word) b::'a::type. SHL f w b = (bit (PRE (WORDLEN w)) w, - WCAT (WSEG (PRE (WORDLEN w)) 0 w, if f then WSEG 1 0 w else WORD [b]))" + WCAT + (WSEG (PRE (WORDLEN w)) (0::nat) w, + if f then WSEG (1::nat) (0::nat) w else WORD [b]))" by (import word_bitop SHL_DEF) -lemma SHR_WSEG: "ALL n. +lemma SHR_WSEG: "ALL n::nat. RES_FORALL (PWORDLEN n) - (%w. ALL m k. - m + k <= n --> - 0 < m --> - (ALL f b. - SHR f b (WSEG m k w) = - (if f - then WCAT (WSEG 1 (k + (m - 1)) w, WSEG (m - 1) (k + 1) w) - else WCAT (WORD [b], WSEG (m - 1) (k + 1) w), - bit k w)))" + (%w::'a::type word. + ALL (m::nat) k::nat. + m + k <= n --> + (0::nat) < m --> + (ALL (f::bool) b::'a::type. + SHR f b (WSEG m k w) = + (if f + then WCAT + (WSEG (1::nat) (k + (m - (1::nat))) w, + WSEG (m - (1::nat)) (k + (1::nat)) w) + else WCAT (WORD [b], WSEG (m - (1::nat)) (k + (1::nat)) w), + bit k w)))" by (import word_bitop SHR_WSEG) -lemma SHR_WSEG_1F: "ALL n. +lemma SHR_WSEG_1F: "ALL n::nat. RES_FORALL (PWORDLEN n) - (%w. ALL b m k. - m + k <= n --> - 0 < m --> - SHR False b (WSEG m k w) = - (WCAT (WORD [b], WSEG (m - 1) (k + 1) w), bit k w))" + (%w::'a::type word. + ALL (b::'a::type) (m::nat) k::nat. + m + k <= n --> + (0::nat) < m --> + SHR False b (WSEG m k w) = + (WCAT (WORD [b], WSEG (m - (1::nat)) (k + (1::nat)) w), bit k w))" by (import word_bitop SHR_WSEG_1F) -lemma SHR_WSEG_NF: "ALL n. +lemma SHR_WSEG_NF: "ALL n::nat. RES_FORALL (PWORDLEN n) - (%w. ALL m k. - m + k < n --> - 0 < m --> - SHR False (bit (m + k) w) (WSEG m k w) = - (WSEG m (k + 1) w, bit k w))" + (%w::'a::type word. + ALL (m::nat) k::nat. + m + k < n --> + (0::nat) < m --> + SHR False (bit (m + k) w) (WSEG m k w) = + (WSEG m (k + (1::nat)) w, bit k w))" by (import word_bitop SHR_WSEG_NF) -lemma SHL_WSEG: "ALL n. +lemma SHL_WSEG: "ALL n::nat. RES_FORALL (PWORDLEN n) - (%w. ALL m k. - m + k <= n --> - 0 < m --> - (ALL f b. - SHL f (WSEG m k w) b = - (bit (k + (m - 1)) w, - if f then WCAT (WSEG (m - 1) k w, WSEG 1 k w) - else WCAT (WSEG (m - 1) k w, WORD [b]))))" + (%w::'a::type word. + ALL (m::nat) k::nat. + m + k <= n --> + (0::nat) < m --> + (ALL (f::bool) b::'a::type. + SHL f (WSEG m k w) b = + (bit (k + (m - (1::nat))) w, + if f then WCAT (WSEG (m - (1::nat)) k w, WSEG (1::nat) k w) + else WCAT (WSEG (m - (1::nat)) k w, WORD [b]))))" by (import word_bitop SHL_WSEG) -lemma SHL_WSEG_1F: "ALL n. +lemma SHL_WSEG_1F: "ALL n::nat. RES_FORALL (PWORDLEN n) - (%w. ALL b m k. - m + k <= n --> - 0 < m --> - SHL False (WSEG m k w) b = - (bit (k + (m - 1)) w, WCAT (WSEG (m - 1) k w, WORD [b])))" + (%w::'a::type word. + ALL (b::'a::type) (m::nat) k::nat. + m + k <= n --> + (0::nat) < m --> + SHL False (WSEG m k w) b = + (bit (k + (m - (1::nat))) w, + WCAT (WSEG (m - (1::nat)) k w, WORD [b])))" by (import word_bitop SHL_WSEG_1F) -lemma SHL_WSEG_NF: "ALL n. +lemma SHL_WSEG_NF: "ALL n::nat. RES_FORALL (PWORDLEN n) - (%w. ALL m k. - m + k <= n --> - 0 < m --> - 0 < k --> - SHL False (WSEG m k w) (bit (k - 1) w) = - (bit (k + (m - 1)) w, WSEG m (k - 1) w))" + (%w::'a::type word. + ALL (m::nat) k::nat. + m + k <= n --> + (0::nat) < m --> + (0::nat) < k --> + SHL False (WSEG m k w) (bit (k - (1::nat)) w) = + (bit (k + (m - (1::nat))) w, WSEG m (k - (1::nat)) w))" by (import word_bitop SHL_WSEG_NF) -lemma WSEG_SHL: "ALL n. +lemma WSEG_SHL: "ALL n::nat. RES_FORALL (PWORDLEN (Suc n)) - (%w. ALL m k. - 0 < k & m + k <= Suc n --> - (ALL b. WSEG m k (snd (SHL f w b)) = WSEG m (k - 1) w))" + (%w::'a::type word. + ALL (m::nat) k::nat. + (0::nat) < k & m + k <= Suc n --> + (ALL b::'a::type. + WSEG m k (snd (SHL (f::bool) w b)) = + WSEG m (k - (1::nat)) w))" by (import word_bitop WSEG_SHL) -lemma WSEG_SHL_0: "ALL n. +lemma WSEG_SHL_0: "ALL n::nat. RES_FORALL (PWORDLEN (Suc n)) - (%w. ALL m b. - 0 < m & m <= Suc n --> - WSEG m 0 (snd (SHL f w b)) = - WCAT (WSEG (m - 1) 0 w, if f then WSEG 1 0 w else WORD [b]))" + (%w::'a::type word. + ALL (m::nat) b::'a::type. + (0::nat) < m & m <= Suc n --> + WSEG m (0::nat) (snd (SHL (f::bool) w b)) = + WCAT + (WSEG (m - (1::nat)) (0::nat) w, + if f then WSEG (1::nat) (0::nat) w else WORD [b]))" by (import word_bitop WSEG_SHL_0) ;end_setup @@ -1013,78 +1216,86 @@ constdefs BV :: "bool => nat" - "BV == %b. if b then Suc 0 else 0" + "(BV == (%(b::bool). (if b then (Suc (0::nat)) else (0::nat))))" -lemma BV_DEF: "ALL b. BV b = (if b then Suc 0 else 0)" +lemma BV_DEF: "(ALL (b::bool). ((BV b) = (if b then (Suc (0::nat)) else (0::nat))))" by (import bword_num BV_DEF) consts BNVAL :: "bool word => nat" -specification (BNVAL) BNVAL_DEF: "ALL l. BNVAL (WORD l) = LVAL BV 2 l" +specification (BNVAL) BNVAL_DEF: "ALL l::bool list. BNVAL (WORD l) = LVAL BV (2::nat) l" by (import bword_num BNVAL_DEF) -lemma BV_LESS_2: "ALL x. BV x < 2" +lemma BV_LESS_2: "ALL x::bool. BV x < (2::nat)" by (import bword_num BV_LESS_2) -lemma BNVAL_NVAL: "ALL w. BNVAL w = NVAL BV 2 w" +lemma BNVAL_NVAL: "ALL w::bool word. BNVAL w = NVAL BV (2::nat) w" by (import bword_num BNVAL_NVAL) -lemma BNVAL0: "BNVAL (WORD []) = 0" +lemma BNVAL0: "BNVAL (WORD []) = (0::nat)" by (import bword_num BNVAL0) -lemma BNVAL_11: "ALL w1 w2. WORDLEN w1 = WORDLEN w2 --> BNVAL w1 = BNVAL w2 --> w1 = w2" +lemma BNVAL_11: "ALL (w1::bool word) w2::bool word. + WORDLEN w1 = WORDLEN w2 --> BNVAL w1 = BNVAL w2 --> w1 = w2" by (import bword_num BNVAL_11) -lemma BNVAL_ONTO: "ALL w. Ex (op = (BNVAL w))" +lemma BNVAL_ONTO: "ALL w::bool word. Ex (op = (BNVAL w))" by (import bword_num BNVAL_ONTO) -lemma BNVAL_MAX: "ALL n. RES_FORALL (PWORDLEN n) (%w. BNVAL w < 2 ^ n)" +lemma BNVAL_MAX: "ALL n::nat. RES_FORALL (PWORDLEN n) (%w::bool word. BNVAL w < (2::nat) ^ n)" by (import bword_num BNVAL_MAX) -lemma BNVAL_WCAT1: "ALL n. +lemma BNVAL_WCAT1: "ALL n::nat. RES_FORALL (PWORDLEN n) - (%w. ALL x. BNVAL (WCAT (w, WORD [x])) = BNVAL w * 2 + BV x)" + (%w::bool word. + ALL x::bool. BNVAL (WCAT (w, WORD [x])) = BNVAL w * (2::nat) + BV x)" by (import bword_num BNVAL_WCAT1) -lemma BNVAL_WCAT2: "ALL n. +lemma BNVAL_WCAT2: "ALL n::nat. RES_FORALL (PWORDLEN n) - (%w. ALL x. BNVAL (WCAT (WORD [x], w)) = BV x * 2 ^ n + BNVAL w)" + (%w::bool word. + ALL x::bool. + BNVAL (WCAT (WORD [x], w)) = BV x * (2::nat) ^ n + BNVAL w)" by (import bword_num BNVAL_WCAT2) -lemma BNVAL_WCAT: "ALL n m. +lemma BNVAL_WCAT: "ALL (n::nat) m::nat. RES_FORALL (PWORDLEN n) - (%w1. RES_FORALL (PWORDLEN m) - (%w2. BNVAL (WCAT (w1, w2)) = BNVAL w1 * 2 ^ m + BNVAL w2))" + (%w1::bool word. + RES_FORALL (PWORDLEN m) + (%w2::bool word. + BNVAL (WCAT (w1, w2)) = BNVAL w1 * (2::nat) ^ m + BNVAL w2))" by (import bword_num BNVAL_WCAT) constdefs VB :: "nat => bool" - "VB == %n. n mod 2 ~= 0" + "VB == %n::nat. n mod (2::nat) ~= (0::nat)" -lemma VB_DEF: "ALL n. VB n = (n mod 2 ~= 0)" +lemma VB_DEF: "ALL n::nat. VB n = (n mod (2::nat) ~= (0::nat))" by (import bword_num VB_DEF) constdefs NBWORD :: "nat => nat => bool word" - "NBWORD == %n m. WORD (NLIST n VB 2 m)" + "NBWORD == %(n::nat) m::nat. WORD (NLIST n VB (2::nat) m)" -lemma NBWORD_DEF: "ALL n m. NBWORD n m = WORD (NLIST n VB 2 m)" +lemma NBWORD_DEF: "ALL (n::nat) m::nat. NBWORD n m = WORD (NLIST n VB (2::nat) m)" by (import bword_num NBWORD_DEF) -lemma NBWORD0: "ALL x. NBWORD 0 x = WORD []" +lemma NBWORD0: "ALL x::nat. NBWORD (0::nat) x = WORD []" by (import bword_num NBWORD0) -lemma WORDLEN_NBWORD: "ALL x xa. WORDLEN (NBWORD x xa) = x" +lemma WORDLEN_NBWORD: "ALL (x::nat) xa::nat. WORDLEN (NBWORD x xa) = x" by (import bword_num WORDLEN_NBWORD) -lemma PWORDLEN_NBWORD: "ALL x xa. IN (NBWORD x xa) (PWORDLEN x)" +lemma PWORDLEN_NBWORD: "ALL (x::nat) xa::nat. IN (NBWORD x xa) (PWORDLEN x)" by (import bword_num PWORDLEN_NBWORD) -lemma NBWORD_SUC: "ALL n m. NBWORD (Suc n) m = WCAT (NBWORD n (m div 2), WORD [VB (m mod 2)])" +lemma NBWORD_SUC: "ALL (n::nat) m::nat. + NBWORD (Suc n) m = + WCAT (NBWORD n (m div (2::nat)), WORD [VB (m mod (2::nat))])" by (import bword_num NBWORD_SUC) -lemma VB_BV: "ALL x. VB (BV x) = x" +lemma VB_BV: "ALL x::bool. VB (BV x) = x" by (import bword_num VB_BV) lemma BV_VB: "(All::(nat => bool) => bool) @@ -1099,19 +1310,24 @@ x))" by (import bword_num BV_VB) -lemma NBWORD_BNVAL: "ALL n. RES_FORALL (PWORDLEN n) (%w. NBWORD n (BNVAL w) = w)" +lemma NBWORD_BNVAL: "ALL n::nat. RES_FORALL (PWORDLEN n) (%w::bool word. NBWORD n (BNVAL w) = w)" by (import bword_num NBWORD_BNVAL) -lemma BNVAL_NBWORD: "ALL n m. m < 2 ^ n --> BNVAL (NBWORD n m) = m" +lemma BNVAL_NBWORD: "ALL (n::nat) m::nat. m < (2::nat) ^ n --> BNVAL (NBWORD n m) = m" by (import bword_num BNVAL_NBWORD) -lemma ZERO_WORD_VAL: "RES_FORALL (PWORDLEN n) (%w. (w = NBWORD n 0) = (BNVAL w = 0))" +lemma ZERO_WORD_VAL: "RES_FORALL (PWORDLEN (n::nat)) + (%w::bool word. (w = NBWORD n (0::nat)) = (BNVAL w = (0::nat)))" by (import bword_num ZERO_WORD_VAL) -lemma WCAT_NBWORD_0: "ALL n1 n2. WCAT (NBWORD n1 0, NBWORD n2 0) = NBWORD (n1 + n2) 0" +lemma WCAT_NBWORD_0: "ALL (n1::nat) n2::nat. + WCAT (NBWORD n1 (0::nat), NBWORD n2 (0::nat)) = NBWORD (n1 + n2) (0::nat)" by (import bword_num WCAT_NBWORD_0) -lemma WSPLIT_NBWORD_0: "ALL n m. m <= n --> WSPLIT m (NBWORD n 0) = (NBWORD (n - m) 0, NBWORD m 0)" +lemma WSPLIT_NBWORD_0: "ALL (n::nat) m::nat. + m <= n --> + WSPLIT m (NBWORD n (0::nat)) = + (NBWORD (n - m) (0::nat), NBWORD m (0::nat))" by (import bword_num WSPLIT_NBWORD_0) lemma EQ_NBWORD0_SPLIT: "(All::(nat => bool) => bool) @@ -1138,60 +1354,75 @@ ((NBWORD::nat => nat => bool word) m (0::nat))))))))" by (import bword_num EQ_NBWORD0_SPLIT) -lemma NBWORD_MOD: "ALL n m. NBWORD n (m mod 2 ^ n) = NBWORD n m" +lemma NBWORD_MOD: "ALL (n::nat) m::nat. NBWORD n (m mod (2::nat) ^ n) = NBWORD n m" by (import bword_num NBWORD_MOD) -lemma WSEG_NBWORD_SUC: "ALL n m. WSEG n 0 (NBWORD (Suc n) m) = NBWORD n m" +lemma WSEG_NBWORD_SUC: "ALL (n::nat) m::nat. WSEG n (0::nat) (NBWORD (Suc n) m) = NBWORD n m" by (import bword_num WSEG_NBWORD_SUC) -lemma NBWORD_SUC_WSEG: "ALL n. RES_FORALL (PWORDLEN (Suc n)) (%w. NBWORD n (BNVAL w) = WSEG n 0 w)" +lemma NBWORD_SUC_WSEG: "ALL n::nat. + RES_FORALL (PWORDLEN (Suc n)) + (%w::bool word. NBWORD n (BNVAL w) = WSEG n (0::nat) w)" by (import bword_num NBWORD_SUC_WSEG) -lemma DOUBL_EQ_SHL: "ALL x>0. +lemma DOUBL_EQ_SHL: "ALL x>0::nat. RES_FORALL (PWORDLEN x) - (%xa. ALL xb. - NBWORD x (BNVAL xa + BNVAL xa + BV xb) = snd (SHL False xa xb))" + (%xa::bool word. + ALL xb::bool. + NBWORD x (BNVAL xa + BNVAL xa + BV xb) = snd (SHL False xa xb))" by (import bword_num DOUBL_EQ_SHL) -lemma MSB_NBWORD: "ALL n m. bit n (NBWORD (Suc n) m) = VB (m div 2 ^ n mod 2)" +lemma MSB_NBWORD: "ALL (n::nat) m::nat. + bit n (NBWORD (Suc n) m) = VB (m div (2::nat) ^ n mod (2::nat))" by (import bword_num MSB_NBWORD) -lemma NBWORD_SPLIT: "ALL n1 n2 m. - NBWORD (n1 + n2) m = WCAT (NBWORD n1 (m div 2 ^ n2), NBWORD n2 m)" +lemma NBWORD_SPLIT: "ALL (n1::nat) (n2::nat) m::nat. + NBWORD (n1 + n2) m = WCAT (NBWORD n1 (m div (2::nat) ^ n2), NBWORD n2 m)" by (import bword_num NBWORD_SPLIT) -lemma WSEG_NBWORD: "ALL m k n. - m + k <= n --> (ALL l. WSEG m k (NBWORD n l) = NBWORD m (l div 2 ^ k))" +lemma WSEG_NBWORD: "ALL (m::nat) (k::nat) n::nat. + m + k <= n --> + (ALL l::nat. WSEG m k (NBWORD n l) = NBWORD m (l div (2::nat) ^ k))" by (import bword_num WSEG_NBWORD) -lemma NBWORD_SUC_FST: "ALL n x. NBWORD (Suc n) x = WCAT (WORD [VB (x div 2 ^ n mod 2)], NBWORD n x)" +lemma NBWORD_SUC_FST: "ALL (n::nat) x::nat. + NBWORD (Suc n) x = + WCAT (WORD [VB (x div (2::nat) ^ n mod (2::nat))], NBWORD n x)" by (import bword_num NBWORD_SUC_FST) -lemma BIT_NBWORD0: "ALL k n. k < n --> bit k (NBWORD n 0) = False" +lemma BIT_NBWORD0: "ALL (k::nat) n::nat. k < n --> bit k (NBWORD n (0::nat)) = False" by (import bword_num BIT_NBWORD0) -lemma ADD_BNVAL_LEFT: "ALL n. +lemma ADD_BNVAL_LEFT: "ALL n::nat. RES_FORALL (PWORDLEN (Suc n)) - (%w1. RES_FORALL (PWORDLEN (Suc n)) - (%w2. BNVAL w1 + BNVAL w2 = - (BV (bit n w1) + BV (bit n w2)) * 2 ^ n + - (BNVAL (WSEG n 0 w1) + BNVAL (WSEG n 0 w2))))" + (%w1::bool word. + RES_FORALL (PWORDLEN (Suc n)) + (%w2::bool word. + BNVAL w1 + BNVAL w2 = + (BV (bit n w1) + BV (bit n w2)) * (2::nat) ^ n + + (BNVAL (WSEG n (0::nat) w1) + BNVAL (WSEG n (0::nat) w2))))" by (import bword_num ADD_BNVAL_LEFT) -lemma ADD_BNVAL_RIGHT: "ALL n. +lemma ADD_BNVAL_RIGHT: "ALL n::nat. RES_FORALL (PWORDLEN (Suc n)) - (%w1. RES_FORALL (PWORDLEN (Suc n)) - (%w2. BNVAL w1 + BNVAL w2 = - (BNVAL (WSEG n 1 w1) + BNVAL (WSEG n 1 w2)) * 2 + - (BV (bit 0 w1) + BV (bit 0 w2))))" + (%w1::bool word. + RES_FORALL (PWORDLEN (Suc n)) + (%w2::bool word. + BNVAL w1 + BNVAL w2 = + (BNVAL (WSEG n (1::nat) w1) + BNVAL (WSEG n (1::nat) w2)) * + (2::nat) + + (BV (bit (0::nat) w1) + BV (bit (0::nat) w2))))" by (import bword_num ADD_BNVAL_RIGHT) -lemma ADD_BNVAL_SPLIT: "ALL n1 n2. +lemma ADD_BNVAL_SPLIT: "ALL (n1::nat) n2::nat. RES_FORALL (PWORDLEN (n1 + n2)) - (%w1. RES_FORALL (PWORDLEN (n1 + n2)) - (%w2. BNVAL w1 + BNVAL w2 = - (BNVAL (WSEG n1 n2 w1) + BNVAL (WSEG n1 n2 w2)) * 2 ^ n2 + - (BNVAL (WSEG n2 0 w1) + BNVAL (WSEG n2 0 w2))))" + (%w1::bool word. + RES_FORALL (PWORDLEN (n1 + n2)) + (%w2::bool word. + BNVAL w1 + BNVAL w2 = + (BNVAL (WSEG n1 n2 w1) + BNVAL (WSEG n1 n2 w2)) * + (2::nat) ^ n2 + + (BNVAL (WSEG n2 (0::nat) w1) + BNVAL (WSEG n2 (0::nat) w2))))" by (import bword_num ADD_BNVAL_SPLIT) ;end_setup @@ -1201,59 +1432,75 @@ consts ACARRY :: "nat => bool word => bool word => bool => bool" -specification (ACARRY) ACARRY_DEF: "(ALL w1 w2 cin. ACARRY 0 w1 w2 cin = cin) & -(ALL n w1 w2 cin. +specification (ACARRY) ACARRY_DEF: "(ALL (w1::bool word) (w2::bool word) cin::bool. + ACARRY (0::nat) w1 w2 cin = cin) & +(ALL (n::nat) (w1::bool word) (w2::bool word) cin::bool. ACARRY (Suc n) w1 w2 cin = - VB ((BV (bit n w1) + BV (bit n w2) + BV (ACARRY n w1 w2 cin)) div 2))" + VB ((BV (bit n w1) + BV (bit n w2) + BV (ACARRY n w1 w2 cin)) div + (2::nat)))" by (import bword_arith ACARRY_DEF) consts ICARRY :: "nat => bool word => bool word => bool => bool" -specification (ICARRY) ICARRY_DEF: "(ALL w1 w2 cin. ICARRY 0 w1 w2 cin = cin) & -(ALL n w1 w2 cin. +specification (ICARRY) ICARRY_DEF: "(ALL (w1::bool word) (w2::bool word) cin::bool. + ICARRY (0::nat) w1 w2 cin = cin) & +(ALL (n::nat) (w1::bool word) (w2::bool word) cin::bool. ICARRY (Suc n) w1 w2 cin = (bit n w1 & bit n w2 | (bit n w1 | bit n w2) & ICARRY n w1 w2 cin))" by (import bword_arith ICARRY_DEF) -lemma ACARRY_EQ_ICARRY: "ALL n. +lemma ACARRY_EQ_ICARRY: "ALL n::nat. RES_FORALL (PWORDLEN n) - (%w1. RES_FORALL (PWORDLEN n) - (%w2. ALL cin k. - k <= n --> ACARRY k w1 w2 cin = ICARRY k w1 w2 cin))" + (%w1::bool word. + RES_FORALL (PWORDLEN n) + (%w2::bool word. + ALL (cin::bool) k::nat. + k <= n --> ACARRY k w1 w2 cin = ICARRY k w1 w2 cin))" by (import bword_arith ACARRY_EQ_ICARRY) -lemma BNVAL_LESS_EQ: "ALL n. RES_FORALL (PWORDLEN n) (%w. BNVAL w <= 2 ^ n - 1)" +lemma BNVAL_LESS_EQ: "ALL n::nat. + RES_FORALL (PWORDLEN n) + (%w::bool word. BNVAL w <= (2::nat) ^ n - (1::nat))" by (import bword_arith BNVAL_LESS_EQ) -lemma ADD_BNVAL_LESS_EQ1: "ALL n cin. +lemma ADD_BNVAL_LESS_EQ1: "ALL (n::nat) cin::bool. RES_FORALL (PWORDLEN n) - (%w1. RES_FORALL (PWORDLEN n) - (%w2. (BNVAL w1 + (BNVAL w2 + BV cin)) div 2 ^ n <= Suc 0))" + (%w1::bool word. + RES_FORALL (PWORDLEN n) + (%w2::bool word. + (BNVAL w1 + (BNVAL w2 + BV cin)) div (2::nat) ^ n + <= Suc (0::nat)))" by (import bword_arith ADD_BNVAL_LESS_EQ1) -lemma ADD_BV_BNVAL_DIV_LESS_EQ1: "ALL n x1 x2 cin. +lemma ADD_BV_BNVAL_DIV_LESS_EQ1: "ALL (n::nat) (x1::bool) (x2::bool) cin::bool. RES_FORALL (PWORDLEN n) - (%w1. RES_FORALL (PWORDLEN n) - (%w2. (BV x1 + BV x2 + - (BNVAL w1 + (BNVAL w2 + BV cin)) div 2 ^ n) div - 2 - <= 1))" + (%w1::bool word. + RES_FORALL (PWORDLEN n) + (%w2::bool word. + (BV x1 + BV x2 + + (BNVAL w1 + (BNVAL w2 + BV cin)) div (2::nat) ^ n) div + (2::nat) + <= (1::nat)))" by (import bword_arith ADD_BV_BNVAL_DIV_LESS_EQ1) -lemma ADD_BV_BNVAL_LESS_EQ: "ALL n x1 x2 cin. +lemma ADD_BV_BNVAL_LESS_EQ: "ALL (n::nat) (x1::bool) (x2::bool) cin::bool. RES_FORALL (PWORDLEN n) - (%w1. RES_FORALL (PWORDLEN n) - (%w2. BV x1 + BV x2 + (BNVAL w1 + (BNVAL w2 + BV cin)) - <= Suc (2 ^ Suc n)))" + (%w1::bool word. + RES_FORALL (PWORDLEN n) + (%w2::bool word. + BV x1 + BV x2 + (BNVAL w1 + (BNVAL w2 + BV cin)) + <= Suc ((2::nat) ^ Suc n)))" by (import bword_arith ADD_BV_BNVAL_LESS_EQ) -lemma ADD_BV_BNVAL_LESS_EQ1: "ALL n x1 x2 cin. +lemma ADD_BV_BNVAL_LESS_EQ1: "ALL (n::nat) (x1::bool) (x2::bool) cin::bool. RES_FORALL (PWORDLEN n) - (%w1. RES_FORALL (PWORDLEN n) - (%w2. (BV x1 + BV x2 + (BNVAL w1 + (BNVAL w2 + BV cin))) div - 2 ^ Suc n - <= 1))" + (%w1::bool word. + RES_FORALL (PWORDLEN n) + (%w2::bool word. + (BV x1 + BV x2 + (BNVAL w1 + (BNVAL w2 + BV cin))) div + (2::nat) ^ Suc n + <= (1::nat)))" by (import bword_arith ADD_BV_BNVAL_LESS_EQ1) lemma ACARRY_EQ_ADD_DIV: "(All::(nat => bool) => bool) @@ -1293,81 +1540,97 @@ k)))))))" by (import bword_arith ACARRY_EQ_ADD_DIV) -lemma ADD_WORD_SPLIT: "ALL n1 n2. +lemma ADD_WORD_SPLIT: "ALL (n1::nat) n2::nat. RES_FORALL (PWORDLEN (n1 + n2)) - (%w1. RES_FORALL (PWORDLEN (n1 + n2)) - (%w2. ALL cin. - NBWORD (n1 + n2) (BNVAL w1 + BNVAL w2 + BV cin) = - WCAT - (NBWORD n1 - (BNVAL (WSEG n1 n2 w1) + BNVAL (WSEG n1 n2 w2) + - BV (ACARRY n2 w1 w2 cin)), - NBWORD n2 - (BNVAL (WSEG n2 0 w1) + BNVAL (WSEG n2 0 w2) + - BV cin))))" + (%w1::bool word. + RES_FORALL (PWORDLEN (n1 + n2)) + (%w2::bool word. + ALL cin::bool. + NBWORD (n1 + n2) (BNVAL w1 + BNVAL w2 + BV cin) = + WCAT + (NBWORD n1 + (BNVAL (WSEG n1 n2 w1) + BNVAL (WSEG n1 n2 w2) + + BV (ACARRY n2 w1 w2 cin)), + NBWORD n2 + (BNVAL (WSEG n2 (0::nat) w1) + + BNVAL (WSEG n2 (0::nat) w2) + + BV cin))))" by (import bword_arith ADD_WORD_SPLIT) -lemma WSEG_NBWORD_ADD: "ALL n. +lemma WSEG_NBWORD_ADD: "ALL n::nat. RES_FORALL (PWORDLEN n) - (%w1. RES_FORALL (PWORDLEN n) - (%w2. ALL m k cin. - m + k <= n --> - WSEG m k (NBWORD n (BNVAL w1 + BNVAL w2 + BV cin)) = - NBWORD m - (BNVAL (WSEG m k w1) + BNVAL (WSEG m k w2) + - BV (ACARRY k w1 w2 cin))))" + (%w1::bool word. + RES_FORALL (PWORDLEN n) + (%w2::bool word. + ALL (m::nat) (k::nat) cin::bool. + m + k <= n --> + WSEG m k (NBWORD n (BNVAL w1 + BNVAL w2 + BV cin)) = + NBWORD m + (BNVAL (WSEG m k w1) + BNVAL (WSEG m k w2) + + BV (ACARRY k w1 w2 cin))))" by (import bword_arith WSEG_NBWORD_ADD) -lemma ADD_NBWORD_EQ0_SPLIT: "ALL n1 n2. +lemma ADD_NBWORD_EQ0_SPLIT: "ALL (n1::nat) n2::nat. RES_FORALL (PWORDLEN (n1 + n2)) - (%w1. RES_FORALL (PWORDLEN (n1 + n2)) - (%w2. ALL cin. - (NBWORD (n1 + n2) (BNVAL w1 + BNVAL w2 + BV cin) = - NBWORD (n1 + n2) 0) = - (NBWORD n1 - (BNVAL (WSEG n1 n2 w1) + BNVAL (WSEG n1 n2 w2) + - BV (ACARRY n2 w1 w2 cin)) = - NBWORD n1 0 & - NBWORD n2 - (BNVAL (WSEG n2 0 w1) + BNVAL (WSEG n2 0 w2) + - BV cin) = - NBWORD n2 0)))" + (%w1::bool word. + RES_FORALL (PWORDLEN (n1 + n2)) + (%w2::bool word. + ALL cin::bool. + (NBWORD (n1 + n2) (BNVAL w1 + BNVAL w2 + BV cin) = + NBWORD (n1 + n2) (0::nat)) = + (NBWORD n1 + (BNVAL (WSEG n1 n2 w1) + BNVAL (WSEG n1 n2 w2) + + BV (ACARRY n2 w1 w2 cin)) = + NBWORD n1 (0::nat) & + NBWORD n2 + (BNVAL (WSEG n2 (0::nat) w1) + + BNVAL (WSEG n2 (0::nat) w2) + + BV cin) = + NBWORD n2 (0::nat))))" by (import bword_arith ADD_NBWORD_EQ0_SPLIT) -lemma ACARRY_MSB: "ALL n. +lemma ACARRY_MSB: "ALL n::nat. RES_FORALL (PWORDLEN n) - (%w1. RES_FORALL (PWORDLEN n) - (%w2. ALL cin. - ACARRY n w1 w2 cin = - bit n (NBWORD (Suc n) (BNVAL w1 + BNVAL w2 + BV cin))))" + (%w1::bool word. + RES_FORALL (PWORDLEN n) + (%w2::bool word. + ALL cin::bool. + ACARRY n w1 w2 cin = + bit n (NBWORD (Suc n) (BNVAL w1 + BNVAL w2 + BV cin))))" by (import bword_arith ACARRY_MSB) -lemma ACARRY_WSEG: "ALL n. +lemma ACARRY_WSEG: "ALL n::nat. RES_FORALL (PWORDLEN n) - (%w1. RES_FORALL (PWORDLEN n) - (%w2. ALL cin k m. - k < m & m <= n --> - ACARRY k (WSEG m 0 w1) (WSEG m 0 w2) cin = - ACARRY k w1 w2 cin))" + (%w1::bool word. + RES_FORALL (PWORDLEN n) + (%w2::bool word. + ALL (cin::bool) (k::nat) m::nat. + k < m & m <= n --> + ACARRY k (WSEG m (0::nat) w1) (WSEG m (0::nat) w2) cin = + ACARRY k w1 w2 cin))" by (import bword_arith ACARRY_WSEG) -lemma ICARRY_WSEG: "ALL n. +lemma ICARRY_WSEG: "ALL n::nat. RES_FORALL (PWORDLEN n) - (%w1. RES_FORALL (PWORDLEN n) - (%w2. ALL cin k m. - k < m & m <= n --> - ICARRY k (WSEG m 0 w1) (WSEG m 0 w2) cin = - ICARRY k w1 w2 cin))" + (%w1::bool word. + RES_FORALL (PWORDLEN n) + (%w2::bool word. + ALL (cin::bool) (k::nat) m::nat. + k < m & m <= n --> + ICARRY k (WSEG m (0::nat) w1) (WSEG m (0::nat) w2) cin = + ICARRY k w1 w2 cin))" by (import bword_arith ICARRY_WSEG) -lemma ACARRY_ACARRY_WSEG: "ALL n. +lemma ACARRY_ACARRY_WSEG: "ALL n::nat. RES_FORALL (PWORDLEN n) - (%w1. RES_FORALL (PWORDLEN n) - (%w2. ALL cin m k1 k2. - k1 < m & k2 < n & m + k2 <= n --> - ACARRY k1 (WSEG m k2 w1) (WSEG m k2 w2) - (ACARRY k2 w1 w2 cin) = - ACARRY (k1 + k2) w1 w2 cin))" + (%w1::bool word. + RES_FORALL (PWORDLEN n) + (%w2::bool word. + ALL (cin::bool) (m::nat) (k1::nat) k2::nat. + k1 < m & k2 < n & m + k2 <= n --> + ACARRY k1 (WSEG m k2 w1) (WSEG m k2 w2) + (ACARRY k2 w1 w2 cin) = + ACARRY (k1 + k2) w1 w2 cin))" by (import bword_arith ACARRY_ACARRY_WSEG) ;end_setup @@ -1377,25 +1640,27 @@ consts WNOT :: "bool word => bool word" -specification (WNOT) WNOT_DEF: "ALL l. WNOT (WORD l) = WORD (map Not l)" +specification (WNOT) WNOT_DEF: "ALL l::bool list. WNOT (WORD l) = WORD (map Not l)" by (import bword_bitop WNOT_DEF) lemma PBITOP_WNOT: "IN WNOT PBITOP" by (import bword_bitop PBITOP_WNOT) -lemma WNOT_WNOT: "ALL w. WNOT (WNOT w) = w" +lemma WNOT_WNOT: "ALL w::bool word. WNOT (WNOT w) = w" by (import bword_bitop WNOT_WNOT) -lemma WCAT_WNOT: "ALL n1 n2. +lemma WCAT_WNOT: "ALL (n1::nat) n2::nat. RES_FORALL (PWORDLEN n1) - (%w1. RES_FORALL (PWORDLEN n2) - (%w2. WCAT (WNOT w1, WNOT w2) = WNOT (WCAT (w1, w2))))" + (%w1::bool word. + RES_FORALL (PWORDLEN n2) + (%w2::bool word. WCAT (WNOT w1, WNOT w2) = WNOT (WCAT (w1, w2))))" by (import bword_bitop WCAT_WNOT) consts WAND :: "bool word => bool word => bool word" -specification (WAND) WAND_DEF: "ALL l1 l2. WAND (WORD l1) (WORD l2) = WORD (map2 op & l1 l2)" +specification (WAND) WAND_DEF: "ALL (l1::bool list) l2::bool list. + WAND (WORD l1) (WORD l2) = WORD (map2 op & l1 l2)" by (import bword_bitop WAND_DEF) lemma PBITBOP_WAND: "IN WAND PBITBOP" @@ -1404,7 +1669,8 @@ consts WOR :: "bool word => bool word => bool word" -specification (WOR) WOR_DEF: "ALL l1 l2. WOR (WORD l1) (WORD l2) = WORD (map2 op | l1 l2)" +specification (WOR) WOR_DEF: "ALL (l1::bool list) l2::bool list. + WOR (WORD l1) (WORD l2) = WORD (map2 op | l1 l2)" by (import bword_bitop WOR_DEF) lemma PBITBOP_WOR: "IN WOR PBITBOP" @@ -1413,7 +1679,8 @@ consts WXOR :: "bool word => bool word => bool word" -specification (WXOR) WXOR_DEF: "ALL l1 l2. WXOR (WORD l1) (WORD l2) = WORD (map2 (%x y. x ~= y) l1 l2)" +specification (WXOR) WXOR_DEF: "ALL (l1::bool list) l2::bool list. + WXOR (WORD l1) (WORD l2) = WORD (map2 (%(x::bool) y::bool. x ~= y) l1 l2)" by (import bword_bitop WXOR_DEF) lemma PBITBOP_WXOR: "IN WXOR PBITBOP" diff -r 7e417a7cbf9f -r bd59bfd4bf37 src/HOL/Import/HOL/HOL4Word32.thy --- a/src/HOL/Import/HOL/HOL4Word32.thy Mon Sep 26 02:06:44 2005 +0200 +++ b/src/HOL/Import/HOL/HOL4Word32.thy Mon Sep 26 02:27:14 2005 +0200 @@ -8,95 +8,122 @@ DIV2 :: "nat => nat" defs - DIV2_primdef: "DIV2 == %n. n div 2" + DIV2_primdef: "DIV2 == %n::nat. n div (2::nat)" -lemma DIV2_def: "ALL n. DIV2 n = n div 2" +lemma DIV2_def: "ALL n::nat. DIV2 n = n div (2::nat)" by (import bits DIV2_def) consts TIMES_2EXP :: "nat => nat => nat" defs - TIMES_2EXP_primdef: "TIMES_2EXP == %x n. n * 2 ^ x" + TIMES_2EXP_primdef: "TIMES_2EXP == %(x::nat) n::nat. n * (2::nat) ^ x" -lemma TIMES_2EXP_def: "ALL x n. TIMES_2EXP x n = n * 2 ^ x" +lemma TIMES_2EXP_def: "ALL (x::nat) n::nat. TIMES_2EXP x n = n * (2::nat) ^ x" by (import bits TIMES_2EXP_def) consts DIV_2EXP :: "nat => nat => nat" defs - DIV_2EXP_primdef: "DIV_2EXP == %x n. n div 2 ^ x" + DIV_2EXP_primdef: "DIV_2EXP == %(x::nat) n::nat. n div (2::nat) ^ x" -lemma DIV_2EXP_def: "ALL x n. DIV_2EXP x n = n div 2 ^ x" +lemma DIV_2EXP_def: "ALL (x::nat) n::nat. DIV_2EXP x n = n div (2::nat) ^ x" by (import bits DIV_2EXP_def) consts MOD_2EXP :: "nat => nat => nat" defs - MOD_2EXP_primdef: "MOD_2EXP == %x n. n mod 2 ^ x" + MOD_2EXP_primdef: "MOD_2EXP == %(x::nat) n::nat. n mod (2::nat) ^ x" -lemma MOD_2EXP_def: "ALL x n. MOD_2EXP x n = n mod 2 ^ x" +lemma MOD_2EXP_def: "ALL (x::nat) n::nat. MOD_2EXP x n = n mod (2::nat) ^ x" by (import bits MOD_2EXP_def) consts DIVMOD_2EXP :: "nat => nat => nat * nat" defs - DIVMOD_2EXP_primdef: "DIVMOD_2EXP == %x n. (n div 2 ^ x, n mod 2 ^ x)" + DIVMOD_2EXP_primdef: "DIVMOD_2EXP == %(x::nat) n::nat. (n div (2::nat) ^ x, n mod (2::nat) ^ x)" -lemma DIVMOD_2EXP_def: "ALL x n. DIVMOD_2EXP x n = (n div 2 ^ x, n mod 2 ^ x)" +lemma DIVMOD_2EXP_def: "ALL (x::nat) n::nat. + DIVMOD_2EXP x n = (n div (2::nat) ^ x, n mod (2::nat) ^ x)" by (import bits DIVMOD_2EXP_def) consts SBIT :: "bool => nat => nat" defs - SBIT_primdef: "SBIT == %b n. if b then 2 ^ n else 0" + SBIT_primdef: "(op ==::(bool => nat => nat) => (bool => nat => nat) => prop) + (SBIT::bool => nat => nat) + (%(b::bool) n::nat. + (If::bool => nat => nat => nat) b + ((op ^::nat => nat => nat) + ((number_of::bin => nat) + ((op BIT::bin => bit => bin) + ((op BIT::bin => bit => bin) (Numeral.Pls::bin) (bit.B1::bit)) + (bit.B0::bit))) + n) + (0::nat))" -lemma SBIT_def: "ALL b n. SBIT b n = (if b then 2 ^ n else 0)" +lemma SBIT_def: "(All::(bool => bool) => bool) + (%b::bool. + (All::(nat => bool) => bool) + (%n::nat. + (op =::nat => nat => bool) ((SBIT::bool => nat => nat) b n) + ((If::bool => nat => nat => nat) b + ((op ^::nat => nat => nat) + ((number_of::bin => nat) + ((op BIT::bin => bit => bin) + ((op BIT::bin => bit => bin) (Numeral.Pls::bin) + (bit.B1::bit)) + (bit.B0::bit))) + n) + (0::nat))))" by (import bits SBIT_def) consts BITS :: "nat => nat => nat => nat" defs - BITS_primdef: "BITS == %h l n. MOD_2EXP (Suc h - l) (DIV_2EXP l n)" + BITS_primdef: "BITS == %(h::nat) (l::nat) n::nat. MOD_2EXP (Suc h - l) (DIV_2EXP l n)" -lemma BITS_def: "ALL h l n. BITS h l n = MOD_2EXP (Suc h - l) (DIV_2EXP l n)" +lemma BITS_def: "ALL (h::nat) (l::nat) n::nat. + BITS h l n = MOD_2EXP (Suc h - l) (DIV_2EXP l n)" by (import bits BITS_def) constdefs bit :: "nat => nat => bool" - "bit == %b n. BITS b b n = 1" + "bit == %(b::nat) n::nat. BITS b b n = (1::nat)" -lemma BIT_def: "ALL b n. bit b n = (BITS b b n = 1)" +lemma BIT_def: "ALL (b::nat) n::nat. bit b n = (BITS b b n = (1::nat))" by (import bits BIT_def) consts SLICE :: "nat => nat => nat => nat" defs - SLICE_primdef: "SLICE == %h l n. MOD_2EXP (Suc h) n - MOD_2EXP l n" + SLICE_primdef: "SLICE == %(h::nat) (l::nat) n::nat. MOD_2EXP (Suc h) n - MOD_2EXP l n" -lemma SLICE_def: "ALL h l n. SLICE h l n = MOD_2EXP (Suc h) n - MOD_2EXP l n" +lemma SLICE_def: "ALL (h::nat) (l::nat) n::nat. + SLICE h l n = MOD_2EXP (Suc h) n - MOD_2EXP l n" by (import bits SLICE_def) consts LSBn :: "nat => bool" defs - LSBn_primdef: "LSBn == bit 0" + LSBn_primdef: "LSBn == bit (0::nat)" -lemma LSBn_def: "LSBn = bit 0" +lemma LSBn_def: "LSBn = bit (0::nat)" by (import bits LSBn_def) consts BITWISE :: "nat => (bool => bool => bool) => nat => nat => nat" -specification (BITWISE_primdef: BITWISE) BITWISE_def: "(ALL oper x y. BITWISE 0 oper x y = 0) & -(ALL n oper x y. +specification (BITWISE_primdef: BITWISE) BITWISE_def: "(ALL (oper::bool => bool => bool) (x::nat) y::nat. + BITWISE (0::nat) oper x y = (0::nat)) & +(ALL (n::nat) (oper::bool => bool => bool) (x::nat) y::nat. BITWISE (Suc n) oper x y = BITWISE n oper x y + SBIT (oper (bit n x) (bit n y)) n)" by (import bits BITWISE_def) @@ -104,7 +131,7 @@ lemma DIV1: "ALL x::nat. x div (1::nat) = x" by (import bits DIV1) -lemma SUC_SUB: "Suc a - a = 1" +lemma SUC_SUB: "Suc (a::nat) - a = (1::nat)" by (import bits SUC_SUB) lemma DIV_MULT_1: "ALL (r::nat) n::nat. r < n --> (n + r) div n = (1::nat)" @@ -129,10 +156,11 @@ lemma EXP_SUB_LESS_EQ: "ALL (a::nat) b::nat. (2::nat) ^ (a - b) <= (2::nat) ^ a" by (import bits EXP_SUB_LESS_EQ) -lemma BITS_THM: "ALL x xa xb. BITS x xa xb = xb div 2 ^ xa mod 2 ^ (Suc x - xa)" +lemma BITS_THM: "ALL (x::nat) (xa::nat) xb::nat. + BITS x xa xb = xb div (2::nat) ^ xa mod (2::nat) ^ (Suc x - xa)" by (import bits BITS_THM) -lemma BITSLT_THM: "ALL h l n. BITS h l n < 2 ^ (Suc h - l)" +lemma BITSLT_THM: "ALL (h::nat) (l::nat) n::nat. BITS h l n < (2::nat) ^ (Suc h - l)" by (import bits BITSLT_THM) lemma DIV_MULT_LEM: "ALL (m::nat) n::nat. (0::nat) < n --> m div n * n <= m" @@ -142,43 +170,46 @@ n mod (2::nat) ^ x = n - n div (2::nat) ^ x * (2::nat) ^ x" by (import bits MOD_2EXP_LEM) -lemma BITS2_THM: "ALL h l n. BITS h l n = n mod 2 ^ Suc h div 2 ^ l" +lemma BITS2_THM: "ALL (h::nat) (l::nat) n::nat. + BITS h l n = n mod (2::nat) ^ Suc h div (2::nat) ^ l" by (import bits BITS2_THM) -lemma BITS_COMP_THM: "ALL h1 l1 h2 l2 n. +lemma BITS_COMP_THM: "ALL (h1::nat) (l1::nat) (h2::nat) (l2::nat) n::nat. h2 + l1 <= h1 --> BITS h2 l2 (BITS h1 l1 n) = BITS (h2 + l1) (l2 + l1) n" by (import bits BITS_COMP_THM) -lemma BITS_DIV_THM: "ALL h l x n. BITS h l x div 2 ^ n = BITS h (l + n) x" +lemma BITS_DIV_THM: "ALL (h::nat) (l::nat) (x::nat) n::nat. + BITS h l x div (2::nat) ^ n = BITS h (l + n) x" by (import bits BITS_DIV_THM) -lemma BITS_LT_HIGH: "ALL h l n. n < 2 ^ Suc h --> BITS h l n = n div 2 ^ l" +lemma BITS_LT_HIGH: "ALL (h::nat) (l::nat) n::nat. + n < (2::nat) ^ Suc h --> BITS h l n = n div (2::nat) ^ l" by (import bits BITS_LT_HIGH) -lemma BITS_ZERO: "ALL h l n. h < l --> BITS h l n = 0" +lemma BITS_ZERO: "ALL (h::nat) (l::nat) n::nat. h < l --> BITS h l n = (0::nat)" by (import bits BITS_ZERO) -lemma BITS_ZERO2: "ALL h l. BITS h l 0 = 0" +lemma BITS_ZERO2: "ALL (h::nat) l::nat. BITS h l (0::nat) = (0::nat)" by (import bits BITS_ZERO2) -lemma BITS_ZERO3: "ALL h x. BITS h 0 x = x mod 2 ^ Suc h" +lemma BITS_ZERO3: "ALL (h::nat) x::nat. BITS h (0::nat) x = x mod (2::nat) ^ Suc h" by (import bits BITS_ZERO3) -lemma BITS_COMP_THM2: "ALL h1 l1 h2 l2 n. +lemma BITS_COMP_THM2: "ALL (h1::nat) (l1::nat) (h2::nat) (l2::nat) n::nat. BITS h2 l2 (BITS h1 l1 n) = BITS (min h1 (h2 + l1)) (l2 + l1) n" by (import bits BITS_COMP_THM2) lemma NOT_MOD2_LEM: "ALL n::nat. (n mod (2::nat) ~= (0::nat)) = (n mod (2::nat) = (1::nat))" by (import bits NOT_MOD2_LEM) -lemma NOT_MOD2_LEM2: "ALL (n::nat) a::'a. +lemma NOT_MOD2_LEM2: "ALL (n::nat) a::'a::type. (n mod (2::nat) ~= (1::nat)) = (n mod (2::nat) = (0::nat))" by (import bits NOT_MOD2_LEM2) -lemma EVEN_MOD2_LEM: "ALL n. EVEN n = (n mod 2 = 0)" +lemma EVEN_MOD2_LEM: "ALL n::nat. EVEN n = (n mod (2::nat) = (0::nat))" by (import bits EVEN_MOD2_LEM) -lemma ODD_MOD2_LEM: "ALL n. ODD n = (n mod 2 = 1)" +lemma ODD_MOD2_LEM: "ALL n::nat. ODD n = (n mod (2::nat) = (1::nat))" by (import bits ODD_MOD2_LEM) lemma LSB_ODD: "LSBn = ODD" @@ -200,7 +231,7 @@ a div (2::nat) ^ x mod (2::nat) ^ y * (2::nat) ^ x" by (import bits SLICE_LEM1) -lemma SLICE_LEM2: "ALL (a::'a) (x::nat) y::nat. +lemma SLICE_LEM2: "ALL (a::'a::type) (x::nat) y::nat. (n::nat) mod (2::nat) ^ (x + y) = n mod (2::nat) ^ x + n div (2::nat) ^ x mod (2::nat) ^ y * (2::nat) ^ x" by (import bits SLICE_LEM2) @@ -209,71 +240,77 @@ l < h --> n mod (2::nat) ^ Suc l <= n mod (2::nat) ^ h" by (import bits SLICE_LEM3) -lemma SLICE_THM: "ALL n h l. SLICE h l n = BITS h l n * 2 ^ l" +lemma SLICE_THM: "ALL (n::nat) (h::nat) l::nat. SLICE h l n = BITS h l n * (2::nat) ^ l" by (import bits SLICE_THM) -lemma SLICELT_THM: "ALL h l n. SLICE h l n < 2 ^ Suc h" +lemma SLICELT_THM: "ALL (h::nat) (l::nat) n::nat. SLICE h l n < (2::nat) ^ Suc h" by (import bits SLICELT_THM) -lemma BITS_SLICE_THM: "ALL h l n. BITS h l (SLICE h l n) = BITS h l n" +lemma BITS_SLICE_THM: "ALL (h::nat) (l::nat) n::nat. BITS h l (SLICE h l n) = BITS h l n" by (import bits BITS_SLICE_THM) -lemma BITS_SLICE_THM2: "ALL h l n. h <= h2 --> BITS h2 l (SLICE h l n) = BITS h l n" +lemma BITS_SLICE_THM2: "ALL (h::nat) (l::nat) n::nat. + h <= (h2::nat) --> BITS h2 l (SLICE h l n) = BITS h l n" by (import bits BITS_SLICE_THM2) lemma MOD_2EXP_MONO: "ALL (n::nat) (h::nat) l::nat. l <= h --> n mod (2::nat) ^ l <= n mod (2::nat) ^ Suc h" by (import bits MOD_2EXP_MONO) -lemma SLICE_COMP_THM: "ALL h m l n. +lemma SLICE_COMP_THM: "ALL (h::nat) (m::nat) (l::nat) n::nat. Suc m <= h & l <= m --> SLICE h (Suc m) n + SLICE m l n = SLICE h l n" by (import bits SLICE_COMP_THM) -lemma SLICE_ZERO: "ALL h l n. h < l --> SLICE h l n = 0" +lemma SLICE_ZERO: "ALL (h::nat) (l::nat) n::nat. h < l --> SLICE h l n = (0::nat)" by (import bits SLICE_ZERO) -lemma BIT_COMP_THM3: "ALL h m l n. +lemma BIT_COMP_THM3: "ALL (h::nat) (m::nat) (l::nat) n::nat. Suc m <= h & l <= m --> - BITS h (Suc m) n * 2 ^ (Suc m - l) + BITS m l n = BITS h l n" + BITS h (Suc m) n * (2::nat) ^ (Suc m - l) + BITS m l n = BITS h l n" by (import bits BIT_COMP_THM3) -lemma NOT_BIT: "ALL n a. (~ bit n a) = (BITS n n a = 0)" +lemma NOT_BIT: "ALL (n::nat) a::nat. (~ bit n a) = (BITS n n a = (0::nat))" by (import bits NOT_BIT) -lemma NOT_BITS: "ALL n a. (BITS n n a ~= 0) = (BITS n n a = 1)" +lemma NOT_BITS: "ALL (n::nat) a::nat. (BITS n n a ~= (0::nat)) = (BITS n n a = (1::nat))" by (import bits NOT_BITS) -lemma NOT_BITS2: "ALL n a. (BITS n n a ~= 1) = (BITS n n a = 0)" +lemma NOT_BITS2: "ALL (n::nat) a::nat. (BITS n n a ~= (1::nat)) = (BITS n n a = (0::nat))" by (import bits NOT_BITS2) -lemma BIT_SLICE: "ALL n a b. (bit n a = bit n b) = (SLICE n n a = SLICE n n b)" +lemma BIT_SLICE: "ALL (n::nat) (a::nat) b::nat. + (bit n a = bit n b) = (SLICE n n a = SLICE n n b)" by (import bits BIT_SLICE) -lemma BIT_SLICE_LEM: "ALL y x n. SBIT (bit x n) (x + y) = SLICE x x n * 2 ^ y" +lemma BIT_SLICE_LEM: "ALL (y::nat) (x::nat) n::nat. + SBIT (bit x n) (x + y) = SLICE x x n * (2::nat) ^ y" by (import bits BIT_SLICE_LEM) -lemma BIT_SLICE_THM: "ALL x xa. SBIT (bit x xa) x = SLICE x x xa" +lemma BIT_SLICE_THM: "ALL (x::nat) xa::nat. SBIT (bit x xa) x = SLICE x x xa" by (import bits BIT_SLICE_THM) -lemma SBIT_DIV: "ALL b m n. n < m --> SBIT b (m - n) = SBIT b m div 2 ^ n" +lemma SBIT_DIV: "ALL (b::bool) (m::nat) n::nat. + n < m --> SBIT b (m - n) = SBIT b m div (2::nat) ^ n" by (import bits SBIT_DIV) -lemma BITS_SUC: "ALL h l n. +lemma BITS_SUC: "ALL (h::nat) (l::nat) n::nat. l <= Suc h --> SBIT (bit (Suc h) n) (Suc h - l) + BITS h l n = BITS (Suc h) l n" by (import bits BITS_SUC) -lemma BITS_SUC_THM: "ALL h l n. +lemma BITS_SUC_THM: "ALL (h::nat) (l::nat) n::nat. BITS (Suc h) l n = - (if Suc h < l then 0 else SBIT (bit (Suc h) n) (Suc h - l) + BITS h l n)" + (if Suc h < l then 0::nat + else SBIT (bit (Suc h) n) (Suc h - l) + BITS h l n)" by (import bits BITS_SUC_THM) -lemma BIT_BITS_THM: "ALL h l a b. - (ALL x. l <= x & x <= h --> bit x a = bit x b) = +lemma BIT_BITS_THM: "ALL (h::nat) (l::nat) (a::nat) b::nat. + (ALL x::nat. l <= x & x <= h --> bit x a = bit x b) = (BITS h l a = BITS h l b)" by (import bits BIT_BITS_THM) -lemma BITWISE_LT_2EXP: "ALL n oper a b. BITWISE n oper a b < 2 ^ n" +lemma BITWISE_LT_2EXP: "ALL (n::nat) (oper::bool => bool => bool) (a::nat) b::nat. + BITWISE n oper a b < (2::nat) ^ n" by (import bits BITWISE_LT_2EXP) lemma LESS_EXP_MULT2: "ALL (a::nat) b::nat. @@ -281,18 +318,20 @@ (EX x::nat. (2::nat) ^ b = (2::nat) ^ (x + (1::nat)) * (2::nat) ^ a)" by (import bits LESS_EXP_MULT2) -lemma BITWISE_THM: "ALL x n oper a b. +lemma BITWISE_THM: "ALL (x::nat) (n::nat) (oper::bool => bool => bool) (a::nat) b::nat. x < n --> bit x (BITWISE n oper a b) = oper (bit x a) (bit x b)" by (import bits BITWISE_THM) -lemma BITWISE_COR: "ALL x n oper a b. +lemma BITWISE_COR: "ALL (x::nat) (n::nat) (oper::bool => bool => bool) (a::nat) b::nat. x < n --> - oper (bit x a) (bit x b) --> BITWISE n oper a b div 2 ^ x mod 2 = 1" + oper (bit x a) (bit x b) --> + BITWISE n oper a b div (2::nat) ^ x mod (2::nat) = (1::nat)" by (import bits BITWISE_COR) -lemma BITWISE_NOT_COR: "ALL x n oper a b. +lemma BITWISE_NOT_COR: "ALL (x::nat) (n::nat) (oper::bool => bool => bool) (a::nat) b::nat. x < n --> - ~ oper (bit x a) (bit x b) --> BITWISE n oper a b div 2 ^ x mod 2 = 0" + ~ oper (bit x a) (bit x b) --> + BITWISE n oper a b div (2::nat) ^ x mod (2::nat) = (0::nat)" by (import bits BITWISE_NOT_COR) lemma MOD_PLUS_RIGHT: "ALL n>0::nat. ALL (j::nat) k::nat. (j + k mod n) mod n = (j + k) mod n" @@ -340,125 +379,131 @@ MODw :: "nat => nat" defs - MODw_primdef: "MODw == %n. n mod 2 ^ WL" + MODw_primdef: "MODw == %n::nat. n mod (2::nat) ^ WL" -lemma MODw_def: "ALL n. MODw n = n mod 2 ^ WL" +lemma MODw_def: "ALL n::nat. MODw n = n mod (2::nat) ^ WL" by (import word32 MODw_def) consts INw :: "nat => bool" defs - INw_primdef: "INw == %n. n < 2 ^ WL" + INw_primdef: "INw == %n::nat. n < (2::nat) ^ WL" -lemma INw_def: "ALL n. INw n = (n < 2 ^ WL)" +lemma INw_def: "ALL n::nat. INw n = (n < (2::nat) ^ WL)" by (import word32 INw_def) consts EQUIV :: "nat => nat => bool" defs - EQUIV_primdef: "EQUIV == %x y. MODw x = MODw y" + EQUIV_primdef: "EQUIV == %(x::nat) y::nat. MODw x = MODw y" -lemma EQUIV_def: "ALL x y. EQUIV x y = (MODw x = MODw y)" +lemma EQUIV_def: "ALL (x::nat) y::nat. EQUIV x y = (MODw x = MODw y)" by (import word32 EQUIV_def) -lemma EQUIV_QT: "ALL x y. EQUIV x y = (EQUIV x = EQUIV y)" +lemma EQUIV_QT: "ALL (x::nat) y::nat. EQUIV x y = (EQUIV x = EQUIV y)" by (import word32 EQUIV_QT) -lemma FUNPOW_THM: "ALL f n x. (f ^ n) (f x) = f ((f ^ n) x)" +lemma FUNPOW_THM: "ALL (f::'a::type => 'a::type) (n::nat) x::'a::type. + (f ^ n) (f x) = f ((f ^ n) x)" by (import word32 FUNPOW_THM) -lemma FUNPOW_THM2: "ALL f n x. (f ^ Suc n) x = f ((f ^ n) x)" +lemma FUNPOW_THM2: "ALL (f::'a::type => 'a::type) (n::nat) x::'a::type. + (f ^ Suc n) x = f ((f ^ n) x)" by (import word32 FUNPOW_THM2) -lemma FUNPOW_COMP: "ALL f m n a. (f ^ m) ((f ^ n) a) = (f ^ (m + n)) a" +lemma FUNPOW_COMP: "ALL (f::'a::type => 'a::type) (m::nat) (n::nat) a::'a::type. + (f ^ m) ((f ^ n) a) = (f ^ (m + n)) a" by (import word32 FUNPOW_COMP) -lemma INw_MODw: "ALL n. INw (MODw n)" +lemma INw_MODw: "ALL n::nat. INw (MODw n)" by (import word32 INw_MODw) -lemma TOw_IDEM: "ALL a. INw a --> MODw a = a" +lemma TOw_IDEM: "ALL a::nat. INw a --> MODw a = a" by (import word32 TOw_IDEM) -lemma MODw_IDEM2: "ALL a. MODw (MODw a) = MODw a" +lemma MODw_IDEM2: "ALL a::nat. MODw (MODw a) = MODw a" by (import word32 MODw_IDEM2) -lemma TOw_QT: "ALL a. EQUIV (MODw a) a" +lemma TOw_QT: "ALL a::nat. EQUIV (MODw a) a" by (import word32 TOw_QT) -lemma MODw_THM: "MODw = BITS HB 0" +lemma MODw_THM: "MODw = BITS HB (0::nat)" by (import word32 MODw_THM) -lemma MOD_ADD: "ALL a b. MODw (a + b) = MODw (MODw a + MODw b)" +lemma MOD_ADD: "ALL (a::nat) b::nat. MODw (a + b) = MODw (MODw a + MODw b)" by (import word32 MOD_ADD) -lemma MODw_MULT: "ALL a b. MODw (a * b) = MODw (MODw a * MODw b)" +lemma MODw_MULT: "ALL (a::nat) b::nat. MODw (a * b) = MODw (MODw a * MODw b)" by (import word32 MODw_MULT) consts AONE :: "nat" defs - AONE_primdef: "AONE == 1" + AONE_primdef: "AONE == 1::nat" -lemma AONE_def: "AONE = 1" +lemma AONE_def: "AONE = (1::nat)" by (import word32 AONE_def) -lemma ADD_QT: "(ALL n. EQUIV (0 + n) n) & (ALL m n. EQUIV (Suc m + n) (Suc (m + n)))" +lemma ADD_QT: "(ALL n::nat. EQUIV ((0::nat) + n) n) & +(ALL (m::nat) n::nat. EQUIV (Suc m + n) (Suc (m + n)))" by (import word32 ADD_QT) -lemma ADD_0_QT: "ALL a. EQUIV (a + 0) a" +lemma ADD_0_QT: "ALL a::nat. EQUIV (a + (0::nat)) a" by (import word32 ADD_0_QT) -lemma ADD_COMM_QT: "ALL a b. EQUIV (a + b) (b + a)" +lemma ADD_COMM_QT: "ALL (a::nat) b::nat. EQUIV (a + b) (b + a)" by (import word32 ADD_COMM_QT) -lemma ADD_ASSOC_QT: "ALL a b c. EQUIV (a + (b + c)) (a + b + c)" +lemma ADD_ASSOC_QT: "ALL (a::nat) (b::nat) c::nat. EQUIV (a + (b + c)) (a + b + c)" by (import word32 ADD_ASSOC_QT) -lemma MULT_QT: "(ALL n. EQUIV (0 * n) 0) & (ALL m n. EQUIV (Suc m * n) (m * n + n))" +lemma MULT_QT: "(ALL n::nat. EQUIV ((0::nat) * n) (0::nat)) & +(ALL (m::nat) n::nat. EQUIV (Suc m * n) (m * n + n))" by (import word32 MULT_QT) -lemma ADD1_QT: "ALL m. EQUIV (Suc m) (m + AONE)" +lemma ADD1_QT: "ALL m::nat. EQUIV (Suc m) (m + AONE)" by (import word32 ADD1_QT) -lemma ADD_CLAUSES_QT: "(ALL m. EQUIV (0 + m) m) & -(ALL m. EQUIV (m + 0) m) & -(ALL m n. EQUIV (Suc m + n) (Suc (m + n))) & -(ALL m n. EQUIV (m + Suc n) (Suc (m + n)))" +lemma ADD_CLAUSES_QT: "(ALL m::nat. EQUIV ((0::nat) + m) m) & +(ALL m::nat. EQUIV (m + (0::nat)) m) & +(ALL (m::nat) n::nat. EQUIV (Suc m + n) (Suc (m + n))) & +(ALL (m::nat) n::nat. EQUIV (m + Suc n) (Suc (m + n)))" by (import word32 ADD_CLAUSES_QT) -lemma SUC_EQUIV_COMP: "ALL a b. EQUIV (Suc a) b --> EQUIV a (b + (2 ^ WL - 1))" +lemma SUC_EQUIV_COMP: "ALL (a::nat) b::nat. + EQUIV (Suc a) b --> EQUIV a (b + ((2::nat) ^ WL - (1::nat)))" by (import word32 SUC_EQUIV_COMP) -lemma INV_SUC_EQ_QT: "ALL m n. EQUIV (Suc m) (Suc n) = EQUIV m n" +lemma INV_SUC_EQ_QT: "ALL (m::nat) n::nat. EQUIV (Suc m) (Suc n) = EQUIV m n" by (import word32 INV_SUC_EQ_QT) -lemma ADD_INV_0_QT: "ALL m n. EQUIV (m + n) m --> EQUIV n 0" +lemma ADD_INV_0_QT: "ALL (m::nat) n::nat. EQUIV (m + n) m --> EQUIV n (0::nat)" by (import word32 ADD_INV_0_QT) -lemma ADD_INV_0_EQ_QT: "ALL m n. EQUIV (m + n) m = EQUIV n 0" +lemma ADD_INV_0_EQ_QT: "ALL (m::nat) n::nat. EQUIV (m + n) m = EQUIV n (0::nat)" by (import word32 ADD_INV_0_EQ_QT) -lemma EQ_ADD_LCANCEL_QT: "ALL m n p. EQUIV (m + n) (m + p) = EQUIV n p" +lemma EQ_ADD_LCANCEL_QT: "ALL (m::nat) (n::nat) p::nat. EQUIV (m + n) (m + p) = EQUIV n p" by (import word32 EQ_ADD_LCANCEL_QT) -lemma EQ_ADD_RCANCEL_QT: "ALL x xa xb. EQUIV (x + xb) (xa + xb) = EQUIV x xa" +lemma EQ_ADD_RCANCEL_QT: "ALL (x::nat) (xa::nat) xb::nat. EQUIV (x + xb) (xa + xb) = EQUIV x xa" by (import word32 EQ_ADD_RCANCEL_QT) -lemma LEFT_ADD_DISTRIB_QT: "ALL m n p. EQUIV (p * (m + n)) (p * m + p * n)" +lemma LEFT_ADD_DISTRIB_QT: "ALL (m::nat) (n::nat) p::nat. EQUIV (p * (m + n)) (p * m + p * n)" by (import word32 LEFT_ADD_DISTRIB_QT) -lemma MULT_ASSOC_QT: "ALL m n p. EQUIV (m * (n * p)) (m * n * p)" +lemma MULT_ASSOC_QT: "ALL (m::nat) (n::nat) p::nat. EQUIV (m * (n * p)) (m * n * p)" by (import word32 MULT_ASSOC_QT) -lemma MULT_COMM_QT: "ALL m n. EQUIV (m * n) (n * m)" +lemma MULT_COMM_QT: "ALL (m::nat) n::nat. EQUIV (m * n) (n * m)" by (import word32 MULT_COMM_QT) -lemma MULT_CLAUSES_QT: "ALL m n. - EQUIV (0 * m) 0 & - EQUIV (m * 0) 0 & +lemma MULT_CLAUSES_QT: "ALL (m::nat) n::nat. + EQUIV ((0::nat) * m) (0::nat) & + EQUIV (m * (0::nat)) (0::nat) & EQUIV (AONE * m) m & EQUIV (m * AONE) m & EQUIV (Suc m * n) (m * n + n) & EQUIV (m * Suc n) (m + m * n)" @@ -477,24 +522,24 @@ ONE_COMP :: "nat => nat" defs - ONE_COMP_primdef: "ONE_COMP == %x. 2 ^ WL - 1 - MODw x" + ONE_COMP_primdef: "ONE_COMP == %x::nat. (2::nat) ^ WL - (1::nat) - MODw x" -lemma ONE_COMP_def: "ALL x. ONE_COMP x = 2 ^ WL - 1 - MODw x" +lemma ONE_COMP_def: "ALL x::nat. ONE_COMP x = (2::nat) ^ WL - (1::nat) - MODw x" by (import word32 ONE_COMP_def) consts TWO_COMP :: "nat => nat" defs - TWO_COMP_primdef: "TWO_COMP == %x. 2 ^ WL - MODw x" + TWO_COMP_primdef: "TWO_COMP == %x::nat. (2::nat) ^ WL - MODw x" -lemma TWO_COMP_def: "ALL x. TWO_COMP x = 2 ^ WL - MODw x" +lemma TWO_COMP_def: "ALL x::nat. TWO_COMP x = (2::nat) ^ WL - MODw x" by (import word32 TWO_COMP_def) -lemma ADD_TWO_COMP_QT: "ALL a. EQUIV (MODw a + TWO_COMP a) 0" +lemma ADD_TWO_COMP_QT: "ALL a::nat. EQUIV (MODw a + TWO_COMP a) (0::nat)" by (import word32 ADD_TWO_COMP_QT) -lemma TWO_COMP_ONE_COMP_QT: "ALL a. EQUIV (TWO_COMP a) (ONE_COMP a + AONE)" +lemma TWO_COMP_ONE_COMP_QT: "ALL a::nat. EQUIV (TWO_COMP a) (ONE_COMP a + AONE)" by (import word32 TWO_COMP_ONE_COMP_QT) lemma BIT_EQUIV_THM: "(All::(nat => bool) => bool) @@ -512,13 +557,14 @@ ((EQUIV::nat => nat => bool) x xa)))" by (import word32 BIT_EQUIV_THM) -lemma BITS_SUC2: "ALL n a. BITS (Suc n) 0 a = SLICE (Suc n) (Suc n) a + BITS n 0 a" +lemma BITS_SUC2: "ALL (n::nat) a::nat. + BITS (Suc n) (0::nat) a = SLICE (Suc n) (Suc n) a + BITS n (0::nat) a" by (import word32 BITS_SUC2) -lemma BITWISE_ONE_COMP_THM: "ALL a b. BITWISE WL (%x y. ~ x) a b = ONE_COMP a" +lemma BITWISE_ONE_COMP_THM: "ALL (a::nat) b::nat. BITWISE WL (%(x::bool) y::bool. ~ x) a b = ONE_COMP a" by (import word32 BITWISE_ONE_COMP_THM) -lemma ONE_COMP_THM: "ALL x xa. xa < WL --> bit xa (ONE_COMP x) = (~ bit xa x)" +lemma ONE_COMP_THM: "ALL (x::nat) xa::nat. xa < WL --> bit xa (ONE_COMP x) = (~ bit xa x)" by (import word32 ONE_COMP_THM) consts @@ -543,18 +589,18 @@ EOR :: "nat => nat => nat" defs - EOR_primdef: "EOR == BITWISE WL (%x y. x ~= y)" + EOR_primdef: "EOR == BITWISE WL (%(x::bool) y::bool. x ~= y)" -lemma EOR_def: "EOR = BITWISE WL (%x y. x ~= y)" +lemma EOR_def: "EOR = BITWISE WL (%(x::bool) y::bool. x ~= y)" by (import word32 EOR_def) consts COMP0 :: "nat" defs - COMP0_primdef: "COMP0 == ONE_COMP 0" + COMP0_primdef: "COMP0 == ONE_COMP (0::nat)" -lemma COMP0_def: "COMP0 = ONE_COMP 0" +lemma COMP0_def: "COMP0 = ONE_COMP (0::nat)" by (import word32 COMP0_def) lemma BITWISE_THM2: "(All::(nat => bool) => bool) @@ -582,149 +628,154 @@ y)))))" by (import word32 BITWISE_THM2) -lemma OR_ASSOC_QT: "ALL a b c. EQUIV (OR a (OR b c)) (OR (OR a b) c)" +lemma OR_ASSOC_QT: "ALL (a::nat) (b::nat) c::nat. EQUIV (OR a (OR b c)) (OR (OR a b) c)" by (import word32 OR_ASSOC_QT) -lemma OR_COMM_QT: "ALL a b. EQUIV (OR a b) (OR b a)" +lemma OR_COMM_QT: "ALL (a::nat) b::nat. EQUIV (OR a b) (OR b a)" by (import word32 OR_COMM_QT) -lemma OR_ABSORB_QT: "ALL a b. EQUIV (AND a (OR a b)) a" +lemma OR_ABSORB_QT: "ALL (a::nat) b::nat. EQUIV (AND a (OR a b)) a" by (import word32 OR_ABSORB_QT) -lemma OR_IDEM_QT: "ALL a. EQUIV (OR a a) a" +lemma OR_IDEM_QT: "ALL a::nat. EQUIV (OR a a) a" by (import word32 OR_IDEM_QT) -lemma AND_ASSOC_QT: "ALL a b c. EQUIV (AND a (AND b c)) (AND (AND a b) c)" +lemma AND_ASSOC_QT: "ALL (a::nat) (b::nat) c::nat. EQUIV (AND a (AND b c)) (AND (AND a b) c)" by (import word32 AND_ASSOC_QT) -lemma AND_COMM_QT: "ALL a b. EQUIV (AND a b) (AND b a)" +lemma AND_COMM_QT: "ALL (a::nat) b::nat. EQUIV (AND a b) (AND b a)" by (import word32 AND_COMM_QT) -lemma AND_ABSORB_QT: "ALL a b. EQUIV (OR a (AND a b)) a" +lemma AND_ABSORB_QT: "ALL (a::nat) b::nat. EQUIV (OR a (AND a b)) a" by (import word32 AND_ABSORB_QT) -lemma AND_IDEM_QT: "ALL a. EQUIV (AND a a) a" +lemma AND_IDEM_QT: "ALL a::nat. EQUIV (AND a a) a" by (import word32 AND_IDEM_QT) -lemma OR_COMP_QT: "ALL a. EQUIV (OR a (ONE_COMP a)) COMP0" +lemma OR_COMP_QT: "ALL a::nat. EQUIV (OR a (ONE_COMP a)) COMP0" by (import word32 OR_COMP_QT) -lemma AND_COMP_QT: "ALL a. EQUIV (AND a (ONE_COMP a)) 0" +lemma AND_COMP_QT: "ALL a::nat. EQUIV (AND a (ONE_COMP a)) (0::nat)" by (import word32 AND_COMP_QT) -lemma ONE_COMP_QT: "ALL a. EQUIV (ONE_COMP (ONE_COMP a)) a" +lemma ONE_COMP_QT: "ALL a::nat. EQUIV (ONE_COMP (ONE_COMP a)) a" by (import word32 ONE_COMP_QT) -lemma RIGHT_AND_OVER_OR_QT: "ALL a b c. EQUIV (AND (OR a b) c) (OR (AND a c) (AND b c))" +lemma RIGHT_AND_OVER_OR_QT: "ALL (a::nat) (b::nat) c::nat. + EQUIV (AND (OR a b) c) (OR (AND a c) (AND b c))" by (import word32 RIGHT_AND_OVER_OR_QT) -lemma RIGHT_OR_OVER_AND_QT: "ALL a b c. EQUIV (OR (AND a b) c) (AND (OR a c) (OR b c))" +lemma RIGHT_OR_OVER_AND_QT: "ALL (a::nat) (b::nat) c::nat. EQUIV (OR (AND a b) c) (AND (OR a c) (OR b c))" by (import word32 RIGHT_OR_OVER_AND_QT) -lemma DE_MORGAN_THM_QT: "ALL a b. +lemma DE_MORGAN_THM_QT: "ALL (a::nat) b::nat. EQUIV (ONE_COMP (AND a b)) (OR (ONE_COMP a) (ONE_COMP b)) & EQUIV (ONE_COMP (OR a b)) (AND (ONE_COMP a) (ONE_COMP b))" by (import word32 DE_MORGAN_THM_QT) -lemma BIT_EQUIV: "ALL n a b. n < WL --> EQUIV a b --> bit n a = bit n b" +lemma BIT_EQUIV: "ALL (n::nat) (a::nat) b::nat. n < WL --> EQUIV a b --> bit n a = bit n b" by (import word32 BIT_EQUIV) -lemma LSB_WELLDEF: "ALL a b. EQUIV a b --> LSBn a = LSBn b" +lemma LSB_WELLDEF: "ALL (a::nat) b::nat. EQUIV a b --> LSBn a = LSBn b" by (import word32 LSB_WELLDEF) -lemma MSB_WELLDEF: "ALL a b. EQUIV a b --> MSBn a = MSBn b" +lemma MSB_WELLDEF: "ALL (a::nat) b::nat. EQUIV a b --> MSBn a = MSBn b" by (import word32 MSB_WELLDEF) -lemma BITWISE_ISTEP: "ALL n oper a b. - 0 < n --> - BITWISE n oper (a div 2) (b div 2) = - BITWISE n oper a b div 2 + SBIT (oper (bit n a) (bit n b)) (n - 1)" +lemma BITWISE_ISTEP: "ALL (n::nat) (oper::bool => bool => bool) (a::nat) b::nat. + (0::nat) < n --> + BITWISE n oper (a div (2::nat)) (b div (2::nat)) = + BITWISE n oper a b div (2::nat) + + SBIT (oper (bit n a) (bit n b)) (n - (1::nat))" by (import word32 BITWISE_ISTEP) -lemma BITWISE_EVAL: "ALL n oper a b. +lemma BITWISE_EVAL: "ALL (n::nat) (oper::bool => bool => bool) (a::nat) b::nat. BITWISE (Suc n) oper a b = - 2 * BITWISE n oper (a div 2) (b div 2) + SBIT (oper (LSBn a) (LSBn b)) 0" + (2::nat) * BITWISE n oper (a div (2::nat)) (b div (2::nat)) + + SBIT (oper (LSBn a) (LSBn b)) (0::nat)" by (import word32 BITWISE_EVAL) -lemma BITWISE_WELLDEF: "ALL n oper a b c d. +lemma BITWISE_WELLDEF: "ALL (n::nat) (oper::bool => bool => bool) (a::nat) (b::nat) (c::nat) d::nat. EQUIV a b & EQUIV c d --> EQUIV (BITWISE n oper a c) (BITWISE n oper b d)" by (import word32 BITWISE_WELLDEF) -lemma BITWISEw_WELLDEF: "ALL oper a b c d. +lemma BITWISEw_WELLDEF: "ALL (oper::bool => bool => bool) (a::nat) (b::nat) (c::nat) d::nat. EQUIV a b & EQUIV c d --> EQUIV (BITWISE WL oper a c) (BITWISE WL oper b d)" by (import word32 BITWISEw_WELLDEF) -lemma SUC_WELLDEF: "ALL a b. EQUIV a b --> EQUIV (Suc a) (Suc b)" +lemma SUC_WELLDEF: "ALL (a::nat) b::nat. EQUIV a b --> EQUIV (Suc a) (Suc b)" by (import word32 SUC_WELLDEF) -lemma ADD_WELLDEF: "ALL a b c d. EQUIV a b & EQUIV c d --> EQUIV (a + c) (b + d)" +lemma ADD_WELLDEF: "ALL (a::nat) (b::nat) (c::nat) d::nat. + EQUIV a b & EQUIV c d --> EQUIV (a + c) (b + d)" by (import word32 ADD_WELLDEF) -lemma MUL_WELLDEF: "ALL a b c d. EQUIV a b & EQUIV c d --> EQUIV (a * c) (b * d)" +lemma MUL_WELLDEF: "ALL (a::nat) (b::nat) (c::nat) d::nat. + EQUIV a b & EQUIV c d --> EQUIV (a * c) (b * d)" by (import word32 MUL_WELLDEF) -lemma ONE_COMP_WELLDEF: "ALL a b. EQUIV a b --> EQUIV (ONE_COMP a) (ONE_COMP b)" +lemma ONE_COMP_WELLDEF: "ALL (a::nat) b::nat. EQUIV a b --> EQUIV (ONE_COMP a) (ONE_COMP b)" by (import word32 ONE_COMP_WELLDEF) -lemma TWO_COMP_WELLDEF: "ALL a b. EQUIV a b --> EQUIV (TWO_COMP a) (TWO_COMP b)" +lemma TWO_COMP_WELLDEF: "ALL (a::nat) b::nat. EQUIV a b --> EQUIV (TWO_COMP a) (TWO_COMP b)" by (import word32 TWO_COMP_WELLDEF) -lemma TOw_WELLDEF: "ALL a b. EQUIV a b --> EQUIV (MODw a) (MODw b)" +lemma TOw_WELLDEF: "ALL (a::nat) b::nat. EQUIV a b --> EQUIV (MODw a) (MODw b)" by (import word32 TOw_WELLDEF) consts LSR_ONE :: "nat => nat" defs - LSR_ONE_primdef: "LSR_ONE == %a. MODw a div 2" + LSR_ONE_primdef: "LSR_ONE == %a::nat. MODw a div (2::nat)" -lemma LSR_ONE_def: "ALL a. LSR_ONE a = MODw a div 2" +lemma LSR_ONE_def: "ALL a::nat. LSR_ONE a = MODw a div (2::nat)" by (import word32 LSR_ONE_def) consts ASR_ONE :: "nat => nat" defs - ASR_ONE_primdef: "ASR_ONE == %a. LSR_ONE a + SBIT (MSBn a) HB" + ASR_ONE_primdef: "ASR_ONE == %a::nat. LSR_ONE a + SBIT (MSBn a) HB" -lemma ASR_ONE_def: "ALL a. ASR_ONE a = LSR_ONE a + SBIT (MSBn a) HB" +lemma ASR_ONE_def: "ALL a::nat. ASR_ONE a = LSR_ONE a + SBIT (MSBn a) HB" by (import word32 ASR_ONE_def) consts ROR_ONE :: "nat => nat" defs - ROR_ONE_primdef: "ROR_ONE == %a. LSR_ONE a + SBIT (LSBn a) HB" + ROR_ONE_primdef: "ROR_ONE == %a::nat. LSR_ONE a + SBIT (LSBn a) HB" -lemma ROR_ONE_def: "ALL a. ROR_ONE a = LSR_ONE a + SBIT (LSBn a) HB" +lemma ROR_ONE_def: "ALL a::nat. ROR_ONE a = LSR_ONE a + SBIT (LSBn a) HB" by (import word32 ROR_ONE_def) consts RRXn :: "bool => nat => nat" defs - RRXn_primdef: "RRXn == %c a. LSR_ONE a + SBIT c HB" + RRXn_primdef: "RRXn == %(c::bool) a::nat. LSR_ONE a + SBIT c HB" -lemma RRXn_def: "ALL c a. RRXn c a = LSR_ONE a + SBIT c HB" +lemma RRXn_def: "ALL (c::bool) a::nat. RRXn c a = LSR_ONE a + SBIT c HB" by (import word32 RRXn_def) -lemma LSR_ONE_WELLDEF: "ALL a b. EQUIV a b --> EQUIV (LSR_ONE a) (LSR_ONE b)" +lemma LSR_ONE_WELLDEF: "ALL (a::nat) b::nat. EQUIV a b --> EQUIV (LSR_ONE a) (LSR_ONE b)" by (import word32 LSR_ONE_WELLDEF) -lemma ASR_ONE_WELLDEF: "ALL a b. EQUIV a b --> EQUIV (ASR_ONE a) (ASR_ONE b)" +lemma ASR_ONE_WELLDEF: "ALL (a::nat) b::nat. EQUIV a b --> EQUIV (ASR_ONE a) (ASR_ONE b)" by (import word32 ASR_ONE_WELLDEF) -lemma ROR_ONE_WELLDEF: "ALL a b. EQUIV a b --> EQUIV (ROR_ONE a) (ROR_ONE b)" +lemma ROR_ONE_WELLDEF: "ALL (a::nat) b::nat. EQUIV a b --> EQUIV (ROR_ONE a) (ROR_ONE b)" by (import word32 ROR_ONE_WELLDEF) -lemma RRX_WELLDEF: "ALL a b c. EQUIV a b --> EQUIV (RRXn c a) (RRXn c b)" +lemma RRX_WELLDEF: "ALL (a::nat) (b::nat) c::bool. EQUIV a b --> EQUIV (RRXn c a) (RRXn c b)" by (import word32 RRX_WELLDEF) -lemma LSR_ONE: "LSR_ONE = BITS HB 1" +lemma LSR_ONE: "LSR_ONE = BITS HB (1::nat)" by (import word32 LSR_ONE) -typedef (open) word32 = "{x. EX xa. x = EQUIV xa}" +typedef (open) word32 = "{x::nat => bool. EX xa::nat. x = EQUIV xa}" by (rule typedef_helper,import word32 word32_TY_DEF) lemmas word32_TY_DEF = typedef_hol2hol4 [OF type_definition_word32] @@ -733,17 +784,18 @@ mk_word32 :: "(nat => bool) => word32" dest_word32 :: "word32 => nat => bool" -specification (dest_word32 mk_word32) word32_tybij: "(ALL a. mk_word32 (dest_word32 a) = a) & -(ALL r. (EX x. r = EQUIV x) = (dest_word32 (mk_word32 r) = r))" +specification (dest_word32 mk_word32) word32_tybij: "(ALL a::word32. mk_word32 (dest_word32 a) = a) & +(ALL r::nat => bool. + (EX x::nat. r = EQUIV x) = (dest_word32 (mk_word32 r) = r))" by (import word32 word32_tybij) consts w_0 :: "word32" defs - w_0_primdef: "w_0 == mk_word32 (EQUIV 0)" + w_0_primdef: "w_0 == mk_word32 (EQUIV (0::nat))" -lemma w_0_def: "w_0 = mk_word32 (EQUIV 0)" +lemma w_0_def: "w_0 = mk_word32 (EQUIV (0::nat))" by (import word32 w_0_def) consts @@ -766,17 +818,18 @@ constdefs word_suc :: "word32 => word32" - "word_suc == %T1. mk_word32 (EQUIV (Suc (Eps (dest_word32 T1))))" + "word_suc == %T1::word32. mk_word32 (EQUIV (Suc (Eps (dest_word32 T1))))" -lemma word_suc: "ALL T1. word_suc T1 = mk_word32 (EQUIV (Suc (Eps (dest_word32 T1))))" +lemma word_suc: "ALL T1::word32. word_suc T1 = mk_word32 (EQUIV (Suc (Eps (dest_word32 T1))))" by (import word32 word_suc) constdefs word_add :: "word32 => word32 => word32" "word_add == -%T1 T2. mk_word32 (EQUIV (Eps (dest_word32 T1) + Eps (dest_word32 T2)))" +%(T1::word32) T2::word32. + mk_word32 (EQUIV (Eps (dest_word32 T1) + Eps (dest_word32 T2)))" -lemma word_add: "ALL T1 T2. +lemma word_add: "ALL (T1::word32) T2::word32. word_add T1 T2 = mk_word32 (EQUIV (Eps (dest_word32 T1) + Eps (dest_word32 T2)))" by (import word32 word_add) @@ -784,81 +837,92 @@ constdefs word_mul :: "word32 => word32 => word32" "word_mul == -%T1 T2. mk_word32 (EQUIV (Eps (dest_word32 T1) * Eps (dest_word32 T2)))" +%(T1::word32) T2::word32. + mk_word32 (EQUIV (Eps (dest_word32 T1) * Eps (dest_word32 T2)))" -lemma word_mul: "ALL T1 T2. +lemma word_mul: "ALL (T1::word32) T2::word32. word_mul T1 T2 = mk_word32 (EQUIV (Eps (dest_word32 T1) * Eps (dest_word32 T2)))" by (import word32 word_mul) constdefs word_1comp :: "word32 => word32" - "word_1comp == %T1. mk_word32 (EQUIV (ONE_COMP (Eps (dest_word32 T1))))" + "word_1comp == +%T1::word32. mk_word32 (EQUIV (ONE_COMP (Eps (dest_word32 T1))))" -lemma word_1comp: "ALL T1. word_1comp T1 = mk_word32 (EQUIV (ONE_COMP (Eps (dest_word32 T1))))" +lemma word_1comp: "ALL T1::word32. + word_1comp T1 = mk_word32 (EQUIV (ONE_COMP (Eps (dest_word32 T1))))" by (import word32 word_1comp) constdefs word_2comp :: "word32 => word32" - "word_2comp == %T1. mk_word32 (EQUIV (TWO_COMP (Eps (dest_word32 T1))))" + "word_2comp == +%T1::word32. mk_word32 (EQUIV (TWO_COMP (Eps (dest_word32 T1))))" -lemma word_2comp: "ALL T1. word_2comp T1 = mk_word32 (EQUIV (TWO_COMP (Eps (dest_word32 T1))))" +lemma word_2comp: "ALL T1::word32. + word_2comp T1 = mk_word32 (EQUIV (TWO_COMP (Eps (dest_word32 T1))))" by (import word32 word_2comp) constdefs word_lsr1 :: "word32 => word32" - "word_lsr1 == %T1. mk_word32 (EQUIV (LSR_ONE (Eps (dest_word32 T1))))" + "word_lsr1 == %T1::word32. mk_word32 (EQUIV (LSR_ONE (Eps (dest_word32 T1))))" -lemma word_lsr1: "ALL T1. word_lsr1 T1 = mk_word32 (EQUIV (LSR_ONE (Eps (dest_word32 T1))))" +lemma word_lsr1: "ALL T1::word32. + word_lsr1 T1 = mk_word32 (EQUIV (LSR_ONE (Eps (dest_word32 T1))))" by (import word32 word_lsr1) constdefs word_asr1 :: "word32 => word32" - "word_asr1 == %T1. mk_word32 (EQUIV (ASR_ONE (Eps (dest_word32 T1))))" + "word_asr1 == %T1::word32. mk_word32 (EQUIV (ASR_ONE (Eps (dest_word32 T1))))" -lemma word_asr1: "ALL T1. word_asr1 T1 = mk_word32 (EQUIV (ASR_ONE (Eps (dest_word32 T1))))" +lemma word_asr1: "ALL T1::word32. + word_asr1 T1 = mk_word32 (EQUIV (ASR_ONE (Eps (dest_word32 T1))))" by (import word32 word_asr1) constdefs word_ror1 :: "word32 => word32" - "word_ror1 == %T1. mk_word32 (EQUIV (ROR_ONE (Eps (dest_word32 T1))))" + "word_ror1 == %T1::word32. mk_word32 (EQUIV (ROR_ONE (Eps (dest_word32 T1))))" -lemma word_ror1: "ALL T1. word_ror1 T1 = mk_word32 (EQUIV (ROR_ONE (Eps (dest_word32 T1))))" +lemma word_ror1: "ALL T1::word32. + word_ror1 T1 = mk_word32 (EQUIV (ROR_ONE (Eps (dest_word32 T1))))" by (import word32 word_ror1) consts RRX :: "bool => word32 => word32" defs - RRX_primdef: "RRX == %T1 T2. mk_word32 (EQUIV (RRXn T1 (Eps (dest_word32 T2))))" + RRX_primdef: "RRX == +%(T1::bool) T2::word32. mk_word32 (EQUIV (RRXn T1 (Eps (dest_word32 T2))))" -lemma RRX_def: "ALL T1 T2. RRX T1 T2 = mk_word32 (EQUIV (RRXn T1 (Eps (dest_word32 T2))))" +lemma RRX_def: "ALL (T1::bool) T2::word32. + RRX T1 T2 = mk_word32 (EQUIV (RRXn T1 (Eps (dest_word32 T2))))" by (import word32 RRX_def) consts LSB :: "word32 => bool" defs - LSB_primdef: "LSB == %T1. LSBn (Eps (dest_word32 T1))" + LSB_primdef: "LSB == %T1::word32. LSBn (Eps (dest_word32 T1))" -lemma LSB_def: "ALL T1. LSB T1 = LSBn (Eps (dest_word32 T1))" +lemma LSB_def: "ALL T1::word32. LSB T1 = LSBn (Eps (dest_word32 T1))" by (import word32 LSB_def) consts MSB :: "word32 => bool" defs - MSB_primdef: "MSB == %T1. MSBn (Eps (dest_word32 T1))" + MSB_primdef: "MSB == %T1::word32. MSBn (Eps (dest_word32 T1))" -lemma MSB_def: "ALL T1. MSB T1 = MSBn (Eps (dest_word32 T1))" +lemma MSB_def: "ALL T1::word32. MSB T1 = MSBn (Eps (dest_word32 T1))" by (import word32 MSB_def) constdefs bitwise_or :: "word32 => word32 => word32" "bitwise_or == -%T1 T2. mk_word32 (EQUIV (OR (Eps (dest_word32 T1)) (Eps (dest_word32 T2))))" +%(T1::word32) T2::word32. + mk_word32 (EQUIV (OR (Eps (dest_word32 T1)) (Eps (dest_word32 T2))))" -lemma bitwise_or: "ALL T1 T2. +lemma bitwise_or: "ALL (T1::word32) T2::word32. bitwise_or T1 T2 = mk_word32 (EQUIV (OR (Eps (dest_word32 T1)) (Eps (dest_word32 T2))))" by (import word32 bitwise_or) @@ -866,10 +930,10 @@ constdefs bitwise_eor :: "word32 => word32 => word32" "bitwise_eor == -%T1 T2. +%(T1::word32) T2::word32. mk_word32 (EQUIV (EOR (Eps (dest_word32 T1)) (Eps (dest_word32 T2))))" -lemma bitwise_eor: "ALL T1 T2. +lemma bitwise_eor: "ALL (T1::word32) T2::word32. bitwise_eor T1 T2 = mk_word32 (EQUIV (EOR (Eps (dest_word32 T1)) (Eps (dest_word32 T2))))" by (import word32 bitwise_eor) @@ -877,10 +941,10 @@ constdefs bitwise_and :: "word32 => word32 => word32" "bitwise_and == -%T1 T2. +%(T1::word32) T2::word32. mk_word32 (EQUIV (AND (Eps (dest_word32 T1)) (Eps (dest_word32 T2))))" -lemma bitwise_and: "ALL T1 T2. +lemma bitwise_and: "ALL (T1::word32) T2::word32. bitwise_and T1 T2 = mk_word32 (EQUIV (AND (Eps (dest_word32 T1)) (Eps (dest_word32 T2))))" by (import word32 bitwise_and) @@ -889,71 +953,78 @@ TOw :: "word32 => word32" defs - TOw_primdef: "TOw == %T1. mk_word32 (EQUIV (MODw (Eps (dest_word32 T1))))" + TOw_primdef: "TOw == %T1::word32. mk_word32 (EQUIV (MODw (Eps (dest_word32 T1))))" -lemma TOw_def: "ALL T1. TOw T1 = mk_word32 (EQUIV (MODw (Eps (dest_word32 T1))))" +lemma TOw_def: "ALL T1::word32. TOw T1 = mk_word32 (EQUIV (MODw (Eps (dest_word32 T1))))" by (import word32 TOw_def) consts n2w :: "nat => word32" defs - n2w_primdef: "n2w == %n. mk_word32 (EQUIV n)" + n2w_primdef: "n2w == %n::nat. mk_word32 (EQUIV n)" -lemma n2w_def: "ALL n. n2w n = mk_word32 (EQUIV n)" +lemma n2w_def: "ALL n::nat. n2w n = mk_word32 (EQUIV n)" by (import word32 n2w_def) consts w2n :: "word32 => nat" defs - w2n_primdef: "w2n == %w. MODw (Eps (dest_word32 w))" + w2n_primdef: "w2n == %w::word32. MODw (Eps (dest_word32 w))" -lemma w2n_def: "ALL w. w2n w = MODw (Eps (dest_word32 w))" +lemma w2n_def: "ALL w::word32. w2n w = MODw (Eps (dest_word32 w))" by (import word32 w2n_def) -lemma ADDw: "(ALL x. word_add w_0 x = x) & -(ALL x xa. word_add (word_suc x) xa = word_suc (word_add x xa))" +lemma ADDw: "(ALL x::word32. word_add w_0 x = x) & +(ALL (x::word32) xa::word32. + word_add (word_suc x) xa = word_suc (word_add x xa))" by (import word32 ADDw) -lemma ADD_0w: "ALL x. word_add x w_0 = x" +lemma ADD_0w: "ALL x::word32. word_add x w_0 = x" by (import word32 ADD_0w) -lemma ADD1w: "ALL x. word_suc x = word_add x w_1" +lemma ADD1w: "ALL x::word32. word_suc x = word_add x w_1" by (import word32 ADD1w) -lemma ADD_ASSOCw: "ALL x xa xb. word_add x (word_add xa xb) = word_add (word_add x xa) xb" +lemma ADD_ASSOCw: "ALL (x::word32) (xa::word32) xb::word32. + word_add x (word_add xa xb) = word_add (word_add x xa) xb" by (import word32 ADD_ASSOCw) -lemma ADD_CLAUSESw: "(ALL x. word_add w_0 x = x) & -(ALL x. word_add x w_0 = x) & -(ALL x xa. word_add (word_suc x) xa = word_suc (word_add x xa)) & -(ALL x xa. word_add x (word_suc xa) = word_suc (word_add x xa))" +lemma ADD_CLAUSESw: "(ALL x::word32. word_add w_0 x = x) & +(ALL x::word32. word_add x w_0 = x) & +(ALL (x::word32) xa::word32. + word_add (word_suc x) xa = word_suc (word_add x xa)) & +(ALL (x::word32) xa::word32. + word_add x (word_suc xa) = word_suc (word_add x xa))" by (import word32 ADD_CLAUSESw) -lemma ADD_COMMw: "ALL x xa. word_add x xa = word_add xa x" +lemma ADD_COMMw: "ALL (x::word32) xa::word32. word_add x xa = word_add xa x" by (import word32 ADD_COMMw) -lemma ADD_INV_0_EQw: "ALL x xa. (word_add x xa = x) = (xa = w_0)" +lemma ADD_INV_0_EQw: "ALL (x::word32) xa::word32. (word_add x xa = x) = (xa = w_0)" by (import word32 ADD_INV_0_EQw) -lemma EQ_ADD_LCANCELw: "ALL x xa xb. (word_add x xa = word_add x xb) = (xa = xb)" +lemma EQ_ADD_LCANCELw: "ALL (x::word32) (xa::word32) xb::word32. + (word_add x xa = word_add x xb) = (xa = xb)" by (import word32 EQ_ADD_LCANCELw) -lemma EQ_ADD_RCANCELw: "ALL x xa xb. (word_add x xb = word_add xa xb) = (x = xa)" +lemma EQ_ADD_RCANCELw: "ALL (x::word32) (xa::word32) xb::word32. + (word_add x xb = word_add xa xb) = (x = xa)" by (import word32 EQ_ADD_RCANCELw) -lemma LEFT_ADD_DISTRIBw: "ALL x xa xb. +lemma LEFT_ADD_DISTRIBw: "ALL (x::word32) (xa::word32) xb::word32. word_mul xb (word_add x xa) = word_add (word_mul xb x) (word_mul xb xa)" by (import word32 LEFT_ADD_DISTRIBw) -lemma MULT_ASSOCw: "ALL x xa xb. word_mul x (word_mul xa xb) = word_mul (word_mul x xa) xb" +lemma MULT_ASSOCw: "ALL (x::word32) (xa::word32) xb::word32. + word_mul x (word_mul xa xb) = word_mul (word_mul x xa) xb" by (import word32 MULT_ASSOCw) -lemma MULT_COMMw: "ALL x xa. word_mul x xa = word_mul xa x" +lemma MULT_COMMw: "ALL (x::word32) xa::word32. word_mul x xa = word_mul xa x" by (import word32 MULT_COMMw) -lemma MULT_CLAUSESw: "ALL x xa. +lemma MULT_CLAUSESw: "ALL (x::word32) xa::word32. word_mul w_0 x = w_0 & word_mul x w_0 = w_0 & word_mul w_1 x = x & @@ -962,58 +1033,58 @@ word_mul x (word_suc xa) = word_add x (word_mul x xa)" by (import word32 MULT_CLAUSESw) -lemma TWO_COMP_ONE_COMP: "ALL x. word_2comp x = word_add (word_1comp x) w_1" +lemma TWO_COMP_ONE_COMP: "ALL x::word32. word_2comp x = word_add (word_1comp x) w_1" by (import word32 TWO_COMP_ONE_COMP) -lemma OR_ASSOCw: "ALL x xa xb. +lemma OR_ASSOCw: "ALL (x::word32) (xa::word32) xb::word32. bitwise_or x (bitwise_or xa xb) = bitwise_or (bitwise_or x xa) xb" by (import word32 OR_ASSOCw) -lemma OR_COMMw: "ALL x xa. bitwise_or x xa = bitwise_or xa x" +lemma OR_COMMw: "ALL (x::word32) xa::word32. bitwise_or x xa = bitwise_or xa x" by (import word32 OR_COMMw) -lemma OR_IDEMw: "ALL x. bitwise_or x x = x" +lemma OR_IDEMw: "ALL x::word32. bitwise_or x x = x" by (import word32 OR_IDEMw) -lemma OR_ABSORBw: "ALL x xa. bitwise_and x (bitwise_or x xa) = x" +lemma OR_ABSORBw: "ALL (x::word32) xa::word32. bitwise_and x (bitwise_or x xa) = x" by (import word32 OR_ABSORBw) -lemma AND_ASSOCw: "ALL x xa xb. +lemma AND_ASSOCw: "ALL (x::word32) (xa::word32) xb::word32. bitwise_and x (bitwise_and xa xb) = bitwise_and (bitwise_and x xa) xb" by (import word32 AND_ASSOCw) -lemma AND_COMMw: "ALL x xa. bitwise_and x xa = bitwise_and xa x" +lemma AND_COMMw: "ALL (x::word32) xa::word32. bitwise_and x xa = bitwise_and xa x" by (import word32 AND_COMMw) -lemma AND_IDEMw: "ALL x. bitwise_and x x = x" +lemma AND_IDEMw: "ALL x::word32. bitwise_and x x = x" by (import word32 AND_IDEMw) -lemma AND_ABSORBw: "ALL x xa. bitwise_or x (bitwise_and x xa) = x" +lemma AND_ABSORBw: "ALL (x::word32) xa::word32. bitwise_or x (bitwise_and x xa) = x" by (import word32 AND_ABSORBw) -lemma ONE_COMPw: "ALL x. word_1comp (word_1comp x) = x" +lemma ONE_COMPw: "ALL x::word32. word_1comp (word_1comp x) = x" by (import word32 ONE_COMPw) -lemma RIGHT_AND_OVER_ORw: "ALL x xa xb. +lemma RIGHT_AND_OVER_ORw: "ALL (x::word32) (xa::word32) xb::word32. bitwise_and (bitwise_or x xa) xb = bitwise_or (bitwise_and x xb) (bitwise_and xa xb)" by (import word32 RIGHT_AND_OVER_ORw) -lemma RIGHT_OR_OVER_ANDw: "ALL x xa xb. +lemma RIGHT_OR_OVER_ANDw: "ALL (x::word32) (xa::word32) xb::word32. bitwise_or (bitwise_and x xa) xb = bitwise_and (bitwise_or x xb) (bitwise_or xa xb)" by (import word32 RIGHT_OR_OVER_ANDw) -lemma DE_MORGAN_THMw: "ALL x xa. +lemma DE_MORGAN_THMw: "ALL (x::word32) xa::word32. word_1comp (bitwise_and x xa) = bitwise_or (word_1comp x) (word_1comp xa) & word_1comp (bitwise_or x xa) = bitwise_and (word_1comp x) (word_1comp xa)" by (import word32 DE_MORGAN_THMw) -lemma w_0: "w_0 = n2w 0" +lemma w_0: "w_0 = n2w (0::nat)" by (import word32 w_0) -lemma w_1: "w_1 = n2w 1" +lemma w_1: "w_1 = n2w (1::nat)" by (import word32 w_1) lemma w_T: "w_T = @@ -1053,204 +1124,216 @@ ALT_ZERO)))))))))))))))))))))))))))))))))" by (import word32 w_T) -lemma ADD_TWO_COMP: "ALL x. word_add x (word_2comp x) = w_0" +lemma ADD_TWO_COMP: "ALL x::word32. word_add x (word_2comp x) = w_0" by (import word32 ADD_TWO_COMP) -lemma ADD_TWO_COMP2: "ALL x. word_add (word_2comp x) x = w_0" +lemma ADD_TWO_COMP2: "ALL x::word32. word_add (word_2comp x) x = w_0" by (import word32 ADD_TWO_COMP2) constdefs word_sub :: "word32 => word32 => word32" - "word_sub == %a b. word_add a (word_2comp b)" + "word_sub == %(a::word32) b::word32. word_add a (word_2comp b)" -lemma word_sub: "ALL a b. word_sub a b = word_add a (word_2comp b)" +lemma word_sub: "ALL (a::word32) b::word32. word_sub a b = word_add a (word_2comp b)" by (import word32 word_sub) constdefs word_lsl :: "word32 => nat => word32" - "word_lsl == %a n. word_mul a (n2w (2 ^ n))" + "word_lsl == %(a::word32) n::nat. word_mul a (n2w ((2::nat) ^ n))" -lemma word_lsl: "ALL a n. word_lsl a n = word_mul a (n2w (2 ^ n))" +lemma word_lsl: "ALL (a::word32) n::nat. word_lsl a n = word_mul a (n2w ((2::nat) ^ n))" by (import word32 word_lsl) constdefs word_lsr :: "word32 => nat => word32" - "word_lsr == %a n. (word_lsr1 ^ n) a" + "word_lsr == %(a::word32) n::nat. (word_lsr1 ^ n) a" -lemma word_lsr: "ALL a n. word_lsr a n = (word_lsr1 ^ n) a" +lemma word_lsr: "ALL (a::word32) n::nat. word_lsr a n = (word_lsr1 ^ n) a" by (import word32 word_lsr) constdefs word_asr :: "word32 => nat => word32" - "word_asr == %a n. (word_asr1 ^ n) a" + "word_asr == %(a::word32) n::nat. (word_asr1 ^ n) a" -lemma word_asr: "ALL a n. word_asr a n = (word_asr1 ^ n) a" +lemma word_asr: "ALL (a::word32) n::nat. word_asr a n = (word_asr1 ^ n) a" by (import word32 word_asr) constdefs word_ror :: "word32 => nat => word32" - "word_ror == %a n. (word_ror1 ^ n) a" + "word_ror == %(a::word32) n::nat. (word_ror1 ^ n) a" -lemma word_ror: "ALL a n. word_ror a n = (word_ror1 ^ n) a" +lemma word_ror: "ALL (a::word32) n::nat. word_ror a n = (word_ror1 ^ n) a" by (import word32 word_ror) consts BITw :: "nat => word32 => bool" defs - BITw_primdef: "BITw == %b n. bit b (w2n n)" + BITw_primdef: "BITw == %(b::nat) n::word32. bit b (w2n n)" -lemma BITw_def: "ALL b n. BITw b n = bit b (w2n n)" +lemma BITw_def: "ALL (b::nat) n::word32. BITw b n = bit b (w2n n)" by (import word32 BITw_def) consts BITSw :: "nat => nat => word32 => nat" defs - BITSw_primdef: "BITSw == %h l n. BITS h l (w2n n)" + BITSw_primdef: "BITSw == %(h::nat) (l::nat) n::word32. BITS h l (w2n n)" -lemma BITSw_def: "ALL h l n. BITSw h l n = BITS h l (w2n n)" +lemma BITSw_def: "ALL (h::nat) (l::nat) n::word32. BITSw h l n = BITS h l (w2n n)" by (import word32 BITSw_def) consts SLICEw :: "nat => nat => word32 => nat" defs - SLICEw_primdef: "SLICEw == %h l n. SLICE h l (w2n n)" + SLICEw_primdef: "SLICEw == %(h::nat) (l::nat) n::word32. SLICE h l (w2n n)" -lemma SLICEw_def: "ALL h l n. SLICEw h l n = SLICE h l (w2n n)" +lemma SLICEw_def: "ALL (h::nat) (l::nat) n::word32. SLICEw h l n = SLICE h l (w2n n)" by (import word32 SLICEw_def) -lemma TWO_COMP_ADD: "ALL a b. word_2comp (word_add a b) = word_add (word_2comp a) (word_2comp b)" +lemma TWO_COMP_ADD: "ALL (a::word32) b::word32. + word_2comp (word_add a b) = word_add (word_2comp a) (word_2comp b)" by (import word32 TWO_COMP_ADD) -lemma TWO_COMP_ELIM: "ALL a. word_2comp (word_2comp a) = a" +lemma TWO_COMP_ELIM: "ALL a::word32. word_2comp (word_2comp a) = a" by (import word32 TWO_COMP_ELIM) -lemma ADD_SUB_ASSOC: "ALL a b c. word_sub (word_add a b) c = word_add a (word_sub b c)" +lemma ADD_SUB_ASSOC: "ALL (a::word32) (b::word32) c::word32. + word_sub (word_add a b) c = word_add a (word_sub b c)" by (import word32 ADD_SUB_ASSOC) -lemma ADD_SUB_SYM: "ALL a b c. word_sub (word_add a b) c = word_add (word_sub a c) b" +lemma ADD_SUB_SYM: "ALL (a::word32) (b::word32) c::word32. + word_sub (word_add a b) c = word_add (word_sub a c) b" by (import word32 ADD_SUB_SYM) -lemma SUB_EQUALw: "ALL a. word_sub a a = w_0" +lemma SUB_EQUALw: "ALL a::word32. word_sub a a = w_0" by (import word32 SUB_EQUALw) -lemma ADD_SUBw: "ALL a b. word_sub (word_add a b) b = a" +lemma ADD_SUBw: "ALL (a::word32) b::word32. word_sub (word_add a b) b = a" by (import word32 ADD_SUBw) -lemma SUB_SUBw: "ALL a b c. word_sub a (word_sub b c) = word_sub (word_add a c) b" +lemma SUB_SUBw: "ALL (a::word32) (b::word32) c::word32. + word_sub a (word_sub b c) = word_sub (word_add a c) b" by (import word32 SUB_SUBw) -lemma ONE_COMP_TWO_COMP: "ALL a. word_1comp a = word_sub (word_2comp a) w_1" +lemma ONE_COMP_TWO_COMP: "ALL a::word32. word_1comp a = word_sub (word_2comp a) w_1" by (import word32 ONE_COMP_TWO_COMP) -lemma SUBw: "ALL m n. word_sub (word_suc m) n = word_suc (word_sub m n)" +lemma SUBw: "ALL (m::word32) n::word32. word_sub (word_suc m) n = word_suc (word_sub m n)" by (import word32 SUBw) -lemma ADD_EQ_SUBw: "ALL m n p. (word_add m n = p) = (m = word_sub p n)" +lemma ADD_EQ_SUBw: "ALL (m::word32) (n::word32) p::word32. + (word_add m n = p) = (m = word_sub p n)" by (import word32 ADD_EQ_SUBw) -lemma CANCEL_SUBw: "ALL m n p. (word_sub n p = word_sub m p) = (n = m)" +lemma CANCEL_SUBw: "ALL (m::word32) (n::word32) p::word32. + (word_sub n p = word_sub m p) = (n = m)" by (import word32 CANCEL_SUBw) -lemma SUB_PLUSw: "ALL a b c. word_sub a (word_add b c) = word_sub (word_sub a b) c" +lemma SUB_PLUSw: "ALL (a::word32) (b::word32) c::word32. + word_sub a (word_add b c) = word_sub (word_sub a b) c" by (import word32 SUB_PLUSw) -lemma word_nchotomy: "ALL w. EX n. w = n2w n" +lemma word_nchotomy: "ALL w::word32. EX n::nat. w = n2w n" by (import word32 word_nchotomy) -lemma dest_word_mk_word_eq3: "ALL a. dest_word32 (mk_word32 (EQUIV a)) = EQUIV a" +lemma dest_word_mk_word_eq3: "ALL a::nat. dest_word32 (mk_word32 (EQUIV a)) = EQUIV a" by (import word32 dest_word_mk_word_eq3) -lemma MODw_ELIM: "ALL n. n2w (MODw n) = n2w n" +lemma MODw_ELIM: "ALL n::nat. n2w (MODw n) = n2w n" by (import word32 MODw_ELIM) -lemma w2n_EVAL: "ALL n. w2n (n2w n) = MODw n" +lemma w2n_EVAL: "ALL n::nat. w2n (n2w n) = MODw n" by (import word32 w2n_EVAL) -lemma w2n_ELIM: "ALL a. n2w (w2n a) = a" +lemma w2n_ELIM: "ALL a::word32. n2w (w2n a) = a" by (import word32 w2n_ELIM) -lemma n2w_11: "ALL a b. (n2w a = n2w b) = (MODw a = MODw b)" +lemma n2w_11: "ALL (a::nat) b::nat. (n2w a = n2w b) = (MODw a = MODw b)" by (import word32 n2w_11) -lemma ADD_EVAL: "word_add (n2w a) (n2w b) = n2w (a + b)" +lemma ADD_EVAL: "word_add (n2w (a::nat)) (n2w (b::nat)) = n2w (a + b)" by (import word32 ADD_EVAL) -lemma MUL_EVAL: "word_mul (n2w a) (n2w b) = n2w (a * b)" +lemma MUL_EVAL: "word_mul (n2w (a::nat)) (n2w (b::nat)) = n2w (a * b)" by (import word32 MUL_EVAL) -lemma ONE_COMP_EVAL: "word_1comp (n2w a) = n2w (ONE_COMP a)" +lemma ONE_COMP_EVAL: "word_1comp (n2w (a::nat)) = n2w (ONE_COMP a)" by (import word32 ONE_COMP_EVAL) -lemma TWO_COMP_EVAL: "word_2comp (n2w a) = n2w (TWO_COMP a)" +lemma TWO_COMP_EVAL: "word_2comp (n2w (a::nat)) = n2w (TWO_COMP a)" by (import word32 TWO_COMP_EVAL) -lemma LSR_ONE_EVAL: "word_lsr1 (n2w a) = n2w (LSR_ONE a)" +lemma LSR_ONE_EVAL: "word_lsr1 (n2w (a::nat)) = n2w (LSR_ONE a)" by (import word32 LSR_ONE_EVAL) -lemma ASR_ONE_EVAL: "word_asr1 (n2w a) = n2w (ASR_ONE a)" +lemma ASR_ONE_EVAL: "word_asr1 (n2w (a::nat)) = n2w (ASR_ONE a)" by (import word32 ASR_ONE_EVAL) -lemma ROR_ONE_EVAL: "word_ror1 (n2w a) = n2w (ROR_ONE a)" +lemma ROR_ONE_EVAL: "word_ror1 (n2w (a::nat)) = n2w (ROR_ONE a)" by (import word32 ROR_ONE_EVAL) -lemma RRX_EVAL: "RRX c (n2w a) = n2w (RRXn c a)" +lemma RRX_EVAL: "RRX (c::bool) (n2w (a::nat)) = n2w (RRXn c a)" by (import word32 RRX_EVAL) -lemma LSB_EVAL: "LSB (n2w a) = LSBn a" +lemma LSB_EVAL: "LSB (n2w (a::nat)) = LSBn a" by (import word32 LSB_EVAL) -lemma MSB_EVAL: "MSB (n2w a) = MSBn a" +lemma MSB_EVAL: "MSB (n2w (a::nat)) = MSBn a" by (import word32 MSB_EVAL) -lemma OR_EVAL: "bitwise_or (n2w a) (n2w b) = n2w (OR a b)" +lemma OR_EVAL: "bitwise_or (n2w (a::nat)) (n2w (b::nat)) = n2w (OR a b)" by (import word32 OR_EVAL) -lemma EOR_EVAL: "bitwise_eor (n2w a) (n2w b) = n2w (EOR a b)" +lemma EOR_EVAL: "bitwise_eor (n2w (a::nat)) (n2w (b::nat)) = n2w (EOR a b)" by (import word32 EOR_EVAL) -lemma AND_EVAL: "bitwise_and (n2w a) (n2w b) = n2w (AND a b)" +lemma AND_EVAL: "bitwise_and (n2w (a::nat)) (n2w (b::nat)) = n2w (AND a b)" by (import word32 AND_EVAL) -lemma BITS_EVAL: "ALL h l a. BITSw h l (n2w a) = BITS h l (MODw a)" +lemma BITS_EVAL: "ALL (h::nat) (l::nat) a::nat. BITSw h l (n2w a) = BITS h l (MODw a)" by (import word32 BITS_EVAL) -lemma BIT_EVAL: "ALL b a. BITw b (n2w a) = bit b (MODw a)" +lemma BIT_EVAL: "ALL (b::nat) a::nat. BITw b (n2w a) = bit b (MODw a)" by (import word32 BIT_EVAL) -lemma SLICE_EVAL: "ALL h l a. SLICEw h l (n2w a) = SLICE h l (MODw a)" +lemma SLICE_EVAL: "ALL (h::nat) (l::nat) a::nat. SLICEw h l (n2w a) = SLICE h l (MODw a)" by (import word32 SLICE_EVAL) -lemma LSL_ADD: "ALL a m n. word_lsl (word_lsl a m) n = word_lsl a (m + n)" +lemma LSL_ADD: "ALL (a::word32) (m::nat) n::nat. + word_lsl (word_lsl a m) n = word_lsl a (m + n)" by (import word32 LSL_ADD) -lemma LSR_ADD: "ALL x xa xb. word_lsr (word_lsr x xa) xb = word_lsr x (xa + xb)" +lemma LSR_ADD: "ALL (x::word32) (xa::nat) xb::nat. + word_lsr (word_lsr x xa) xb = word_lsr x (xa + xb)" by (import word32 LSR_ADD) -lemma ASR_ADD: "ALL x xa xb. word_asr (word_asr x xa) xb = word_asr x (xa + xb)" +lemma ASR_ADD: "ALL (x::word32) (xa::nat) xb::nat. + word_asr (word_asr x xa) xb = word_asr x (xa + xb)" by (import word32 ASR_ADD) -lemma ROR_ADD: "ALL x xa xb. word_ror (word_ror x xa) xb = word_ror x (xa + xb)" +lemma ROR_ADD: "ALL (x::word32) (xa::nat) xb::nat. + word_ror (word_ror x xa) xb = word_ror x (xa + xb)" by (import word32 ROR_ADD) -lemma LSL_LIMIT: "ALL w n. HB < n --> word_lsl w n = w_0" +lemma LSL_LIMIT: "ALL (w::word32) n::nat. HB < n --> word_lsl w n = w_0" by (import word32 LSL_LIMIT) -lemma MOD_MOD_DIV: "ALL a b. INw (MODw a div 2 ^ b)" +lemma MOD_MOD_DIV: "ALL (a::nat) b::nat. INw (MODw a div (2::nat) ^ b)" by (import word32 MOD_MOD_DIV) -lemma MOD_MOD_DIV_2EXP: "ALL a n. MODw (MODw a div 2 ^ n) div 2 = MODw a div 2 ^ Suc n" +lemma MOD_MOD_DIV_2EXP: "ALL (a::nat) n::nat. + MODw (MODw a div (2::nat) ^ n) div (2::nat) = MODw a div (2::nat) ^ Suc n" by (import word32 MOD_MOD_DIV_2EXP) -lemma LSR_EVAL: "ALL n. word_lsr (n2w a) n = n2w (MODw a div 2 ^ n)" +lemma LSR_EVAL: "ALL n::nat. word_lsr (n2w (a::nat)) n = n2w (MODw a div (2::nat) ^ n)" by (import word32 LSR_EVAL) -lemma LSR_THM: "ALL x n. word_lsr (n2w n) x = n2w (BITS HB (min WL x) n)" +lemma LSR_THM: "ALL (x::nat) n::nat. word_lsr (n2w n) x = n2w (BITS HB (min WL x) n)" by (import word32 LSR_THM) -lemma LSR_LIMIT: "ALL x w. HB < x --> word_lsr w x = w_0" +lemma LSR_LIMIT: "ALL (x::nat) w::word32. HB < x --> word_lsr w x = w_0" by (import word32 LSR_LIMIT) lemma LEFT_SHIFT_LESS: "ALL (n::nat) (m::nat) a::nat. @@ -1258,41 +1341,44 @@ (2::nat) ^ n + a * (2::nat) ^ n <= (2::nat) ^ (m + n)" by (import word32 LEFT_SHIFT_LESS) -lemma ROR_THM: "ALL x n. +lemma ROR_THM: "ALL (x::nat) n::nat. word_ror (n2w n) x = - (let x' = x mod WL - in n2w (BITS HB x' n + BITS (x' - 1) 0 n * 2 ^ (WL - x')))" + (let x'::nat = x mod WL + in n2w (BITS HB x' n + + BITS (x' - (1::nat)) (0::nat) n * (2::nat) ^ (WL - x')))" by (import word32 ROR_THM) -lemma ROR_CYCLE: "ALL x w. word_ror w (x * WL) = w" +lemma ROR_CYCLE: "ALL (x::nat) w::word32. word_ror w (x * WL) = w" by (import word32 ROR_CYCLE) -lemma ASR_THM: "ALL x n. +lemma ASR_THM: "ALL (x::nat) n::nat. word_asr (n2w n) x = - (let x' = min HB x; s = BITS HB x' n - in n2w (if MSBn n then 2 ^ WL - 2 ^ (WL - x') + s else s))" + (let x'::nat = min HB x; s::nat = BITS HB x' n + in n2w (if MSBn n then (2::nat) ^ WL - (2::nat) ^ (WL - x') + s else s))" by (import word32 ASR_THM) -lemma ASR_LIMIT: "ALL x w. HB <= x --> word_asr w x = (if MSB w then w_T else w_0)" +lemma ASR_LIMIT: "ALL (x::nat) w::word32. + HB <= x --> word_asr w x = (if MSB w then w_T else w_0)" by (import word32 ASR_LIMIT) -lemma ZERO_SHIFT: "(ALL n. word_lsl w_0 n = w_0) & -(ALL n. word_asr w_0 n = w_0) & -(ALL n. word_lsr w_0 n = w_0) & (ALL n. word_ror w_0 n = w_0)" +lemma ZERO_SHIFT: "(ALL n::nat. word_lsl w_0 n = w_0) & +(ALL n::nat. word_asr w_0 n = w_0) & +(ALL n::nat. word_lsr w_0 n = w_0) & (ALL n::nat. word_ror w_0 n = w_0)" by (import word32 ZERO_SHIFT) -lemma ZERO_SHIFT2: "(ALL a. word_lsl a 0 = a) & -(ALL a. word_asr a 0 = a) & -(ALL a. word_lsr a 0 = a) & (ALL a. word_ror a 0 = a)" +lemma ZERO_SHIFT2: "(ALL a::word32. word_lsl a (0::nat) = a) & +(ALL a::word32. word_asr a (0::nat) = a) & +(ALL a::word32. word_lsr a (0::nat) = a) & +(ALL a::word32. word_ror a (0::nat) = a)" by (import word32 ZERO_SHIFT2) -lemma ASR_w_T: "ALL n. word_asr w_T n = w_T" +lemma ASR_w_T: "ALL n::nat. word_asr w_T n = w_T" by (import word32 ASR_w_T) -lemma ROR_w_T: "ALL n. word_ror w_T n = w_T" +lemma ROR_w_T: "ALL n::nat. word_ror w_T n = w_T" by (import word32 ROR_w_T) -lemma MODw_EVAL: "ALL x. +lemma MODw_EVAL: "ALL x::nat. MODw x = x mod NUMERAL @@ -1331,27 +1417,27 @@ ALT_ZERO))))))))))))))))))))))))))))))))" by (import word32 MODw_EVAL) -lemma ADD_EVAL2: "ALL b a. word_add (n2w a) (n2w b) = n2w (MODw (a + b))" +lemma ADD_EVAL2: "ALL (b::nat) a::nat. word_add (n2w a) (n2w b) = n2w (MODw (a + b))" by (import word32 ADD_EVAL2) -lemma MUL_EVAL2: "ALL b a. word_mul (n2w a) (n2w b) = n2w (MODw (a * b))" +lemma MUL_EVAL2: "ALL (b::nat) a::nat. word_mul (n2w a) (n2w b) = n2w (MODw (a * b))" by (import word32 MUL_EVAL2) -lemma ONE_COMP_EVAL2: "ALL a. +lemma ONE_COMP_EVAL2: "ALL a::nat. word_1comp (n2w a) = - n2w (2 ^ + n2w ((2::nat) ^ NUMERAL (NUMERAL_BIT2 (NUMERAL_BIT1 (NUMERAL_BIT1 (NUMERAL_BIT1 (NUMERAL_BIT1 ALT_ZERO))))) - - 1 - + (1::nat) - MODw a)" by (import word32 ONE_COMP_EVAL2) -lemma TWO_COMP_EVAL2: "ALL a. +lemma TWO_COMP_EVAL2: "ALL a::nat. word_2comp (n2w a) = n2w (MODw - (2 ^ + ((2::nat) ^ NUMERAL (NUMERAL_BIT2 (NUMERAL_BIT1 @@ -1359,12 +1445,12 @@ MODw a))" by (import word32 TWO_COMP_EVAL2) -lemma LSR_ONE_EVAL2: "ALL a. word_lsr1 (n2w a) = n2w (MODw a div 2)" +lemma LSR_ONE_EVAL2: "ALL a::nat. word_lsr1 (n2w a) = n2w (MODw a div (2::nat))" by (import word32 LSR_ONE_EVAL2) -lemma ASR_ONE_EVAL2: "ALL a. +lemma ASR_ONE_EVAL2: "ALL a::nat. word_asr1 (n2w a) = - n2w (MODw a div 2 + + n2w (MODw a div (2::nat) + SBIT (MSBn a) (NUMERAL (NUMERAL_BIT1 @@ -1372,9 +1458,9 @@ (NUMERAL_BIT1 (NUMERAL_BIT1 (NUMERAL_BIT1 ALT_ZERO)))))))" by (import word32 ASR_ONE_EVAL2) -lemma ROR_ONE_EVAL2: "ALL a. +lemma ROR_ONE_EVAL2: "ALL a::nat. word_ror1 (n2w a) = - n2w (MODw a div 2 + + n2w (MODw a div (2::nat) + SBIT (LSBn a) (NUMERAL (NUMERAL_BIT1 @@ -1382,9 +1468,9 @@ (NUMERAL_BIT1 (NUMERAL_BIT1 (NUMERAL_BIT1 ALT_ZERO)))))))" by (import word32 ROR_ONE_EVAL2) -lemma RRX_EVAL2: "ALL c a. +lemma RRX_EVAL2: "ALL (c::bool) a::nat. RRX c (n2w a) = - n2w (MODw a div 2 + + n2w (MODw a div (2::nat) + SBIT c (NUMERAL (NUMERAL_BIT1 @@ -1392,10 +1478,10 @@ (NUMERAL_BIT1 (NUMERAL_BIT1 (NUMERAL_BIT1 ALT_ZERO)))))))" by (import word32 RRX_EVAL2) -lemma LSB_EVAL2: "ALL a. LSB (n2w a) = ODD a" +lemma LSB_EVAL2: "ALL a::nat. LSB (n2w a) = ODD a" by (import word32 LSB_EVAL2) -lemma MSB_EVAL2: "ALL a. +lemma MSB_EVAL2: "ALL a::nat. MSB (n2w a) = bit (NUMERAL (NUMERAL_BIT1 @@ -1404,7 +1490,7 @@ a" by (import word32 MSB_EVAL2) -lemma OR_EVAL2: "ALL b a. +lemma OR_EVAL2: "ALL (b::nat) a::nat. bitwise_or (n2w a) (n2w b) = n2w (BITWISE (NUMERAL @@ -1414,7 +1500,7 @@ op | a b)" by (import word32 OR_EVAL2) -lemma AND_EVAL2: "ALL b a. +lemma AND_EVAL2: "ALL (b::nat) a::nat. bitwise_and (n2w a) (n2w b) = n2w (BITWISE (NUMERAL @@ -1424,54 +1510,91 @@ op & a b)" by (import word32 AND_EVAL2) -lemma EOR_EVAL2: "ALL b a. +lemma EOR_EVAL2: "ALL (b::nat) a::nat. bitwise_eor (n2w a) (n2w b) = n2w (BITWISE (NUMERAL (NUMERAL_BIT2 (NUMERAL_BIT1 (NUMERAL_BIT1 (NUMERAL_BIT1 (NUMERAL_BIT1 ALT_ZERO)))))) - (%x y. x ~= y) a b)" + (%(x::bool) y::bool. x ~= y) a b)" by (import word32 EOR_EVAL2) -lemma BITWISE_EVAL2: "ALL n oper x y. - BITWISE n oper x y = - (if n = 0 then 0 - else 2 * BITWISE (n - 1) oper (x div 2) (y div 2) + - (if oper (ODD x) (ODD y) then 1 else 0))" +lemma BITWISE_EVAL2: "(All::(nat => bool) => bool) + (%n::nat. + (All::((bool => bool => bool) => bool) => bool) + (%oper::bool => bool => bool. + (All::(nat => bool) => bool) + (%x::nat. + (All::(nat => bool) => bool) + (%y::nat. + (op =::nat => nat => bool) + ((BITWISE::nat + => (bool => bool => bool) + => nat => nat => nat) + n oper x y) + ((If::bool => nat => nat => nat) + ((op =::nat => nat => bool) n (0::nat)) (0::nat) + ((op +::nat => nat => nat) + ((op *::nat => nat => nat) + ((number_of::bin => nat) + ((op BIT::bin => bit => bin) + ((op BIT::bin => bit => bin) + (Numeral.Pls::bin) (bit.B1::bit)) + (bit.B0::bit))) + ((BITWISE::nat +=> (bool => bool => bool) => nat => nat => nat) + ((op -::nat => nat => nat) n (1::nat)) oper + ((op div::nat => nat => nat) x + ((number_of::bin => nat) + ((op BIT::bin => bit => bin) + ((op BIT::bin => bit => bin) + (Numeral.Pls::bin) (bit.B1::bit)) + (bit.B0::bit)))) + ((op div::nat => nat => nat) y + ((number_of::bin => nat) + ((op BIT::bin => bit => bin) + ((op BIT::bin => bit => bin) + (Numeral.Pls::bin) (bit.B1::bit)) + (bit.B0::bit)))))) + ((If::bool => nat => nat => nat) + (oper ((ODD::nat => bool) x) + ((ODD::nat => bool) y)) + (1::nat) (0::nat))))))))" by (import word32 BITWISE_EVAL2) -lemma BITSwLT_THM: "ALL h l n. BITSw h l n < 2 ^ (Suc h - l)" +lemma BITSwLT_THM: "ALL (h::nat) (l::nat) n::word32. BITSw h l n < (2::nat) ^ (Suc h - l)" by (import word32 BITSwLT_THM) -lemma BITSw_COMP_THM: "ALL h1 l1 h2 l2 n. +lemma BITSw_COMP_THM: "ALL (h1::nat) (l1::nat) (h2::nat) (l2::nat) n::word32. h2 + l1 <= h1 --> BITS h2 l2 (BITSw h1 l1 n) = BITSw (h2 + l1) (l2 + l1) n" by (import word32 BITSw_COMP_THM) -lemma BITSw_DIV_THM: "ALL h l n x. BITSw h l x div 2 ^ n = BITSw h (l + n) x" +lemma BITSw_DIV_THM: "ALL (h::nat) (l::nat) (n::nat) x::word32. + BITSw h l x div (2::nat) ^ n = BITSw h (l + n) x" by (import word32 BITSw_DIV_THM) -lemma BITw_THM: "ALL b n. BITw b n = (BITSw b b n = 1)" +lemma BITw_THM: "ALL (b::nat) n::word32. BITw b n = (BITSw b b n = (1::nat))" by (import word32 BITw_THM) -lemma SLICEw_THM: "ALL n h l. SLICEw h l n = BITSw h l n * 2 ^ l" +lemma SLICEw_THM: "ALL (n::word32) (h::nat) l::nat. SLICEw h l n = BITSw h l n * (2::nat) ^ l" by (import word32 SLICEw_THM) -lemma BITS_SLICEw_THM: "ALL h l n. BITS h l (SLICEw h l n) = BITSw h l n" +lemma BITS_SLICEw_THM: "ALL (h::nat) (l::nat) n::word32. BITS h l (SLICEw h l n) = BITSw h l n" by (import word32 BITS_SLICEw_THM) -lemma SLICEw_ZERO_THM: "ALL n h. SLICEw h 0 n = BITSw h 0 n" +lemma SLICEw_ZERO_THM: "ALL (n::word32) h::nat. SLICEw h (0::nat) n = BITSw h (0::nat) n" by (import word32 SLICEw_ZERO_THM) -lemma SLICEw_COMP_THM: "ALL h m l a. +lemma SLICEw_COMP_THM: "ALL (h::nat) (m::nat) (l::nat) a::word32. Suc m <= h & l <= m --> SLICEw h (Suc m) a + SLICEw m l a = SLICEw h l a" by (import word32 SLICEw_COMP_THM) -lemma BITSw_ZERO: "ALL h l n. h < l --> BITSw h l n = 0" +lemma BITSw_ZERO: "ALL (h::nat) (l::nat) n::word32. h < l --> BITSw h l n = (0::nat)" by (import word32 BITSw_ZERO) -lemma SLICEw_ZERO: "ALL h l n. h < l --> SLICEw h l n = 0" +lemma SLICEw_ZERO: "ALL (h::nat) (l::nat) n::word32. h < l --> SLICEw h l n = (0::nat)" by (import word32 SLICEw_ZERO) ;end_setup diff -r 7e417a7cbf9f -r bd59bfd4bf37 src/HOL/Import/HOL/ROOT.ML --- a/src/HOL/Import/HOL/ROOT.ML Mon Sep 26 02:06:44 2005 +0200 +++ b/src/HOL/Import/HOL/ROOT.ML Mon Sep 26 02:27:14 2005 +0200 @@ -3,5 +3,6 @@ Author: Sebastian Skalberg (TU Muenchen) *) -setmp quick_and_dirty true use_thy "HOL4Prob"; -setmp quick_and_dirty true use_thy "HOL4"; +set show_types; set show_sorts; +use_thy "HOL4Prob"; +use_thy "HOL4"; diff -r 7e417a7cbf9f -r bd59bfd4bf37 src/HOL/Import/HOL/bool.imp --- a/src/HOL/Import/HOL/bool.imp Mon Sep 26 02:06:44 2005 +0200 +++ b/src/HOL/Import/HOL/bool.imp Mon Sep 26 02:27:14 2005 +0200 @@ -21,7 +21,7 @@ "RES_FORALL" > "HOL4Base.bool.RES_FORALL" "RES_EXISTS_UNIQUE" > "HOL4Base.bool.RES_EXISTS_UNIQUE" "RES_EXISTS" > "HOL4Base.bool.RES_EXISTS" - "ONTO" > "HOL4Setup.ONTO" + "ONTO" > "Fun.surj" "ONE_ONE" > "HOL4Setup.ONE_ONE" "LET" > "HOL4Compat.LET" "IN" > "HOL4Base.bool.IN" @@ -37,11 +37,11 @@ "bool_case_thm" > "HOL4Base.bool.bool_case_thm" "bool_case_ID" > "HOL4Base.bool.bool_case_ID" "bool_case_DEF" > "HOL4Compat.bool_case_DEF" - "bool_INDUCT" > "Datatype.bool.induct_correctness" + "bool_INDUCT" > "Datatype.bool.induct" "boolAxiom" > "HOL4Base.bool.boolAxiom" "UNWIND_THM2" > "HOL.simp_thms_39" "UNWIND_THM1" > "HOL.simp_thms_40" - "UNWIND_FORALL_THM2" > "HOL.simp_thms_41" + "UNWIND_FORALL_THM2" > "HOL4Base.bool.UNWIND_FORALL_THM2" "UNWIND_FORALL_THM1" > "HOL.simp_thms_42" "UEXISTS_SIMP" > "HOL4Base.bool.UEXISTS_SIMP" "UEXISTS_OR_THM" > "HOL4Base.bool.UEXISTS_OR_THM" @@ -82,8 +82,8 @@ "OR_DEF" > "HOL.or_def" "OR_CONG" > "HOL4Base.bool.OR_CONG" "OR_CLAUSES" > "HOL4Base.bool.OR_CLAUSES" - "ONTO_THM" > "HOL4Setup.ONTO_DEF" - "ONTO_DEF" > "HOL4Setup.ONTO_DEF" + "ONTO_THM" > "Fun.surj_def" + "ONTO_DEF" > "Fun.surj_def" "ONE_ONE_THM" > "HOL4Base.bool.ONE_ONE_THM" "ONE_ONE_DEF" > "HOL4Setup.ONE_ONE_DEF" "NOT_IMP" > "HOL.not_imp" @@ -95,7 +95,7 @@ "NOT_AND" > "HOL4Base.bool.NOT_AND" "MONO_OR" > "Inductive.basic_monos_3" "MONO_NOT" > "HOL.rev_contrapos" - "MONO_IMP" > "Set.imp_mono" + "MONO_IMP" > "HOL4Base.bool.MONO_IMP" "MONO_EXISTS" > "Inductive.basic_monos_5" "MONO_COND" > "HOL4Base.bool.MONO_COND" "MONO_AND" > "Inductive.basic_monos_4" @@ -108,11 +108,11 @@ "LEFT_OR_OVER_AND" > "HOL.disj_conj_distribL" "LEFT_OR_EXISTS_THM" > "HOL.ex_simps_3" "LEFT_FORALL_OR_THM" > "HOL.all_simps_3" - "LEFT_FORALL_IMP_THM" > "HOL.imp_ex" + "LEFT_FORALL_IMP_THM" > "HOL4Base.bool.LEFT_FORALL_IMP_THM" "LEFT_EXISTS_IMP_THM" > "HOL.imp_all" "LEFT_EXISTS_AND_THM" > "HOL.ex_simps_1" "LEFT_AND_OVER_OR" > "HOL.conj_disj_distribL" - "LEFT_AND_FORALL_THM" > "HOL.all_simps_1" + "LEFT_AND_FORALL_THM" > "HOL4Base.bool.LEFT_AND_FORALL_THM" "IN_def" > "HOL4Base.bool.IN_def" "IN_DEF" > "HOL4Base.bool.IN_DEF" "INFINITY_AX" > "HOL4Setup.INFINITY_AX" @@ -146,7 +146,7 @@ "EQ_SYM" > "HOL.meta_eq_to_obj_eq" "EQ_REFL" > "Presburger.fm_modd_pinf" "EQ_IMP_THM" > "HOL.iff_conv_conj_imp" - "EQ_EXT" > "HOL.meta_eq_to_obj_eq" + "EQ_EXT" > "HOL4Base.bool.EQ_EXT" "EQ_EXPAND" > "HOL4Base.bool.EQ_EXPAND" "EQ_CLAUSES" > "HOL4Base.bool.EQ_CLAUSES" "DISJ_SYM" > "HOL.disj_comms_1" diff -r 7e417a7cbf9f -r bd59bfd4bf37 src/HOL/Import/HOL/pair.imp --- a/src/HOL/Import/HOL/pair.imp Mon Sep 26 02:06:44 2005 +0200 +++ b/src/HOL/Import/HOL/pair.imp Mon Sep 26 02:27:14 2005 +0200 @@ -21,7 +21,7 @@ "##" > "prod_fun" thm_maps - "pair_induction" > "Datatype.prod.induct_correctness" + "pair_induction" > "Datatype.prod.induct" "pair_case_thm" > "Product_Type.split" "pair_case_def" > "HOL4Compat.pair_case_def" "pair_case_cong" > "HOL4Base.pair.pair_case_cong" diff -r 7e417a7cbf9f -r bd59bfd4bf37 src/HOL/Import/HOL/real.imp --- a/src/HOL/Import/HOL/real.imp Mon Sep 26 02:06:44 2005 +0200 +++ b/src/HOL/Import/HOL/real.imp Mon Sep 26 02:27:14 2005 +0200 @@ -31,7 +31,7 @@ "real_lt" > "Orderings.linorder_not_le" "real_gt" > "HOL4Compat.real_gt" "real_ge" > "HOL4Compat.real_ge" - "real_div" > "Ring_and_Field.field.divide_inverse" + "real_div" > "Ring_and_Field.field_class.divide_inverse" "pow" > "HOL4Compat.pow" "abs" > "HOL4Compat.abs" "SUP_LEMMA3" > "HOL4Real.real.SUP_LEMMA3" @@ -147,7 +147,7 @@ "REAL_LT_TOTAL" > "HOL4Compat.REAL_LT_TOTAL" "REAL_LT_SUB_RADD" > "OrderedGroup.compare_rls_6" "REAL_LT_SUB_LADD" > "OrderedGroup.compare_rls_7" - "REAL_LT_RMUL_IMP" > "Ring_and_Field.ordered_semiring_strict.mult_strict_right_mono" + "REAL_LT_RMUL_IMP" > "Ring_and_Field.ordered_semiring_strict_class.mult_strict_right_mono" "REAL_LT_RMUL_0" > "HOL4Real.real.REAL_LT_RMUL_0" "REAL_LT_RMUL" > "RealDef.real_mult_less_iff1" "REAL_LT_REFL" > "Orderings.order_less_irrefl" @@ -161,10 +161,10 @@ "REAL_LT_MULTIPLE" > "HOL4Real.real.REAL_LT_MULTIPLE" "REAL_LT_MUL2" > "Ring_and_Field.mult_strict_mono'" "REAL_LT_MUL" > "Ring_and_Field.mult_pos_pos" - "REAL_LT_LMUL_IMP" > "Ring_and_Field.ordered_comm_semiring_strict.mult_strict_mono" + "REAL_LT_LMUL_IMP" > "Ring_and_Field.ordered_comm_semiring_strict_class.mult_strict_mono" "REAL_LT_LMUL_0" > "HOL4Real.real.REAL_LT_LMUL_0" "REAL_LT_LMUL" > "HOL4Real.real.REAL_LT_LMUL" - "REAL_LT_LE" > "Orderings.order.order_less_le" + "REAL_LT_LE" > "Orderings.order_class.order_less_le" "REAL_LT_LDIV_EQ" > "Ring_and_Field.pos_divide_less_eq" "REAL_LT_LADD" > "OrderedGroup.add_less_cancel_left" "REAL_LT_INV_EQ" > "Ring_and_Field.inverse_positive_iff_positive" @@ -188,7 +188,7 @@ "REAL_LT_ADD1" > "HOL4Real.real.REAL_LT_ADD1" "REAL_LT_ADD" > "OrderedGroup.add_pos_pos" "REAL_LT_1" > "HOL4Real.real.REAL_LT_1" - "REAL_LT_01" > "Ring_and_Field.ordered_semidom.zero_less_one" + "REAL_LT_01" > "Ring_and_Field.ordered_semidom_class.zero_less_one" "REAL_LTE_TRANS" > "Set.basic_trans_rules_24" "REAL_LTE_TOTAL" > "HOL4Real.real.REAL_LTE_TOTAL" "REAL_LTE_ANTSYM" > "HOL4Real.real.REAL_LTE_ANTSYM" @@ -199,12 +199,12 @@ "REAL_LNEG_UNIQ" > "HOL4Real.real.REAL_LNEG_UNIQ" "REAL_LINV_UNIQ" > "HOL4Real.real.REAL_LINV_UNIQ" "REAL_LE_TRANS" > "Set.basic_trans_rules_25" - "REAL_LE_TOTAL" > "Orderings.linorder.linorder_linear" + "REAL_LE_TOTAL" > "Orderings.linorder_class.linorder_linear" "REAL_LE_SUB_RADD" > "OrderedGroup.compare_rls_8" "REAL_LE_SUB_LADD" > "OrderedGroup.compare_rls_9" "REAL_LE_SQUARE" > "Ring_and_Field.zero_le_square" "REAL_LE_RNEG" > "OrderedGroup.le_eq_neg" - "REAL_LE_RMUL_IMP" > "Ring_and_Field.pordered_semiring.mult_right_mono" + "REAL_LE_RMUL_IMP" > "Ring_and_Field.pordered_semiring_class.mult_right_mono" "REAL_LE_RMUL" > "RealDef.real_mult_le_cancel_iff1" "REAL_LE_REFL" > "Finite_Set.max.f_below.below_refl" "REAL_LE_RDIV_EQ" > "Ring_and_Field.pos_le_divide_eq" @@ -220,11 +220,11 @@ "REAL_LE_MUL" > "Ring_and_Field.mult_nonneg_nonneg" "REAL_LE_LT" > "Orderings.order_le_less" "REAL_LE_LNEG" > "RealDef.real_0_le_add_iff" - "REAL_LE_LMUL_IMP" > "Ring_and_Field.pordered_comm_semiring.mult_mono" + "REAL_LE_LMUL_IMP" > "Ring_and_Field.pordered_comm_semiring_class.mult_mono" "REAL_LE_LMUL" > "RealDef.real_mult_le_cancel_iff2" "REAL_LE_LDIV_EQ" > "Ring_and_Field.pos_divide_le_eq" "REAL_LE_LDIV" > "Ring_and_Field.mult_imp_div_pos_le" - "REAL_LE_LADD_IMP" > "OrderedGroup.pordered_ab_semigroup_add.add_left_mono" + "REAL_LE_LADD_IMP" > "OrderedGroup.pordered_ab_semigroup_add_class.add_left_mono" "REAL_LE_LADD" > "OrderedGroup.add_le_cancel_left" "REAL_LE_INV_EQ" > "Ring_and_Field.inverse_nonnegative_iff_nonnegative" "REAL_LE_INV" > "HOL4Real.real.REAL_LE_INV" @@ -251,7 +251,7 @@ "REAL_INV_INV" > "Ring_and_Field.inverse_inverse_eq" "REAL_INV_EQ_0" > "Ring_and_Field.inverse_nonzero_iff_nonzero" "REAL_INV_1OVER" > "Ring_and_Field.inverse_eq_divide" - "REAL_INV_0" > "Ring_and_Field.division_by_zero.inverse_zero" + "REAL_INV_0" > "Ring_and_Field.division_by_zero_class.inverse_zero" "REAL_INVINV" > "Ring_and_Field.nonzero_inverse_inverse_eq" "REAL_INV1" > "Ring_and_Field.inverse_1" "REAL_INJ" > "RealDef.real_of_nat_inject" diff -r 7e417a7cbf9f -r bd59bfd4bf37 src/HOL/Import/HOL/realax.imp --- a/src/HOL/Import/HOL/realax.imp Mon Sep 26 02:06:44 2005 +0200 +++ b/src/HOL/Import/HOL/realax.imp Mon Sep 26 02:27:14 2005 +0200 @@ -101,7 +101,7 @@ "REAL_LT_MUL" > "Ring_and_Field.mult_pos_pos" "REAL_LT_IADD" > "OrderedGroup.add_strict_left_mono" "REAL_LDISTRIB" > "Ring_and_Field.ring_eq_simps_2" - "REAL_INV_0" > "Ring_and_Field.division_by_zero.inverse_zero" + "REAL_INV_0" > "Ring_and_Field.division_by_zero_class.inverse_zero" "REAL_ADD_SYM" > "Finite_Set.AC_add.f.AC_2" "REAL_ADD_LINV" > "HOL4Compat.REAL_ADD_LINV" "REAL_ADD_LID" > "Finite_Set.AC_add.f_e.left_ident" diff -r 7e417a7cbf9f -r bd59bfd4bf37 src/HOL/Import/HOLLight/HOLLight.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Import/HOLLight/HOLLight.thy Mon Sep 26 02:27:14 2005 +0200 @@ -0,0 +1,13857 @@ +(* AUTOMATICALLY GENERATED, DO NOT EDIT! *) + +theory HOLLight = "../HOLLightCompat" + "../HOL4Syntax": + +;setup_theory hollight + +consts + "_FALSITY_" :: "bool" ("'_FALSITY'_") + +defs + "_FALSITY__def": "_FALSITY_ == False" + +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) & +(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) & +(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)" + 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" + 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" + 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)" + by (import hollight IMP_CLAUSES) + +lemma IMP_EQ_CLAUSE: "((x::'q_864::type) = 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))" + 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)" + 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))" + 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)" + 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))" + 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))" + 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))" + 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" + by (import hollight SELECT_UNIQUE) + +lemma EXCLUDED_MIDDLE: "ALL t::bool. t | ~ t" + by (import hollight EXCLUDED_MIDDLE) + +constdefs + COND :: "bool => 'A::type => 'A::type => 'A::type" + "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" + by (import hollight COND_CLAUSES) + +lemma COND_EXPAND: "ALL (b::bool) (t1::bool) t2::bool. COND b t1 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)" + 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" + 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" + 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))" + 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))" + 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" + 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" + by (import hollight o_ASSOC) + +lemma I_O_ID: "ALL f::'A::type => 'B::type. id o f = f & f o id = f" + by (import hollight I_O_ID) + +lemma EXISTS_ONE_REP: "EX x::bool. x" + by (import hollight EXISTS_ONE_REP) + +lemma one_axiom: "ALL f::'A::type => unit. All (op = f)" + by (import hollight one_axiom) + +lemma one_RECURSION: "ALL e::'A::type. EX x::unit => 'A::type. x () = e" + by (import hollight one_RECURSION) + +lemma one_Axiom: "ALL e::'A::type. EX! fn::unit => 'A::type. fn () = e" + 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)" + by (import hollight th_cond) + +constdefs + LET_END :: "'A::type => 'A::type" + "LET_END == %t::'A::type. t" + +lemma DEF_LET_END: "LET_END = (%t::'A::type. t)" + by (import hollight DEF_LET_END) + +constdefs + GABS :: "('A::type => bool) => 'A::type" + "(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) + +constdefs + GEQ :: "'A::type => 'A::type => bool" + "(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" + by (import hollight PAIR_EXISTS_THM) + +constdefs + CURRY :: "('A::type * 'B::type => 'C::type) => 'A::type => 'B::type => 'C::type" + "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) + +constdefs + UNCURRY :: "('A::type => 'B::type => 'C::type) => 'A::type * 'B::type => 'C::type" + "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))" + by (import hollight DEF_UNCURRY) + +constdefs + PASSOC :: "(('A::type * 'B::type) * 'C::type => 'D::type) +=> 'A::type * 'B::type * 'C::type => 'D::type" + "PASSOC == +%(u::('A::type * 'B::type) * 'C::type => 'D::type) + ua::'A::type * 'B::type * 'C::type. + 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 ((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::nat) = e & (ALL n::nat. fn (Suc n) = f (fn n) n)" + by (import hollight num_Axiom) + +lemma ADD_CLAUSES: "(ALL x::nat. (0::nat) + x = x) & +(ALL x::nat. x + (0::nat) = x) & +(ALL (x::nat) xa::nat. Suc x + xa = Suc (x + xa)) & +(ALL (x::nat) xa::nat. x + Suc xa = Suc (x + xa))" + by (import hollight ADD_CLAUSES) + +lemma ADD_AC: "(m::nat) + (n::nat) = n + m & +m + n + (p::nat) = m + (n + p) & m + (n + p) = n + (m + p)" + by (import hollight ADD_AC) + +lemma EQ_ADD_LCANCEL_0: "ALL (m::nat) n::nat. (m + n = 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::nat))" + by (import hollight EQ_ADD_RCANCEL_0) + +lemma ONE: "NUMERAL_BIT1 (0::nat) = Suc (0::nat)" + by (import hollight ONE) + +lemma TWO: "NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)) = Suc (NUMERAL_BIT1 (0::nat))" + by (import hollight TWO) + +lemma ADD1: "ALL x::nat. Suc x = x + NUMERAL_BIT1 (0::nat)" + by (import hollight ADD1) + +lemma MULT_CLAUSES: "(ALL x::nat. (0::nat) * x = (0::nat)) & +(ALL x::nat. x * (0::nat) = (0::nat)) & +(ALL x::nat. NUMERAL_BIT1 (0::nat) * x = x) & +(ALL x::nat. x * NUMERAL_BIT1 (0::nat) = x) & +(ALL (x::nat) xa::nat. Suc x * xa = x * xa + xa) & +(ALL (x::nat) xa::nat. x * Suc xa = x + x * xa)" + by (import hollight MULT_CLAUSES) + +lemma MULT_AC: "(m::nat) * (n::nat) = n * m & +m * n * (p::nat) = m * (n * p) & m * (n * p) = n * (m * p)" + by (import hollight MULT_AC) + +lemma MULT_2: "ALL n::nat. NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)) * n = n + n" + by (import hollight MULT_2) + +lemma MULT_EQ_1: "ALL (m::nat) n::nat. + (m * n = NUMERAL_BIT1 (0::nat)) = + (m = NUMERAL_BIT1 (0::nat) & n = NUMERAL_BIT1 (0::nat))" + by (import hollight MULT_EQ_1) + +constdefs + EXP :: "nat => nat => nat" + "EXP == +SOME EXP::nat => nat => nat. + (ALL m::nat. EXP m (0::nat) = NUMERAL_BIT1 (0::nat)) & + (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::nat) = NUMERAL_BIT1 (0::nat)) & + (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::nat)) = (m = (0::nat) & n ~= (0::nat))" + 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::nat)) n = NUMERAL_BIT1 (0::nat)" + by (import hollight EXP_ONE) + +lemma EXP_1: "ALL x::nat. EXP x (NUMERAL_BIT1 (0::nat)) = x" + by (import hollight EXP_1) + +lemma EXP_2: "ALL x::nat. EXP x (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))) = 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::nat) = (m = (0::nat))) & + (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::nat) = (m = (0::nat))) & + (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::nat) = 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::nat) = 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::nat))" + by (import hollight LE_0) + +lemma LT_0: "ALL x::nat. < (0::nat) (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)" + by (import hollight LT_ANTISYM) + +lemma LET_ANTISYM: "ALL (m::nat) n::nat. ~ (<= m n & < n m)" + by (import hollight LET_ANTISYM) + +lemma LTE_ANTISYM: "ALL (x::nat) xa::nat. ~ (< x xa & <= 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" + 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" + by (import hollight LTE_CASES) + +lemma LT_NZ: "ALL n::nat. < (0::nat) n = (n ~= (0::nat))" + 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)" + 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::nat) n" + by (import hollight LT_ADD) + +lemma LT_ADDR: "ALL (x::nat) xa::nat. < xa (x + 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::nat) (m * n) = (< (0::nat) m & < (0::nat) 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::nat) & < n p --> < (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::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::nat))" + by (import hollight LE_MULT_RCANCEL) + +lemma LT_MULT_LCANCEL: "ALL (m::nat) (n::nat) p::nat. < (m * n) (m * p) = (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::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)" + 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))" + 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))" + 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))" + by (import hollight num_WOP) + +lemma num_MAX: "ALL P::nat => 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))" + by (import hollight num_MAX) + +constdefs + EVEN :: "nat => bool" + "EVEN == +SOME EVEN::nat => bool. + EVEN (0::nat) = True & (ALL n::nat. EVEN (Suc n) = (~ EVEN n))" + +lemma DEF_EVEN: "EVEN = +(SOME EVEN::nat => bool. + EVEN (0::nat) = True & (ALL n::nat. EVEN (Suc n) = (~ EVEN n)))" + by (import hollight DEF_EVEN) + +constdefs + ODD :: "nat => bool" + "ODD == +SOME ODD::nat => bool. + ODD (0::nat) = False & (ALL n::nat. ODD (Suc n) = (~ ODD n))" + +lemma DEF_ODD: "ODD = +(SOME ODD::nat => bool. + ODD (0::nat) = False & (ALL n::nat. ODD (Suc n) = (~ ODD n)))" + by (import hollight DEF_ODD) + +lemma NOT_EVEN: "ALL n::nat. (~ EVEN n) = ODD n" + by (import hollight NOT_EVEN) + +lemma NOT_ODD: "ALL n::nat. (~ ODD n) = EVEN n" + by (import hollight NOT_ODD) + +lemma EVEN_OR_ODD: "ALL n::nat. EVEN n | ODD n" + by (import hollight EVEN_OR_ODD) + +lemma EVEN_AND_ODD: "ALL x::nat. ~ (EVEN x & 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::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)" + by (import hollight ODD_MULT) + +lemma ODD_EXP: "ALL (m::nat) n::nat. ODD (EXP m n) = (ODD m | n = (0::nat))" + by (import hollight ODD_EXP) + +lemma EVEN_DOUBLE: "ALL n::nat. EVEN (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)) * n)" + by (import hollight EVEN_DOUBLE) + +lemma ODD_DOUBLE: "ALL x::nat. ODD (Suc (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)) * x))" + by (import hollight ODD_DOUBLE) + +lemma EVEN_EXISTS_LEMMA: "ALL n::nat. + (EVEN n --> (EX m::nat. n = NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)) * m)) & + (~ EVEN n --> + (EX m::nat. n = Suc (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)) * m)))" + by (import hollight EVEN_EXISTS_LEMMA) + +lemma EVEN_EXISTS: "ALL n::nat. + EVEN n = (EX m::nat. n = NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)) * m)" + by (import hollight EVEN_EXISTS) + +lemma ODD_EXISTS: "ALL n::nat. + ODD n = (EX m::nat. n = Suc (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)) * 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::nat))) k * m) = + (n ~= (0::nat))" + by (import hollight EVEN_ODD_DECOMPOSITION) + +lemma SUB_0: "ALL x::nat. (0::nat) - x = (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" + by (import hollight SUB_PRESUC) + +lemma SUB_EQ_0: "ALL (m::nat) n::nat. (m - n = (0::nat)) = <= m n" + by (import hollight SUB_EQ_0) + +lemma ADD_SUBR: "ALL (x::nat) xa::nat. xa - (x + 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::nat) = x" + by (import hollight SUC_SUB1) + +constdefs + FACT :: "nat => nat" + "FACT == +SOME FACT::nat => nat. + FACT (0::nat) = NUMERAL_BIT1 (0::nat) & + (ALL n::nat. FACT (Suc n) = Suc n * FACT n)" + +lemma DEF_FACT: "FACT = +(SOME FACT::nat => nat. + FACT (0::nat) = NUMERAL_BIT1 (0::nat) & + (ALL n::nat. FACT (Suc n) = Suc n * FACT n))" + by (import hollight DEF_FACT) + +lemma FACT_LT: "ALL n::nat. < (0::nat) (FACT n)" + by (import hollight FACT_LT) + +lemma FACT_LE: "ALL x::nat. <= (NUMERAL_BIT1 (0::nat)) (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::nat) --> (EX (q::nat) r::nat. m = 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::nat)) (x = (0::nat) & xa = (0::nat)) + (m = x * n + xa & < xa n)" + by (import hollight DIVMOD_EXIST_0) + +constdefs + DIV :: "nat => nat => nat" + "DIV == +SOME q::nat => nat => nat. + EX r::nat => nat => nat. + ALL (m::nat) n::nat. + COND (n = (0::nat)) (q m n = (0::nat) & r m n = (0::nat)) + (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::nat)) (q m n = (0::nat) & r m n = (0::nat)) + (m = q m n * n + r m n & < (r m n) n))" + by (import hollight DEF_DIV) + +constdefs + MOD :: "nat => nat => nat" + "MOD == +SOME r::nat => nat => nat. + ALL (m::nat) n::nat. + COND (n = (0::nat)) (DIV m n = (0::nat) & r m n = (0::nat)) + (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::nat)) (DIV m n = (0::nat) & r m n = (0::nat)) + (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::nat) --> m = DIV m n * n + MOD m n & < (MOD m 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" + 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" + 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" + 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" + by (import hollight DIV_UNIQ) + +lemma MOD_MULT: "ALL (x::nat) xa::nat. x ~= (0::nat) --> MOD (x * xa) x = (0::nat)" + by (import hollight MOD_MULT) + +lemma DIV_MULT: "ALL (x::nat) xa::nat. x ~= (0::nat) --> DIV (x * xa) x = xa" + by (import hollight DIV_MULT) + +lemma DIV_DIV: "ALL (m::nat) (n::nat) p::nat. + n * p ~= (0::nat) --> 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" + by (import hollight MOD_EQ) + +lemma DIV_MOD: "ALL (m::nat) (n::nat) p::nat. + n * p ~= (0::nat) --> 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::nat)) = n" + by (import hollight DIV_1) + +lemma EXP_LT_0: "ALL (x::nat) xa::nat. + < (0::nat) (EXP xa x) = (xa ~= (0::nat) | x = (0::nat))" + by (import hollight EXP_LT_0) + +lemma DIV_LE: "ALL (m::nat) n::nat. n ~= (0::nat) --> <= (DIV m n) m" + by (import hollight DIV_LE) + +lemma DIV_MUL_LE: "ALL (m::nat) n::nat. <= (n * DIV m n) m" + by (import hollight DIV_MUL_LE) + +lemma DIV_0: "ALL n::nat. n ~= (0::nat) --> DIV (0::nat) n = (0::nat)" + by (import hollight DIV_0) + +lemma MOD_0: "ALL n::nat. n ~= (0::nat) --> MOD (0::nat) n = (0::nat)" + by (import hollight MOD_0) + +lemma DIV_LT: "ALL (m::nat) n::nat. < m n --> DIV m n = (0::nat)" + by (import hollight DIV_LT) + +lemma MOD_MOD: "ALL (m::nat) (n::nat) p::nat. + n * p ~= (0::nat) --> MOD (MOD m (n * p)) n = MOD m n" + by (import hollight MOD_MOD) + +lemma MOD_MOD_REFL: "ALL (m::nat) n::nat. n ~= (0::nat) --> MOD (MOD m n) n = MOD m n" + by (import hollight MOD_MOD_REFL) + +lemma DIV_MULT2: "ALL (x::nat) (xa::nat) xb::nat. + x * xb ~= (0::nat) --> DIV (x * xa) (x * xb) = DIV xa xb" + by (import hollight DIV_MULT2) + +lemma MOD_MULT2: "ALL (x::nat) (xa::nat) xb::nat. + x * xb ~= (0::nat) --> MOD (x * xa) (x * xb) = x * MOD xa xb" + by (import hollight MOD_MULT2) + +lemma MOD_1: "ALL n::nat. MOD n (NUMERAL_BIT1 (0::nat)) = (0::nat)" + by (import hollight MOD_1) + +lemma MOD_EXISTS: "ALL (m::nat) n::nat. + (EX q::nat. m = n * q) = + COND (n = (0::nat)) (m = (0::nat)) (MOD m 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::nat))) x & < m n | + x = (0::nat) & m ~= (0::nat) & n = (0::nat))" + by (import hollight LT_EXP) + +lemma LE_EXP: "ALL (x::nat) (m::nat) n::nat. + <= (EXP x m) (EXP x n) = + COND (x = (0::nat)) (m = (0::nat) --> n = (0::nat)) + (x = NUMERAL_BIT1 (0::nat) | <= m n)" + by (import hollight LE_EXP) + +lemma DIV_MONO: "ALL (m::nat) (n::nat) p::nat. + p ~= (0::nat) & <= m n --> <= (DIV m p) (DIV n p)" + by (import hollight DIV_MONO) + +lemma DIV_MONO_LT: "ALL (m::nat) (n::nat) p::nat. + p ~= (0::nat) & <= (m + p) n --> < (DIV m p) (DIV n p)" + by (import hollight DIV_MONO_LT) + +lemma LE_LDIV: "ALL (a::nat) (b::nat) n::nat. + a ~= (0::nat) & <= 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::nat) --> <= 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::nat) --> <= (DIV b a) n = < b (a * (n + NUMERAL_BIT1 (0::nat)))" + by (import hollight LE_LDIV_EQ) + +lemma DIV_EQ_0: "ALL (m::nat) n::nat. n ~= (0::nat) --> (DIV m n = (0::nat)) = < m n" + by (import hollight DIV_EQ_0) + +lemma MOD_EQ_0: "ALL (m::nat) n::nat. + n ~= (0::nat) --> (MOD m 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::nat))) = (0::nat))" + by (import hollight EVEN_MOD) + +lemma ODD_MOD: "ALL n::nat. + ODD n = + (MOD n (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))) = NUMERAL_BIT1 (0::nat))" + by (import hollight ODD_MOD) + +lemma MOD_MULT_RMOD: "ALL (m::nat) (n::nat) p::nat. + n ~= (0::nat) --> MOD (m * MOD p n) n = MOD (m * p) n" + by (import hollight MOD_MULT_RMOD) + +lemma MOD_MULT_LMOD: "ALL (x::nat) (xa::nat) xb::nat. + xa ~= (0::nat) --> MOD (MOD x xa * xb) xa = MOD (x * xb) xa" + by (import hollight MOD_MULT_LMOD) + +lemma MOD_MULT_MOD2: "ALL (x::nat) (xa::nat) xb::nat. + xa ~= (0::nat) --> MOD (MOD x xa * MOD xb xa) xa = MOD (x * xb) xa" + by (import hollight MOD_MULT_MOD2) + +lemma MOD_EXP_MOD: "ALL (m::nat) (n::nat) p::nat. + n ~= (0::nat) --> MOD (EXP (MOD m n) p) n = MOD (EXP m p) 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::nat) --> MOD (MOD a n + MOD b n) n = MOD (a + b) n" + by (import hollight MOD_ADD_MOD) + +lemma DIV_ADD_MOD: "ALL (a::nat) (b::nat) n::nat. + n ~= (0::nat) --> + (MOD (a + b) n = MOD a n + MOD b n) = (DIV (a + b) n = DIV a n + DIV b n)" + by (import hollight DIV_ADD_MOD) + +lemma DIV_REFL: "ALL n::nat. n ~= (0::nat) --> DIV n n = NUMERAL_BIT1 (0::nat)" + by (import hollight DIV_REFL) + +lemma MOD_LE: "ALL (m::nat) n::nat. n ~= (0::nat) --> <= (MOD m n) m" + by (import hollight MOD_LE) + +lemma DIV_MONO2: "ALL (m::nat) (n::nat) p::nat. + p ~= (0::nat) & <= p m --> <= (DIV n m) (DIV n p)" + by (import hollight DIV_MONO2) + +lemma DIV_LE_EXCLUSION: "ALL (a::nat) (b::nat) (c::nat) d::nat. + b ~= (0::nat) & < (b * c) ((a + NUMERAL_BIT1 (0::nat)) * d) --> + <= (DIV c d) (DIV a b)" + by (import hollight DIV_LE_EXCLUSION) + +lemma DIV_EQ_EXCLUSION: "< ((b::nat) * (c::nat)) (((a::nat) + NUMERAL_BIT1 (0::nat)) * (d::nat)) & +< (a * d) ((c + NUMERAL_BIT1 (0::nat)) * b) --> +DIV a b = DIV c d" + by (import hollight DIV_EQ_EXCLUSION) + +lemma SUB_ELIM_THM: "(P::nat => bool) ((a::nat) - (b::nat)) = +(ALL x::nat. (b = a + x --> P (0::nat)) & (a = b + x --> P x))" + by (import hollight SUB_ELIM_THM) + +lemma PRE_ELIM_THM: "(P::nat => bool) (Pred (n::nat)) = +(ALL m::nat. (n = (0::nat) --> P (0::nat)) & (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::nat) & P (0::nat) (0::nat) | + n ~= (0::nat) & (ALL (q::nat) r::nat. m = q * n + r & < r n --> P q r))" + by (import hollight DIVMOD_ELIM_THM) + +constdefs + eqeq :: "'q_9910::type +=> 'q_9909::type => ('q_9910::type => 'q_9909::type => bool) => bool" + "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_) + +constdefs + mod_nat :: "nat => nat => nat => bool" + "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) + +constdefs + minimal :: "(nat => bool) => nat" + "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))" + by (import hollight DEF_minimal) + +lemma MINIMAL: "ALL P::nat => bool. + Ex P = (P (minimal P) & (ALL x::nat. < x (minimal P) --> ~ P x))" + by (import hollight MINIMAL) + +constdefs + WF :: "('A::type => 'A::type => bool) => bool" + "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_354::'A::type => 'A::type => bool) = +(ALL P::'A::type => bool. + Ex P = (EX x::'A::type. P x & (ALL y::'A::type. u_354 y x --> ~ P y)))" + by (import hollight WF_EQ) + +lemma WF_IND: "WF (u_354::'A::type => 'A::type => bool) = +(ALL P::'A::type => bool. + (ALL x::'A::type. (ALL y::'A::type. u_354 y x --> P y) --> P x) --> + All P)" + by (import hollight WF_IND) + +lemma WF_DCHAIN: "WF (u_354::'A::type => 'A::type => bool) = +(~ (EX s::nat => 'A::type. ALL n::nat. u_354 (s (Suc n)) (s n)))" + by (import hollight WF_DCHAIN) + +lemma WF_UREC: "WF (u_354::'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_354 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))" + 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_354::'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_354" + by (import hollight WF_UREC_WF) + +lemma WF_REC_INVARIANT: "WF (u_354::'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_354 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))" + by (import hollight WF_REC_INVARIANT) + +lemma WF_REC: "WF (u_354::'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_354 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))" + 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_354::'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_354" + by (import hollight WF_REC_WF) + +lemma WF_EREC: "WF (u_354::'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_354 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))" + by (import hollight WF_EREC) + +lemma WF_SUBSET: "(ALL (x::'A::type) y::'A::type. + (u_354::'A::type => 'A::type => bool) x y --> + (u_473::'A::type => 'A::type => bool) x y) & +WF u_473 --> +WF u_354" + by (import hollight WF_SUBSET) + +lemma WF_MEASURE_GEN: "ALL m::'A::type => 'B::type. + WF (u_354::'B::type => 'B::type => bool) --> + WF (%(x::'A::type) x'::'A::type. u_354 (m 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)))))" + 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)))))" + by (import hollight WF_LEX) + +lemma WF_POINTWISE: "WF (u_354::'A::type => 'A::type => bool) & +WF (u_473::'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_354 x1 x2 & u_473 y1 y2)))))" + by (import hollight WF_POINTWISE) + +lemma WF_num: "WF <" + 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)" + by (import hollight WF_REC_num) + +consts + measure :: "('q_11107::type => nat) => 'q_11107::type => 'q_11107::type => 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)" + 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)" + by (import hollight MEASURE_LE) + +lemma WF_REFL: "ALL x::'A::type. WF (u_354::'A::type => 'A::type => bool) --> ~ u_354 x 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)" + 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_354::'A::type => 'A::type => bool) & + (ALL (f::'A::type => 'B::type) (g::'A::type => 'B::type) x::'A::type. + (ALL z::'A::type. u_354 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_354 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_354 y (G f x) --> u_354 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))" + by (import hollight WF_REC_TAIL_GENERAL) + +lemma ARITH_ZERO: "(0::nat) = (0::nat) & NUMERAL_BIT0 (0::nat) = (0::nat)" + by (import hollight ARITH_ZERO) + +lemma ARITH_SUC: "(ALL x::nat. Suc x = Suc x) & +Suc (0::nat) = NUMERAL_BIT1 (0::nat) & +(ALL x::nat. Suc (NUMERAL_BIT0 x) = NUMERAL_BIT1 x) & +(ALL x::nat. Suc (NUMERAL_BIT1 x) = NUMERAL_BIT0 (Suc x))" + by (import hollight ARITH_SUC) + +lemma ARITH_PRE: "(ALL x::nat. Pred x = Pred x) & +Pred (0::nat) = (0::nat) & +(ALL x::nat. + Pred (NUMERAL_BIT0 x) = + COND (x = (0::nat)) (0::nat) (NUMERAL_BIT1 (Pred x))) & +(ALL x::nat. Pred (NUMERAL_BIT1 x) = NUMERAL_BIT0 x)" + by (import hollight ARITH_PRE) + +lemma ARITH_ADD: "(ALL (x::nat) xa::nat. x + xa = x + xa) & +(0::nat) + (0::nat) = (0::nat) & +(ALL x::nat. (0::nat) + NUMERAL_BIT0 x = NUMERAL_BIT0 x) & +(ALL x::nat. (0::nat) + NUMERAL_BIT1 x = NUMERAL_BIT1 x) & +(ALL x::nat. NUMERAL_BIT0 x + (0::nat) = NUMERAL_BIT0 x) & +(ALL x::nat. NUMERAL_BIT1 x + (0::nat) = NUMERAL_BIT1 x) & +(ALL (x::nat) xa::nat. + NUMERAL_BIT0 x + NUMERAL_BIT0 xa = NUMERAL_BIT0 (x + xa)) & +(ALL (x::nat) xa::nat. + NUMERAL_BIT0 x + NUMERAL_BIT1 xa = NUMERAL_BIT1 (x + xa)) & +(ALL (x::nat) xa::nat. + NUMERAL_BIT1 x + NUMERAL_BIT0 xa = NUMERAL_BIT1 (x + xa)) & +(ALL (x::nat) xa::nat. + NUMERAL_BIT1 x + NUMERAL_BIT1 xa = NUMERAL_BIT0 (Suc (x + xa)))" + by (import hollight ARITH_ADD) + +lemma ARITH_MULT: "(ALL (x::nat) xa::nat. x * xa = x * xa) & +(0::nat) * (0::nat) = (0::nat) & +(ALL x::nat. (0::nat) * NUMERAL_BIT0 x = (0::nat)) & +(ALL x::nat. (0::nat) * NUMERAL_BIT1 x = (0::nat)) & +(ALL x::nat. NUMERAL_BIT0 x * (0::nat) = (0::nat)) & +(ALL x::nat. NUMERAL_BIT1 x * (0::nat) = (0::nat)) & +(ALL (x::nat) xa::nat. + NUMERAL_BIT0 x * NUMERAL_BIT0 xa = + NUMERAL_BIT0 (NUMERAL_BIT0 (x * xa))) & +(ALL (x::nat) xa::nat. + NUMERAL_BIT0 x * NUMERAL_BIT1 xa = + NUMERAL_BIT0 x + NUMERAL_BIT0 (NUMERAL_BIT0 (x * xa))) & +(ALL (x::nat) xa::nat. + NUMERAL_BIT1 x * NUMERAL_BIT0 xa = + NUMERAL_BIT0 xa + NUMERAL_BIT0 (NUMERAL_BIT0 (x * xa))) & +(ALL (x::nat) xa::nat. + NUMERAL_BIT1 x * NUMERAL_BIT1 xa = + NUMERAL_BIT1 x + + (NUMERAL_BIT0 xa + NUMERAL_BIT0 (NUMERAL_BIT0 (x * xa))))" + by (import hollight ARITH_MULT) + +lemma ARITH_EXP: "(ALL (x::nat) xa::nat. EXP x xa = EXP x xa) & +EXP (0::nat) (0::nat) = NUMERAL_BIT1 (0::nat) & +(ALL m::nat. EXP (NUMERAL_BIT0 m) (0::nat) = NUMERAL_BIT1 (0::nat)) & +(ALL m::nat. EXP (NUMERAL_BIT1 m) (0::nat) = NUMERAL_BIT1 (0::nat)) & +(ALL n::nat. + EXP (0::nat) (NUMERAL_BIT0 n) = EXP (0::nat) n * EXP (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) & +(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::nat) (NUMERAL_BIT1 n) = (0::nat)) & +(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)) & +(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))" + by (import hollight ARITH_EXP) + +lemma ARITH_EVEN: "(ALL x::nat. EVEN x = EVEN x) & +EVEN (0::nat) = True & +(ALL x::nat. EVEN (NUMERAL_BIT0 x) = True) & +(ALL x::nat. EVEN (NUMERAL_BIT1 x) = False)" + by (import hollight ARITH_EVEN) + +lemma ARITH_ODD: "(ALL x::nat. ODD x = ODD x) & +ODD (0::nat) = False & +(ALL x::nat. ODD (NUMERAL_BIT0 x) = False) & +(ALL x::nat. ODD (NUMERAL_BIT1 x) = True)" + by (import hollight ARITH_ODD) + +lemma ARITH_LE: "(ALL (x::nat) xa::nat. <= x xa = <= x xa) & +<= (0::nat) (0::nat) = True & +(ALL x::nat. <= (NUMERAL_BIT0 x) (0::nat) = (x = (0::nat))) & +(ALL x::nat. <= (NUMERAL_BIT1 x) (0::nat) = False) & +(ALL x::nat. <= (0::nat) (NUMERAL_BIT0 x) = True) & +(ALL x::nat. <= (0::nat) (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)" + by (import hollight ARITH_LE) + +lemma ARITH_LT: "(ALL (x::nat) xa::nat. < x xa = < x xa) & +< (0::nat) (0::nat) = False & +(ALL x::nat. < (NUMERAL_BIT0 x) (0::nat) = False) & +(ALL x::nat. < (NUMERAL_BIT1 x) (0::nat) = False) & +(ALL x::nat. < (0::nat) (NUMERAL_BIT0 x) = < (0::nat) x) & +(ALL x::nat. < (0::nat) (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)" + by (import hollight ARITH_LT) + +lemma ARITH_EQ: "(ALL (x::nat) xa::nat. (x = xa) = (x = xa)) & +((0::nat) = (0::nat)) = True & +(ALL x::nat. (NUMERAL_BIT0 x = (0::nat)) = (x = (0::nat))) & +(ALL x::nat. (NUMERAL_BIT1 x = (0::nat)) = False) & +(ALL x::nat. ((0::nat) = NUMERAL_BIT0 x) = ((0::nat) = x)) & +(ALL x::nat. ((0::nat) = NUMERAL_BIT1 x) = False) & +(ALL (x::nat) xa::nat. (NUMERAL_BIT0 x = NUMERAL_BIT0 xa) = (x = xa)) & +(ALL (x::nat) xa::nat. (NUMERAL_BIT0 x = NUMERAL_BIT1 xa) = False) & +(ALL (x::nat) xa::nat. (NUMERAL_BIT1 x = NUMERAL_BIT0 xa) = False) & +(ALL (x::nat) xa::nat. (NUMERAL_BIT1 x = NUMERAL_BIT1 xa) = (x = xa))" + by (import hollight ARITH_EQ) + +lemma ARITH_SUB: "(ALL (x::nat) xa::nat. x - xa = x - xa) & +(0::nat) - (0::nat) = (0::nat) & +(ALL x::nat. (0::nat) - NUMERAL_BIT0 x = (0::nat)) & +(ALL x::nat. (0::nat) - NUMERAL_BIT1 x = (0::nat)) & +(ALL x::nat. NUMERAL_BIT0 x - (0::nat) = NUMERAL_BIT0 x) & +(ALL x::nat. NUMERAL_BIT1 x - (0::nat) = NUMERAL_BIT1 x) & +(ALL (m::nat) n::nat. + NUMERAL_BIT0 m - NUMERAL_BIT0 n = NUMERAL_BIT0 (m - n)) & +(ALL (m::nat) n::nat. + NUMERAL_BIT0 m - NUMERAL_BIT1 n = Pred (NUMERAL_BIT0 (m - n))) & +(ALL (m::nat) n::nat. + NUMERAL_BIT1 m - NUMERAL_BIT0 n = + COND (<= n m) (NUMERAL_BIT1 (m - n)) (0::nat)) & +(ALL (m::nat) n::nat. + NUMERAL_BIT1 m - NUMERAL_BIT1 n = NUMERAL_BIT0 (m - n))" + by (import hollight ARITH_SUB) + +lemma right_th: "(s::nat) * NUMERAL_BIT1 (x::nat) = s + NUMERAL_BIT0 (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::nat) = 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::nat))) & +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::nat) = r1 & +pwr x (NUMERAL_BIT1 (0::nat)) = x & +mul x (add y (z::'A::type)) = 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::nat) + 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::nat) * x = x) & +(ALL x::nat. (0::nat) * x = (0::nat)) & +(ALL (x::nat) (xa::nat) xb::nat. x * (xa + xb) = x * xa + x * xb) & +(ALL x::nat. EXP x (0::nat) = NUMERAL_BIT1 (0::nat)) & +(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)" + by (import hollight NUM_INTEGRAL_LEMMA) + +lemma NUM_INTEGRAL: "(ALL x::nat. (0::nat) * x = (0::nat)) & +(ALL (x::nat) (xa::nat) xb::nat. (x + xa = x + xb) = (xa = xb)) & +(ALL (w::nat) (x::nat) (y::nat) z::nat. + (w * y + x * z = w * z + x * y) = (w = x | y = z))" + by (import hollight NUM_INTEGRAL) + +lemma INJ_INVERSE2: "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)" + by (import hollight INJ_INVERSE2) + +constdefs + NUMPAIR :: "nat => nat => nat" + "NUMPAIR == +%(u::nat) ua::nat. + EXP (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))) u * + (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)) * ua + NUMERAL_BIT1 (0::nat))" + +lemma DEF_NUMPAIR: "NUMPAIR = +(%(u::nat) ua::nat. + EXP (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))) u * + (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)) * ua + NUMERAL_BIT1 (0::nat)))" + 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" + 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)" + by (import hollight NUMPAIR_INJ) + +constdefs + NUMFST :: "nat => nat" + "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)" + by (import hollight DEF_NUMFST) + +constdefs + NUMSND :: "nat => nat" + "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)" + by (import hollight DEF_NUMSND) + +constdefs + NUMSUM :: "bool => nat => nat" + "NUMSUM == +%(u::bool) ua::nat. + COND u (Suc (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)) * ua)) + (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)) * ua)" + +lemma DEF_NUMSUM: "NUMSUM = +(%(u::bool) ua::nat. + COND u (Suc (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)) * ua)) + (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)) * 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)" + by (import hollight NUMSUM_INJ) + +constdefs + NUMLEFT :: "nat => bool" + "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)" + by (import hollight DEF_NUMLEFT) + +constdefs + NUMRIGHT :: "nat => nat" + "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)" + by (import hollight DEF_NUMRIGHT) + +constdefs + INJN :: "nat => nat => 'A::type => bool" + "INJN == %(u::nat) (n::nat) a::'A::type. n = u" + +lemma DEF_INJN: "INJN = (%(u::nat) (n::nat) a::'A::type. 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)))" + by (import hollight INJN_INJ) + +constdefs + INJA :: "'A::type => nat => 'A::type => bool" + "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)" + by (import hollight DEF_INJA) + +lemma INJA_INJ: "ALL (a1::'A::type) a2::'A::type. (INJA a1 = INJA a2) = (a1 = a2)" + by (import hollight INJA_INJ) + +constdefs + INJF :: "(nat => nat => 'A::type => bool) => nat => 'A::type => bool" + "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))" + 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)" + by (import hollight INJF_INJ) + +constdefs + INJP :: "(nat => 'A::type => bool) +=> (nat => 'A::type => bool) => nat => 'A::type => bool" + "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)" + +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))" + 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')" + by (import hollight INJP_INJ) + +constdefs + ZCONSTR :: "nat +=> 'A::type => (nat => nat => 'A::type => bool) => nat => 'A::type => bool" + "ZCONSTR == +%(u::nat) (ua::'A::type) ub::nat => nat => 'A::type => 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. + INJP (INJN (Suc u)) (INJP (INJA ua) (INJF ub)))" + by (import hollight DEF_ZCONSTR) + +constdefs + ZBOT :: "nat => 'A::type => bool" + "ZBOT == INJP (INJN (0::nat)) (SOME z::nat => 'A::type => bool. True)" + +lemma DEF_ZBOT: "ZBOT = INJP (INJN (0::nat)) (SOME z::nat => 'A::type => 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" + by (import hollight ZCONSTR_ZBOT) + +constdefs + ZRECSPACE :: "(nat => 'A::type => bool) => bool" + "ZRECSPACE == +%a::nat => 'A::type => bool. + ALL ZRECSPACE'::(nat => 'A::type => bool) => bool. + (ALL a::nat => 'A::type => bool. + a = ZBOT | + (EX (c::nat) (i::'A::type) r::nat => nat => 'A::type => 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 = ZBOT | + (EX (c::nat) (i::'A::type) r::nat => nat => 'A::type => 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"]) + by (import hollight TYDEF_recspace) + +syntax + "_dest_rec" :: _ ("'_dest'_rec") + +syntax + "_mk_rec" :: _ ("'_mk'_rec") + +lemmas "TYDEF_recspace_@intern" = typedef_hol2hollight + [where a="a :: 'A::type recspace" and r=r , + OF type_definition_recspace] + +constdefs + BOTTOM :: "'A::type recspace" + "BOTTOM == _mk_rec ZBOT" + +lemma DEF_BOTTOM: "BOTTOM = _mk_rec ZBOT" + by (import hollight DEF_BOTTOM) + +constdefs + CONSTR :: "nat => 'A::type => (nat => 'A::type recspace) => 'A::type recspace" + "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 (x::nat => 'A::type => bool) y::nat => 'A::type => bool. + _mk_rec x = _mk_rec y --> 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" + 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)" + 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" + 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))" + by (import hollight CONSTR_REC) + +constdefs + FCONS :: "'A::type => (nat => 'A::type) => nat => 'A::type" + "FCONS == +SOME FCONS::'A::type => (nat => 'A::type) => nat => 'A::type. + (ALL (a::'A::type) f::nat => 'A::type. FCONS a f (0::nat) = a) & + (ALL (a::'A::type) (f::nat => 'A::type) 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::nat) = a) & + (ALL (a::'A::type) (f::nat => 'A::type) 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::nat)) (f o Suc)" + by (import hollight FCONS_UNDO) + +constdefs + FNIL :: "nat => 'A::type" + "FNIL == %u::nat. SOME x::'A::type. True" + +lemma DEF_FNIL: "FNIL = (%u::nat. SOME x::'A::type. True)" + by (import hollight DEF_FNIL) + +typedef (open) ('A, 'B) sum = "(Collect::(('A::type * 'B::type) recspace => bool) + => ('A::type * 'B::type) recspace set) + (%a::('A::type * 'B::type) recspace. + (All::((('A::type * 'B::type) recspace => bool) => bool) => bool) + (%sum'::('A::type * 'B::type) recspace => bool. + (op -->::bool => bool => bool) + ((All::(('A::type * 'B::type) recspace => bool) => bool) + (%a::('A::type * 'B::type) recspace. + (op -->::bool => bool => bool) + ((op |::bool => bool => bool) + ((Ex::('A::type => bool) => bool) + (%aa::'A::type. + (op =::('A::type * 'B::type) recspace + => ('A::type * 'B::type) recspace => bool) + a ((CONSTR::nat + => 'A::type * 'B::type + => (nat => ('A::type * 'B::type) recspace) + => ('A::type * 'B::type) recspace) + ((NUMERAL::nat => nat) (0::nat)) + ((Pair::'A::type + => 'B::type => 'A::type * 'B::type) + aa ((Eps::('B::type => bool) => 'B::type) +(%v::'B::type. True::bool))) + (%n::nat. + BOTTOM::('A::type * + 'B::type) recspace)))) + ((Ex::('B::type => bool) => bool) + (%aa::'B::type. + (op =::('A::type * 'B::type) recspace + => ('A::type * 'B::type) recspace => bool) + a ((CONSTR::nat + => 'A::type * 'B::type + => (nat => ('A::type * 'B::type) recspace) + => ('A::type * 'B::type) recspace) + ((Suc::nat => nat) + ((NUMERAL::nat => nat) (0::nat))) + ((Pair::'A::type + => 'B::type => 'A::type * 'B::type) + ((Eps::('A::type => bool) => 'A::type) + (%v::'A::type. True::bool)) + aa) + (%n::nat. + BOTTOM::('A::type * + 'B::type) recspace))))) + (sum' a))) + (sum' a)))" morphisms "_dest_sum" "_mk_sum" + apply (rule light_ex_imp_nonempty[where t="(CONSTR::nat + => 'A::type * 'B::type + => (nat => ('A::type * 'B::type) recspace) + => ('A::type * 'B::type) recspace) + ((NUMERAL::nat => nat) (0::nat)) + ((Pair::'A::type => 'B::type => 'A::type * 'B::type) (a::'A::type) + ((Eps::('B::type => bool) => 'B::type) (%v::'B::type. True::bool))) + (%n::nat. BOTTOM::('A::type * 'B::type) recspace)"]) + by (import hollight TYDEF_sum) + +syntax + "_dest_sum" :: _ ("'_dest'_sum") + +syntax + "_mk_sum" :: _ ("'_mk'_sum") + +lemmas "TYDEF_sum_@intern" = typedef_hol2hollight + [where a="a :: ('A::type, 'B::type) sum" and r=r , + OF type_definition_sum] + +constdefs + INL :: "'A::type => ('A::type, 'B::type) sum" + "INL == +%a::'A::type. + _mk_sum (CONSTR (0::nat) (a, SOME v::'B::type. True) (%n::nat. BOTTOM))" + +lemma DEF_INL: "INL = +(%a::'A::type. + _mk_sum (CONSTR (0::nat) (a, SOME v::'B::type. True) (%n::nat. BOTTOM)))" + by (import hollight DEF_INL) + +constdefs + INR :: "'B::type => ('A::type, 'B::type) sum" + "INR == +%a::'B::type. + _mk_sum + (CONSTR (Suc (0::nat)) (SOME v::'A::type. True, a) (%n::nat. BOTTOM))" + +lemma DEF_INR: "INR = +(%a::'B::type. + _mk_sum + (CONSTR (Suc (0::nat)) (SOME v::'A::type. True, a) (%n::nat. BOTTOM)))" + by (import hollight DEF_INR) + +consts + OUTL :: "('A::type, 'B::type) sum => 'A::type" + +defs + OUTL_def: "hollight.OUTL == +SOME OUTL::('A::type, 'B::type) sum => 'A::type. + ALL x::'A::type. OUTL (INL x) = x" + +lemma DEF_OUTL: "hollight.OUTL = +(SOME OUTL::('A::type, 'B::type) sum => 'A::type. + ALL x::'A::type. OUTL (INL x) = x)" + by (import hollight DEF_OUTL) + +consts + OUTR :: "('A::type, 'B::type) sum => 'B::type" + +defs + OUTR_def: "hollight.OUTR == +SOME OUTR::('A::type, 'B::type) sum => 'B::type. + ALL y::'B::type. OUTR (INR y) = y" + +lemma DEF_OUTR: "hollight.OUTR = +(SOME OUTR::('A::type, 'B::type) sum => '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) + +syntax + "_dest_option" :: _ ("'_dest'_option") + +syntax + "_mk_option" :: _ ("'_mk'_option") + +lemmas "TYDEF_option_@intern" = typedef_hol2hollight + [where a="a :: 'A::type hollight.option" and r=r , + OF type_definition_option] + +constdefs + NONE :: "'A::type hollight.option" + "NONE == +_mk_option (CONSTR (0::nat) (SOME v::'A::type. True) (%n::nat. BOTTOM))" + +lemma DEF_NONE: "NONE = +_mk_option (CONSTR (0::nat) (SOME v::'A::type. True) (%n::nat. BOTTOM))" + by (import hollight DEF_NONE) + +consts + SOME :: "'A::type => 'A::type hollight.option" ("SOME") + +defs + SOME_def: "SOME == %a::'A::type. _mk_option (CONSTR (Suc (0::nat)) a (%n::nat. BOTTOM))" + +lemma DEF_SOME: "SOME = +(%a::'A::type. _mk_option (CONSTR (Suc (0::nat)) a (%n::nat. BOTTOM)))" + 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::type hollight.list" and r=r , + OF type_definition_list] + +constdefs + NIL :: "'A::type hollight.list" + "NIL == _mk_list (CONSTR (0::nat) (SOME v::'A::type. True) (%n::nat. BOTTOM))" + +lemma DEF_NIL: "NIL = _mk_list (CONSTR (0::nat) (SOME v::'A::type. True) (%n::nat. BOTTOM))" + by (import hollight DEF_NIL) + +constdefs + CONS :: "'A::type => 'A::type hollight.list => 'A::type hollight.list" + "CONS == +%(a0::'A::type) a1::'A::type hollight.list. + _mk_list + (CONSTR (Suc (0::nat)) a0 (FCONS (_dest_list a1) (%n::nat. BOTTOM)))" + +lemma DEF_CONS: "CONS = +(%(a0::'A::type) a1::'A::type hollight.list. + _mk_list + (CONSTR (Suc (0::nat)) a0 (FCONS (_dest_list a1) (%n::nat. BOTTOM))))" + 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::nat) = e & (ALL n::nat. fn (Suc n) = f n (fn n))" + by (import hollight num_RECURSION_STD) + +constdefs + ISO :: "('A::type => 'B::type) => ('B::type => 'A::type) => bool" + "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_16585::type => 'q_16582::type) + (g::'q_16582::type => 'q_16585::type) --> +(ALL P::'q_16585::type => bool. All P = (ALL x::'q_16582::type. P (g x))) & +(ALL P::'q_16585::type => bool. Ex P = (EX x::'q_16582::type. P (g x))) & +(ALL (a::'q_16585::type) b::'q_16582::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::nat)) (SOME x::bool. True) (%n::nat. BOTTOM) | + a = + CONSTR (Suc (NUMERAL (0::nat))) (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::nat)) (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 + "_10288" :: "N_2" ("'_10288") + +defs + "_10288_def": "_10288 == _mk_2 (CONSTR (0::nat) (SOME x::bool. True) (%n::nat. BOTTOM))" + +lemma DEF__10288: "_10288 = _mk_2 (CONSTR (0::nat) (SOME x::bool. True) (%n::nat. BOTTOM))" + by (import hollight DEF__10288) + +consts + "_10289" :: "N_2" ("'_10289") + +defs + "_10289_def": "_10289 == +_mk_2 (CONSTR (Suc (0::nat)) (SOME x::bool. True) (%n::nat. BOTTOM))" + +lemma DEF__10289: "_10289 = +_mk_2 (CONSTR (Suc (0::nat)) (SOME x::bool. True) (%n::nat. BOTTOM))" + by (import hollight DEF__10289) + +constdefs + two_1 :: "N_2" + "two_1 == _10288" + +lemma DEF_two_1: "two_1 = _10288" + by (import hollight DEF_two_1) + +constdefs + two_2 :: "N_2" + "two_2 == _10289" + +lemma DEF_two_2: "two_2 = _10289" + 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::nat)) (SOME x::bool. True) (%n::nat. BOTTOM) | + a = + CONSTR (Suc (NUMERAL (0::nat))) (SOME x::bool. True) + (%n::nat. BOTTOM) | + a = + CONSTR (Suc (Suc (NUMERAL (0::nat)))) (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::nat)) (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 + "_10312" :: "N_3" ("'_10312") + +defs + "_10312_def": "_10312 == _mk_3 (CONSTR (0::nat) (SOME x::bool. True) (%n::nat. BOTTOM))" + +lemma DEF__10312: "_10312 = _mk_3 (CONSTR (0::nat) (SOME x::bool. True) (%n::nat. BOTTOM))" + by (import hollight DEF__10312) + +consts + "_10313" :: "N_3" ("'_10313") + +defs + "_10313_def": "_10313 == +_mk_3 (CONSTR (Suc (0::nat)) (SOME x::bool. True) (%n::nat. BOTTOM))" + +lemma DEF__10313: "_10313 = +_mk_3 (CONSTR (Suc (0::nat)) (SOME x::bool. True) (%n::nat. BOTTOM))" + by (import hollight DEF__10313) + +consts + "_10314" :: "N_3" ("'_10314") + +defs + "_10314_def": "_10314 == +_mk_3 (CONSTR (Suc (Suc (0::nat))) (SOME x::bool. True) (%n::nat. BOTTOM))" + +lemma DEF__10314: "_10314 = +_mk_3 (CONSTR (Suc (Suc (0::nat))) (SOME x::bool. True) (%n::nat. BOTTOM))" + by (import hollight DEF__10314) + +constdefs + three_1 :: "N_3" + "three_1 == _10312" + +lemma DEF_three_1: "three_1 = _10312" + by (import hollight DEF_three_1) + +constdefs + three_2 :: "N_3" + "three_2 == _10313" + +lemma DEF_three_2: "three_2 = _10313" + by (import hollight DEF_three_2) + +constdefs + three_3 :: "N_3" + "three_3 == _10314" + +lemma DEF_three_3: "three_3 = _10314" + 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) + +constdefs + HD :: "'A::type hollight.list => 'A::type" + "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) + +constdefs + TL :: "'A::type hollight.list => 'A::type hollight.list" + "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) + +constdefs + APPEND :: "'A::type hollight.list => 'A::type hollight.list => 'A::type hollight.list" + "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) + +constdefs + REVERSE :: "'A::type hollight.list => 'A::type hollight.list" + "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) + +constdefs + LENGTH :: "'A::type hollight.list => nat" + "LENGTH == +SOME LENGTH::'A::type hollight.list => nat. + LENGTH NIL = (0::nat) & + (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::nat) & + (ALL (h::'A::type) t::'A::type hollight.list. + LENGTH (CONS h t) = Suc (LENGTH t)))" + by (import hollight DEF_LENGTH) + +constdefs + MAP :: "('A::type => 'B::type) => 'A::type hollight.list => 'B::type hollight.list" + "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) + +constdefs + LAST :: "'A::type hollight.list => 'A::type" + "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) + +constdefs + REPLICATE :: "nat => 'q_16809::type => 'q_16809::type hollight.list" + "REPLICATE == +SOME REPLICATE::nat => 'q_16809::type => 'q_16809::type hollight.list. + (ALL x::'q_16809::type. REPLICATE (0::nat) x = NIL) & + (ALL (n::nat) x::'q_16809::type. + REPLICATE (Suc n) x = CONS x (REPLICATE n x))" + +lemma DEF_REPLICATE: "REPLICATE = +(SOME REPLICATE::nat => 'q_16809::type => 'q_16809::type hollight.list. + (ALL x::'q_16809::type. REPLICATE (0::nat) x = NIL) & + (ALL (n::nat) x::'q_16809::type. + REPLICATE (Suc n) x = CONS x (REPLICATE n x)))" + by (import hollight DEF_REPLICATE) + +constdefs + NULL :: "'q_16824::type hollight.list => bool" + "NULL == +SOME NULL::'q_16824::type hollight.list => bool. + NULL NIL = True & + (ALL (h::'q_16824::type) t::'q_16824::type hollight.list. + NULL (CONS h t) = False)" + +lemma DEF_NULL: "NULL = +(SOME NULL::'q_16824::type hollight.list => bool. + NULL NIL = True & + (ALL (h::'q_16824::type) t::'q_16824::type hollight.list. + NULL (CONS h t) = False))" + by (import hollight DEF_NULL) + +constdefs + ALL_list :: "('q_16844::type => bool) => 'q_16844::type hollight.list => bool" + "ALL_list == +SOME u::('q_16844::type => bool) => 'q_16844::type hollight.list => bool. + (ALL P::'q_16844::type => bool. u P NIL = True) & + (ALL (h::'q_16844::type) (P::'q_16844::type => bool) + t::'q_16844::type hollight.list. u P (CONS h t) = (P h & u P t))" + +lemma DEF_ALL: "ALL_list = +(SOME u::('q_16844::type => bool) => 'q_16844::type hollight.list => bool. + (ALL P::'q_16844::type => bool. u P NIL = True) & + (ALL (h::'q_16844::type) (P::'q_16844::type => bool) + t::'q_16844::type hollight.list. u P (CONS h t) = (P h & u P t)))" + by (import hollight DEF_ALL) + +consts + EX :: "('q_16865::type => bool) => 'q_16865::type hollight.list => bool" ("EX") + +defs + EX_def: "EX == +SOME u::('q_16865::type => bool) => 'q_16865::type hollight.list => bool. + (ALL P::'q_16865::type => bool. u P NIL = False) & + (ALL (h::'q_16865::type) (P::'q_16865::type => bool) + t::'q_16865::type hollight.list. u P (CONS h t) = (P h | u P t))" + +lemma DEF_EX: "EX = +(SOME u::('q_16865::type => bool) => 'q_16865::type hollight.list => bool. + (ALL P::'q_16865::type => bool. u P NIL = False) & + (ALL (h::'q_16865::type) (P::'q_16865::type => bool) + t::'q_16865::type hollight.list. u P (CONS h t) = (P h | u P t)))" + by (import hollight DEF_EX) + +constdefs + ITLIST :: "('q_16888::type => 'q_16887::type => 'q_16887::type) +=> 'q_16888::type hollight.list => 'q_16887::type => 'q_16887::type" + "ITLIST == +SOME ITLIST::('q_16888::type => 'q_16887::type => 'q_16887::type) + => 'q_16888::type hollight.list + => 'q_16887::type => 'q_16887::type. + (ALL (f::'q_16888::type => 'q_16887::type => 'q_16887::type) + b::'q_16887::type. ITLIST f NIL b = b) & + (ALL (h::'q_16888::type) + (f::'q_16888::type => 'q_16887::type => 'q_16887::type) + (t::'q_16888::type hollight.list) b::'q_16887::type. + ITLIST f (CONS h t) b = f h (ITLIST f t b))" + +lemma DEF_ITLIST: "ITLIST = +(SOME ITLIST::('q_16888::type => 'q_16887::type => 'q_16887::type) + => 'q_16888::type hollight.list + => 'q_16887::type => 'q_16887::type. + (ALL (f::'q_16888::type => 'q_16887::type => 'q_16887::type) + b::'q_16887::type. ITLIST f NIL b = b) & + (ALL (h::'q_16888::type) + (f::'q_16888::type => 'q_16887::type => 'q_16887::type) + (t::'q_16888::type hollight.list) b::'q_16887::type. + ITLIST f (CONS h t) b = f h (ITLIST f t b)))" + by (import hollight DEF_ITLIST) + +constdefs + MEM :: "'q_16913::type => 'q_16913::type hollight.list => bool" + "MEM == +SOME MEM::'q_16913::type => 'q_16913::type hollight.list => bool. + (ALL x::'q_16913::type. MEM x NIL = False) & + (ALL (h::'q_16913::type) (x::'q_16913::type) + t::'q_16913::type hollight.list. + MEM x (CONS h t) = (x = h | MEM x t))" + +lemma DEF_MEM: "MEM = +(SOME MEM::'q_16913::type => 'q_16913::type hollight.list => bool. + (ALL x::'q_16913::type. MEM x NIL = False) & + (ALL (h::'q_16913::type) (x::'q_16913::type) + t::'q_16913::type hollight.list. + MEM x (CONS h t) = (x = h | MEM x t)))" + by (import hollight DEF_MEM) + +constdefs + ALL2 :: "('q_16946::type => 'q_16953::type => bool) +=> 'q_16946::type hollight.list => 'q_16953::type hollight.list => bool" + "ALL2 == +SOME ALL2::('q_16946::type => 'q_16953::type => bool) + => 'q_16946::type hollight.list + => 'q_16953::type hollight.list => bool. + (ALL (P::'q_16946::type => 'q_16953::type => bool) + l2::'q_16953::type hollight.list. ALL2 P NIL l2 = (l2 = NIL)) & + (ALL (h1::'q_16946::type) (P::'q_16946::type => 'q_16953::type => bool) + (t1::'q_16946::type hollight.list) l2::'q_16953::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_16946::type => 'q_16953::type => bool) + => 'q_16946::type hollight.list + => 'q_16953::type hollight.list => bool. + (ALL (P::'q_16946::type => 'q_16953::type => bool) + l2::'q_16953::type hollight.list. ALL2 P NIL l2 = (l2 = NIL)) & + (ALL (h1::'q_16946::type) (P::'q_16946::type => 'q_16953::type => bool) + (t1::'q_16946::type hollight.list) l2::'q_16953::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_17008::type => 'q_17007::type => bool) NIL NIL = True & +ALL2 P (CONS (h1::'q_17008::type) (t1::'q_17008::type hollight.list)) NIL = +False & +ALL2 P NIL (CONS (h2::'q_17007::type) (t2::'q_17007::type hollight.list)) = +False & +ALL2 P (CONS h1 t1) (CONS h2 t2) = (P h1 h2 & ALL2 P t1 t2)" + by (import hollight ALL2) + +constdefs + MAP2 :: "('q_17038::type => 'q_17045::type => 'q_17035::type) +=> 'q_17038::type hollight.list + => 'q_17045::type hollight.list => 'q_17035::type hollight.list" + "MAP2 == +SOME MAP2::('q_17038::type => 'q_17045::type => 'q_17035::type) + => 'q_17038::type hollight.list + => 'q_17045::type hollight.list + => 'q_17035::type hollight.list. + (ALL (f::'q_17038::type => 'q_17045::type => 'q_17035::type) + l::'q_17045::type hollight.list. MAP2 f NIL l = NIL) & + (ALL (h1::'q_17038::type) + (f::'q_17038::type => 'q_17045::type => 'q_17035::type) + (t1::'q_17038::type hollight.list) l::'q_17045::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_17038::type => 'q_17045::type => 'q_17035::type) + => 'q_17038::type hollight.list + => 'q_17045::type hollight.list + => 'q_17035::type hollight.list. + (ALL (f::'q_17038::type => 'q_17045::type => 'q_17035::type) + l::'q_17045::type hollight.list. MAP2 f NIL l = NIL) & + (ALL (h1::'q_17038::type) + (f::'q_17038::type => 'q_17045::type => 'q_17035::type) + (t1::'q_17038::type hollight.list) l::'q_17045::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_17080::type => 'q_17079::type => 'q_17086::type) NIL NIL = NIL & +MAP2 f (CONS (h1::'q_17080::type) (t1::'q_17080::type hollight.list)) + (CONS (h2::'q_17079::type) (t2::'q_17079::type hollight.list)) = +CONS (f h1 h2) (MAP2 f t1 t2)" + by (import hollight MAP2) + +constdefs + EL :: "nat => 'q_17106::type hollight.list => 'q_17106::type" + "EL == +SOME EL::nat => 'q_17106::type hollight.list => 'q_17106::type. + (ALL l::'q_17106::type hollight.list. EL (0::nat) l = HD l) & + (ALL (n::nat) l::'q_17106::type hollight.list. + EL (Suc n) l = EL n (TL l))" + +lemma DEF_EL: "EL = +(SOME EL::nat => 'q_17106::type hollight.list => 'q_17106::type. + (ALL l::'q_17106::type hollight.list. EL (0::nat) l = HD l) & + (ALL (n::nat) l::'q_17106::type hollight.list. + EL (Suc n) l = EL n (TL l)))" + by (import hollight DEF_EL) + +constdefs + FILTER :: "('q_17131::type => bool) +=> 'q_17131::type hollight.list => 'q_17131::type hollight.list" + "FILTER == +SOME FILTER::('q_17131::type => bool) + => 'q_17131::type hollight.list + => 'q_17131::type hollight.list. + (ALL P::'q_17131::type => bool. FILTER P NIL = NIL) & + (ALL (h::'q_17131::type) (P::'q_17131::type => bool) + t::'q_17131::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_17131::type => bool) + => 'q_17131::type hollight.list + => 'q_17131::type hollight.list. + (ALL P::'q_17131::type => bool. FILTER P NIL = NIL) & + (ALL (h::'q_17131::type) (P::'q_17131::type => bool) + t::'q_17131::type hollight.list. + FILTER P (CONS h t) = + COND (P h) (CONS h (FILTER P t)) (FILTER P t)))" + by (import hollight DEF_FILTER) + +constdefs + ASSOC :: "'q_17160::type +=> ('q_17160::type * 'q_17154::type) hollight.list => 'q_17154::type" + "ASSOC == +SOME ASSOC::'q_17160::type + => ('q_17160::type * 'q_17154::type) hollight.list + => 'q_17154::type. + ALL (h::'q_17160::type * 'q_17154::type) (a::'q_17160::type) + t::('q_17160::type * 'q_17154::type) hollight.list. + ASSOC a (CONS h t) = COND (fst h = a) (snd h) (ASSOC a t)" + +lemma DEF_ASSOC: "ASSOC = +(SOME ASSOC::'q_17160::type + => ('q_17160::type * 'q_17154::type) hollight.list + => 'q_17154::type. + ALL (h::'q_17160::type * 'q_17154::type) (a::'q_17160::type) + t::('q_17160::type * 'q_17154::type) hollight.list. + ASSOC a (CONS h t) = COND (fst h = a) (snd h) (ASSOC a t))" + by (import hollight DEF_ASSOC) + +constdefs + ITLIST2 :: "('q_17184::type => 'q_17192::type => 'q_17182::type => 'q_17182::type) +=> 'q_17184::type hollight.list + => 'q_17192::type hollight.list => 'q_17182::type => 'q_17182::type" + "ITLIST2 == +SOME ITLIST2::('q_17184::type + => 'q_17192::type => 'q_17182::type => 'q_17182::type) + => 'q_17184::type hollight.list + => 'q_17192::type hollight.list + => 'q_17182::type => 'q_17182::type. + (ALL (f::'q_17184::type + => 'q_17192::type => 'q_17182::type => 'q_17182::type) + (l2::'q_17192::type hollight.list) b::'q_17182::type. + ITLIST2 f NIL l2 b = b) & + (ALL (h1::'q_17184::type) + (f::'q_17184::type + => 'q_17192::type => 'q_17182::type => 'q_17182::type) + (t1::'q_17184::type hollight.list) (l2::'q_17192::type hollight.list) + b::'q_17182::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_17184::type + => 'q_17192::type => 'q_17182::type => 'q_17182::type) + => 'q_17184::type hollight.list + => 'q_17192::type hollight.list + => 'q_17182::type => 'q_17182::type. + (ALL (f::'q_17184::type + => 'q_17192::type => 'q_17182::type => 'q_17182::type) + (l2::'q_17192::type hollight.list) b::'q_17182::type. + ITLIST2 f NIL l2 b = b) & + (ALL (h1::'q_17184::type) + (f::'q_17184::type + => 'q_17192::type => 'q_17182::type => 'q_17182::type) + (t1::'q_17184::type hollight.list) + (l2::'q_17192::type hollight.list) b::'q_17182::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_17226::type => 'q_17225::type => 'q_17224::type => 'q_17224::type) + NIL NIL (b::'q_17224::type) = +b & +ITLIST2 f (CONS (h1::'q_17226::type) (t1::'q_17226::type hollight.list)) + (CONS (h2::'q_17225::type) (t2::'q_17225::type hollight.list)) b = +f h1 h2 (ITLIST2 f t1 t2 b)" + by (import hollight ITLIST2) + +consts + ZIP :: "'q_17256::type hollight.list +=> 'q_17264::type hollight.list + => ('q_17256::type * 'q_17264::type) hollight.list" + +defs + ZIP_def: "hollight.ZIP == +SOME ZIP::'q_17256::type hollight.list + => 'q_17264::type hollight.list + => ('q_17256::type * 'q_17264::type) hollight.list. + (ALL l2::'q_17264::type hollight.list. ZIP NIL l2 = NIL) & + (ALL (h1::'q_17256::type) (t1::'q_17256::type hollight.list) + l2::'q_17264::type hollight.list. + ZIP (CONS h1 t1) l2 = CONS (h1, HD l2) (ZIP t1 (TL l2)))" + +lemma DEF_ZIP: "hollight.ZIP = +(SOME ZIP::'q_17256::type hollight.list + => 'q_17264::type hollight.list + => ('q_17256::type * 'q_17264::type) hollight.list. + (ALL l2::'q_17264::type hollight.list. ZIP NIL l2 = NIL) & + (ALL (h1::'q_17256::type) (t1::'q_17256::type hollight.list) + l2::'q_17264::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_17275::type * 'q_17276::type) hollight.list + => ('q_17275::type * 'q_17276::type) hollight.list => bool) + ((hollight.ZIP::'q_17275::type hollight.list + => 'q_17276::type hollight.list + => ('q_17275::type * 'q_17276::type) hollight.list) + (NIL::'q_17275::type hollight.list) + (NIL::'q_17276::type hollight.list)) + (NIL::('q_17275::type * 'q_17276::type) hollight.list)) + ((op =::('q_17300::type * 'q_17301::type) hollight.list + => ('q_17300::type * 'q_17301::type) hollight.list => bool) + ((hollight.ZIP::'q_17300::type hollight.list + => 'q_17301::type hollight.list + => ('q_17300::type * 'q_17301::type) hollight.list) + ((CONS::'q_17300::type + => 'q_17300::type hollight.list + => 'q_17300::type hollight.list) + (h1::'q_17300::type) (t1::'q_17300::type hollight.list)) + ((CONS::'q_17301::type + => 'q_17301::type hollight.list + => 'q_17301::type hollight.list) + (h2::'q_17301::type) (t2::'q_17301::type hollight.list))) + ((CONS::'q_17300::type * 'q_17301::type + => ('q_17300::type * 'q_17301::type) hollight.list + => ('q_17300::type * 'q_17301::type) hollight.list) + ((Pair::'q_17300::type + => 'q_17301::type => 'q_17300::type * 'q_17301::type) + h1 h2) + ((hollight.ZIP::'q_17300::type hollight.list + => 'q_17301::type hollight.list + => ('q_17300::type * 'q_17301::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::nat)) = (l = NIL)" + by (import hollight LENGTH_EQ_NIL) + +lemma LENGTH_EQ_CONS: "ALL (l::'q_17608::type hollight.list) n::nat. + (LENGTH l = Suc n) = + (EX (h::'q_17608::type) t::'q_17608::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_17672::type => 'q_17683::type) + (g::'q_17672::type => 'q_17683::type) l::'q_17672::type hollight.list. + ALL_list (%x::'q_17672::type. f x = g x) l --> MAP f l = MAP g l" + by (import hollight MAP_EQ) + +lemma ALL_IMP: "ALL (P::'q_17713::type => bool) (Q::'q_17713::type => bool) + l::'q_17713::type hollight.list. + (ALL x::'q_17713::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_17741::type => bool) l::'q_17741::type hollight.list. + (~ EX P l) = ALL_list (%x::'q_17741::type. ~ P x) l" + by (import hollight NOT_EX) + +lemma NOT_ALL: "ALL (P::'q_17763::type => bool) l::'q_17763::type hollight.list. + (~ ALL_list P l) = EX (%x::'q_17763::type. ~ P x) l" + by (import hollight NOT_ALL) + +lemma ALL_MAP: "ALL (P::'q_17785::type => bool) (f::'q_17784::type => 'q_17785::type) + l::'q_17784::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_17803::type. True))" + by (import hollight ALL_T) + +lemma MAP_EQ_ALL2: "ALL (l::'q_17828::type hollight.list) m::'q_17828::type hollight.list. + ALL2 + (%(x::'q_17828::type) y::'q_17828::type. + (f::'q_17828::type => 'q_17839::type) x = f y) + l m --> + MAP f l = MAP f m" + by (import hollight MAP_EQ_ALL2) + +lemma ALL2_MAP: "ALL (P::'q_17870::type => 'q_17871::type => bool) + (f::'q_17871::type => 'q_17870::type) l::'q_17871::type hollight.list. + ALL2 P (MAP f l) l = ALL_list (%a::'q_17871::type. P (f a) a) l" + by (import hollight ALL2_MAP) + +lemma MAP_EQ_DEGEN: "ALL (l::'q_17888::type hollight.list) f::'q_17888::type => 'q_17888::type. + ALL_list (%x::'q_17888::type. f x = x) l --> MAP f l = l" + by (import hollight MAP_EQ_DEGEN) + +lemma ALL2_AND_RIGHT: "ALL (l::'q_17931::type hollight.list) (m::'q_17930::type hollight.list) + (P::'q_17931::type => bool) Q::'q_17931::type => 'q_17930::type => bool. + ALL2 (%(x::'q_17931::type) y::'q_17930::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_17968::type hollight.list. + ITLIST (f::'q_17968::type => 'q_17967::type => 'q_17967::type) + (APPEND l (CONS (a::'q_17968::type) NIL)) (b::'q_17967::type) = + ITLIST f l (f a b)" + by (import hollight ITLIST_EXTRA) + +lemma ALL_MP: "ALL (P::'q_17994::type => bool) (Q::'q_17994::type => bool) + l::'q_17994::type hollight.list. + ALL_list (%x::'q_17994::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_18024::type hollight.list. + (ALL_list (P::'q_18024::type => bool) x & + ALL_list (Q::'q_18024::type => bool) x) = + ALL_list (%x::'q_18024::type. P x & Q x) x" + by (import hollight AND_ALL) + +lemma EX_IMP: "ALL (P::'q_18054::type => bool) (Q::'q_18054::type => bool) + l::'q_18054::type hollight.list. + (ALL x::'q_18054::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_18081::type => bool) l::'q_18081::type hollight.list. + (ALL x::'q_18081::type. MEM x l --> P x) = ALL_list P l" + by (import hollight ALL_MEM) + +lemma LENGTH_REPLICATE: "ALL (n::nat) x::'q_18099::type. LENGTH (REPLICATE n x) = n" + by (import hollight LENGTH_REPLICATE) + +lemma EX_MAP: "ALL (P::'q_18123::type => bool) (f::'q_18122::type => 'q_18123::type) + l::'q_18122::type hollight.list. EX P (MAP f l) = EX (P o f) l" + by (import hollight EX_MAP) + +lemma EXISTS_EX: "ALL (P::'q_18161::type => 'q_18160::type => bool) + l::'q_18160::type hollight.list. + (EX x::'q_18161::type. EX (P x) l) = + EX (%s::'q_18160::type. EX x::'q_18161::type. P x s) l" + by (import hollight EXISTS_EX) + +lemma FORALL_ALL: "ALL (P::'q_18191::type => 'q_18190::type => bool) + l::'q_18190::type hollight.list. + (ALL x::'q_18191::type. ALL_list (P x) l) = + ALL_list (%s::'q_18190::type. ALL x::'q_18191::type. P x s) l" + by (import hollight FORALL_ALL) + +lemma MEM_APPEND: "ALL (x::'q_18219::type) (l1::'q_18219::type hollight.list) + l2::'q_18219::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_18255::type => 'q_18252::type) (y::'q_18252::type) + l::'q_18255::type hollight.list. + MEM y (MAP f l) = (EX x::'q_18255::type. MEM x l & y = f x)" + by (import hollight MEM_MAP) + +lemma FILTER_APPEND: "ALL (P::'q_18286::type => bool) (l1::'q_18286::type hollight.list) + l2::'q_18286::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_18313::type => bool) (f::'q_18320::type => 'q_18313::type) + l::'q_18320::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_18347::type => bool) (l::'q_18347::type hollight.list) + x::'q_18347::type. MEM x (FILTER P l) = (P x & MEM x l)" + by (import hollight MEM_FILTER) + +lemma EX_MEM: "ALL (P::'q_18368::type => bool) l::'q_18368::type hollight.list. + (EX x::'q_18368::type. P x & MEM x l) = EX P l" + by (import hollight EX_MEM) + +lemma MAP_FST_ZIP: "ALL (l1::'q_18388::type hollight.list) l2::'q_18390::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_18419::type hollight.list) l2::'q_18421::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_18465::type * 'q_18449::type) hollight.list) x::'q_18465::type. + MEM (x, ASSOC x l) l = MEM x (MAP fst l)" + by (import hollight MEM_ASSOC) + +lemma ALL_APPEND: "ALL (P::'q_18486::type => bool) (l1::'q_18486::type hollight.list) + l2::'q_18486::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_18509::type hollight.list) n::nat. + < n (LENGTH l) --> MEM (EL n l) l" + by (import hollight MEM_EL) + +lemma ALL2_MAP2: "ALL (l::'q_18552::type hollight.list) m::'q_18553::type hollight.list. + ALL2 (P::'q_18551::type => 'q_18550::type => bool) + (MAP (f::'q_18552::type => 'q_18551::type) l) + (MAP (g::'q_18553::type => 'q_18550::type) m) = + ALL2 (%(x::'q_18552::type) y::'q_18553::type. P (f x) (g y)) l m" + by (import hollight ALL2_MAP2) + +lemma AND_ALL2: "ALL (P::'q_18599::type => 'q_18598::type => bool) + (Q::'q_18599::type => 'q_18598::type => bool) + (x::'q_18599::type hollight.list) xa::'q_18598::type hollight.list. + (ALL2 P x xa & ALL2 Q x xa) = + ALL2 (%(x::'q_18599::type) y::'q_18598::type. P x y & Q x y) x xa" + by (import hollight AND_ALL2) + +lemma ALL2_ALL: "ALL (P::'q_18621::type => 'q_18621::type => bool) + l::'q_18621::type hollight.list. + ALL2 P l l = ALL_list (%x::'q_18621::type. P x x) l" + by (import hollight ALL2_ALL) + +lemma APPEND_EQ_NIL: "ALL (x::'q_18650::type hollight.list) xa::'q_18650::type hollight.list. + (APPEND x xa = NIL) = (x = NIL & xa = NIL)" + by (import hollight APPEND_EQ_NIL) + +lemma LENGTH_MAP2: "ALL (f::'q_18670::type => 'q_18672::type => 'q_18683::type) + (l::'q_18670::type hollight.list) m::'q_18672::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) + +constdefs + dist :: "nat * nat => nat" + "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))" + by (import hollight DEF_dist) + +lemma DIST_REFL: "ALL x::nat. dist (x, x) = (0::nat)" + by (import hollight DIST_REFL) + +lemma DIST_LZERO: "ALL x::nat. dist (0::nat, x) = x" + by (import hollight DIST_LZERO) + +lemma DIST_RZERO: "ALL x::nat. dist (x, 0::nat) = x" + by (import hollight DIST_RZERO) + +lemma DIST_SYM: "ALL (x::nat) xa::nat. dist (x, xa) = 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)" + by (import hollight DIST_LADD) + +lemma DIST_RADD: "ALL (x::nat) (xa::nat) xb::nat. dist (x + xa, xb + xa) = dist (x, xb)" + by (import hollight DIST_RADD) + +lemma DIST_LADD_0: "ALL (x::nat) xa::nat. 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" + 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)" + by (import hollight DIST_LMUL) + +lemma DIST_RMUL: "ALL (x::nat) (xa::nat) xb::nat. dist (x, xa) * xb = dist (x * xb, xa * xb)" + by (import hollight DIST_RMUL) + +lemma DIST_EQ_0: "ALL (x::nat) xa::nat. (dist (x, xa) = (0::nat)) = (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))" + 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))" + 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" + 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))" + 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" + by (import hollight BOUNDS_LINEAR) + +lemma BOUNDS_LINEAR_0: "ALL (A::nat) B::nat. (ALL n::nat. <= (A * n) B) = (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))" + by (import hollight BOUNDS_DIVIDED) + +lemma BOUNDS_NOTZERO: "ALL (P::nat => nat => nat) (A::nat) B::nat. + P (0::nat) (0::nat) = (0::nat) & + (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)))" + 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))" + by (import hollight BOUNDS_IGNORE) + +constdefs + is_nadd :: "(nat => nat) => bool" + "is_nadd == +%u::nat => nat. + EX B::nat. + ALL (m::nat) n::nat. <= (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)))" + by (import hollight DEF_is_nadd) + +lemma is_nadd_0: "is_nadd (%n::nat. 0::nat)" + 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::nat)"]) + by (import hollight TYDEF_nadd) + +syntax + dest_nadd :: _ + +syntax + mk_nadd :: _ + +lemmas "TYDEF_nadd_@intern" = typedef_hol2hollight + [where a="a :: nadd" and r=r , + OF type_definition_nadd] + +lemma NADD_CAUCHY: "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))" + 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)" + 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)" + 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" + 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" + 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)" + 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))" + 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)" + by (import hollight NADD_ALTMUL) + +constdefs + nadd_eq :: "nadd => nadd => bool" + "nadd_eq == +%(u::nadd) ua::nadd. + EX B::nat. ALL n::nat. <= (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)" + by (import hollight DEF_nadd_eq) + +lemma NADD_EQ_REFL: "ALL x::nadd. 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" + 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" + by (import hollight NADD_EQ_TRANS) + +constdefs + nadd_of_num :: "nat => nadd" + "nadd_of_num == %u::nat. mk_nadd (op * u)" + +lemma DEF_nadd_of_num: "nadd_of_num = (%u::nat. 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" + 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)" + 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)" + by (import hollight NADD_OF_NUM_EQ) + +constdefs + nadd_le :: "nadd => nadd => bool" + "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))" + 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'" + 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'" + by (import hollight NADD_LE_WELLDEF) + +lemma NADD_LE_REFL: "ALL x::nadd. 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" + 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" + 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::nat) & < (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" + by (import hollight NADD_LE_TOTAL) + +lemma NADD_ARCH: "ALL x::nadd. EX xa::nat. 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" + by (import hollight NADD_OF_NUM_LE) + +constdefs + nadd_add :: "nadd => nadd => nadd" + "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))" + 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)" + 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')" + 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)" + 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)" + by (import hollight NADD_ADD_ASSOC) + +lemma NADD_ADD_LID: "ALL x::nadd. nadd_eq (nadd_add (nadd_of_num (0::nat)) 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" + by (import hollight NADD_ADD_LCANCEL) + +lemma NADD_LE_ADD: "ALL (x::nadd) y::nadd. 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))" + 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))" + by (import hollight NADD_OF_NUM_ADD) + +constdefs + nadd_mul :: "nadd => nadd => nadd" + "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)))" + 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))" + by (import hollight NADD_MUL) + +lemma NADD_MUL_SYM: "ALL (x::nadd) y::nadd. 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)" + by (import hollight NADD_MUL_ASSOC) + +lemma NADD_MUL_LID: "ALL x::nadd. nadd_eq (nadd_mul (nadd_of_num (NUMERAL_BIT1 (0::nat))) 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))" + 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')" + 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')" + 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))" + by (import hollight NADD_OF_NUM_MUL) + +lemma NADD_LE_0: "All (nadd_le (nadd_of_num (0::nat)))" + by (import hollight NADD_LE_0) + +lemma NADD_EQ_IMP_LE: "ALL (x::nadd) y::nadd. 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)" + 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)" + 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" + 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" + 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))" + by (import hollight NADD_RDISTRIB) + +lemma NADD_ARCH_MULT: "ALL (x::nadd) k::nat. + ~ nadd_eq x (nadd_of_num (0::nat)) --> + (EX xa::nat. 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::nat))" + 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" + 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'))" + 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)" + by (import hollight NADD_UBOUND) + +lemma NADD_NONZERO: "ALL x::nadd. + ~ nadd_eq x (nadd_of_num (0::nat)) --> + (EX N::nat. ALL n::nat. <= N n --> dest_nadd x n ~= (0::nat))" + by (import hollight NADD_NONZERO) + +lemma NADD_LBOUND: "ALL x::nadd. + ~ nadd_eq x (nadd_of_num (0::nat)) --> + (EX (A::nat) N::nat. ALL n::nat. <= N n --> <= n (A * dest_nadd x n))" + by (import hollight NADD_LBOUND) + +constdefs + nadd_rinv :: "nadd => nat => nat" + "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))" + by (import hollight DEF_nadd_rinv) + +lemma NADD_MUL_LINV_LEMMA0: "ALL x::nadd. + ~ nadd_eq x (nadd_of_num (0::nat)) --> + (EX (xa::nat) B::nat. ALL i::nat. <= (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::nat) --> + <= (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::nat)) --> + (EX N::nat. + ALL n::nat. + <= N n --> + <= (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::nat)) --> + (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)))" + by (import hollight NADD_MUL_LINV_LEMMA3) + +lemma NADD_MUL_LINV_LEMMA4: "ALL x::nadd. + ~ nadd_eq x (nadd_of_num (0::nat)) --> + (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)))" + by (import hollight NADD_MUL_LINV_LEMMA4) + +lemma NADD_MUL_LINV_LEMMA5: "ALL x::nadd. + ~ nadd_eq x (nadd_of_num (0::nat)) --> + (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))))" + by (import hollight NADD_MUL_LINV_LEMMA5) + +lemma NADD_MUL_LINV_LEMMA6: "ALL x::nadd. + ~ nadd_eq x (nadd_of_num (0::nat)) --> + (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))))" + by (import hollight NADD_MUL_LINV_LEMMA6) + +lemma NADD_MUL_LINV_LEMMA7: "ALL x::nadd. + ~ nadd_eq x (nadd_of_num (0::nat)) --> + (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)))" + by (import hollight NADD_MUL_LINV_LEMMA7) + +lemma NADD_MUL_LINV_LEMMA7a: "ALL x::nadd. + ~ nadd_eq x (nadd_of_num (0::nat)) --> + (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))" + by (import hollight NADD_MUL_LINV_LEMMA7a) + +lemma NADD_MUL_LINV_LEMMA8: "ALL x::nadd. + ~ nadd_eq x (nadd_of_num (0::nat)) --> + (EX B::nat. + ALL (m::nat) n::nat. + <= (dist (m * nadd_rinv x n, n * nadd_rinv x m)) (B * (m + n)))" + by (import hollight NADD_MUL_LINV_LEMMA8) + +constdefs + nadd_inv :: "nadd => nadd" + "nadd_inv == +%u::nadd. + COND (nadd_eq u (nadd_of_num (0::nat))) (nadd_of_num (0::nat)) + (mk_nadd (nadd_rinv u))" + +lemma DEF_nadd_inv: "nadd_inv = +(%u::nadd. + COND (nadd_eq u (nadd_of_num (0::nat))) (nadd_of_num (0::nat)) + (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::nat))) (%n::nat. 0::nat) (nadd_rinv x)" + by (import hollight NADD_INV) + +lemma NADD_MUL_LINV: "ALL x::nadd. + ~ nadd_eq x (nadd_of_num (0::nat)) --> + nadd_eq (nadd_mul (nadd_inv x) x) (nadd_of_num (NUMERAL_BIT1 (0::nat)))" + by (import hollight NADD_MUL_LINV) + +lemma NADD_INV_0: "nadd_eq (nadd_inv (nadd_of_num (0::nat))) (nadd_of_num (0::nat))" + 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)" + 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)"]) + by (import hollight TYDEF_hreal) + +syntax + dest_hreal :: _ + +syntax + mk_hreal :: _ + +lemmas "TYDEF_hreal_@intern" = typedef_hol2hollight + [where a="a :: hreal" and r=r , + OF type_definition_hreal] + +constdefs + hreal_of_num :: "nat => hreal" + "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)))" + by (import hollight DEF_hreal_of_num) + +constdefs + hreal_add :: "hreal => hreal => hreal" + "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)" + +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))" + by (import hollight DEF_hreal_add) + +constdefs + hreal_mul :: "hreal => hreal => hreal" + "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)" + +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))" + by (import hollight DEF_hreal_mul) + +constdefs + hreal_le :: "hreal => hreal => bool" + "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" + +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)" + by (import hollight DEF_hreal_le) + +constdefs + hreal_inv :: "hreal => hreal" + "hreal_inv == +%x::hreal. + mk_hreal + (%u::nadd. EX xa::nadd. 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))" + 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)" + 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)" + 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)" + 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" + 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" + by (import hollight HREAL_LE_ADD_RCANCEL) + +lemma HREAL_ADD_RID: "ALL x::hreal. hreal_add x (hreal_of_num (0::nat)) = 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)" + by (import hollight HREAL_ADD_RDISTRIB) + +lemma HREAL_MUL_LZERO: "ALL m::hreal. hreal_mul (hreal_of_num (0::nat)) m = hreal_of_num (0::nat)" + by (import hollight HREAL_MUL_LZERO) + +lemma HREAL_MUL_RZERO: "ALL x::hreal. hreal_mul x (hreal_of_num (0::nat)) = hreal_of_num (0::nat)" + 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) & +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)" + 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)" + by (import hollight HREAL_LE_MUL_RCANCEL_IMP) + +constdefs + treal_of_num :: "nat => hreal * hreal" + "treal_of_num == %u::nat. (hreal_of_num u, hreal_of_num (0::nat))" + +lemma DEF_treal_of_num: "treal_of_num = (%u::nat. (hreal_of_num u, hreal_of_num (0::nat)))" + by (import hollight DEF_treal_of_num) + +constdefs + treal_neg :: "hreal * hreal => hreal * hreal" + "treal_neg == %u::hreal * hreal. (snd u, fst u)" + +lemma DEF_treal_neg: "treal_neg = (%u::hreal * hreal. (snd u, fst u))" + by (import hollight DEF_treal_neg) + +constdefs + treal_add :: "hreal * hreal => hreal * hreal => hreal * hreal" + "treal_add == +%(u::hreal * hreal) ua::hreal * hreal. + (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)))" + by (import hollight DEF_treal_add) + +constdefs + treal_mul :: "hreal * hreal => hreal * hreal => hreal * hreal" + "treal_mul == +%(u::hreal * hreal) ua::hreal * hreal. + (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. + (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) + +constdefs + treal_le :: "hreal * hreal => hreal * hreal => bool" + "treal_le == +%(u::hreal * hreal) ua::hreal * hreal. + 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)))" + by (import hollight DEF_treal_le) + +constdefs + treal_inv :: "hreal * hreal => hreal * hreal" + "treal_inv == +%u::hreal * hreal. + COND (fst u = snd u) (hreal_of_num (0::nat), hreal_of_num (0::nat)) + (COND (hreal_le (snd u) (fst u)) + (hreal_inv (SOME d::hreal. fst u = hreal_add (snd u) d), + hreal_of_num (0::nat)) + (hreal_of_num (0::nat), + hreal_inv (SOME d::hreal. 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::nat), hreal_of_num (0::nat)) + (COND (hreal_le (snd u) (fst u)) + (hreal_inv (SOME d::hreal. fst u = hreal_add (snd u) d), + hreal_of_num (0::nat)) + (hreal_of_num (0::nat), + hreal_inv (SOME d::hreal. snd u = hreal_add (fst u) d))))" + by (import hollight DEF_treal_inv) + +constdefs + treal_eq :: "hreal * hreal => hreal * hreal => bool" + "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))" + by (import hollight DEF_treal_eq) + +lemma TREAL_EQ_REFL: "ALL x::hreal * hreal. 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" + 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" + by (import hollight TREAL_EQ_TRANS) + +lemma TREAL_EQ_AP: "ALL (x::hreal * hreal) y::hreal * hreal. x = y --> treal_eq x y" + 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)" + 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" + 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))" + 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))" + 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" + 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" + 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)" + 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)" + by (import hollight TREAL_ADD_ASSOC) + +lemma TREAL_ADD_LID: "ALL x::hreal * hreal. treal_eq (treal_add (treal_of_num (0::nat)) 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::nat))" + 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)" + 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)" + by (import hollight TREAL_MUL_ASSOC) + +lemma TREAL_MUL_LID: "ALL x::hreal * hreal. + treal_eq (treal_mul (treal_of_num (NUMERAL_BIT1 (0::nat))) 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))" + by (import hollight TREAL_ADD_LDISTRIB) + +lemma TREAL_LE_REFL: "ALL x::hreal * hreal. 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" + 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" + 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" + 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)" + by (import hollight TREAL_LE_LADD_IMP) + +lemma TREAL_LE_MUL: "ALL (x::hreal * hreal) y::hreal * hreal. + treal_le (treal_of_num (0::nat)) x & + treal_le (treal_of_num (0::nat)) y --> + treal_le (treal_of_num (0::nat)) (treal_mul x y)" + by (import hollight TREAL_LE_MUL) + +lemma TREAL_INV_0: "treal_eq (treal_inv (treal_of_num (0::nat))) (treal_of_num (0::nat))" + by (import hollight TREAL_INV_0) + +lemma TREAL_MUL_LINV: "ALL x::hreal * hreal. + ~ treal_eq x (treal_of_num (0::nat)) --> + treal_eq (treal_mul (treal_inv x) x) + (treal_of_num (NUMERAL_BIT1 (0::nat)))" + 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)" + 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)" + 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)" + 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)" + 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)" + 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)" + 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" + 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" + 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)" + 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)"]) + by (import hollight TYDEF_real) + +syntax + dest_real :: _ + +syntax + mk_real :: _ + +lemmas "TYDEF_real_@intern" = typedef_hol2hollight + [where a="a :: hollight.real" and r=r , + OF type_definition_real] + +constdefs + real_of_num :: "nat => hollight.real" + "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)))" + by (import hollight DEF_real_of_num) + +constdefs + real_neg :: "hollight.real => hollight.real" + "real_neg == +%x1::hollight.real. + mk_real + (%u::hreal * hreal. + EX x1a::hreal * hreal. + 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))" + by (import hollight DEF_real_neg) + +constdefs + real_add :: "hollight.real => hollight.real => hollight.real" + "real_add == +%(x1::hollight.real) y1::hollight.real. + 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)" + +lemma DEF_real_add: "real_add = +(%(x1::hollight.real) y1::hollight.real. + 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))" + by (import hollight DEF_real_add) + +constdefs + real_mul :: "hollight.real => hollight.real => hollight.real" + "real_mul == +%(x1::hollight.real) y1::hollight.real. + 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)" + +lemma DEF_real_mul: "real_mul = +(%(x1::hollight.real) y1::hollight.real. + 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))" + by (import hollight DEF_real_mul) + +constdefs + real_le :: "hollight.real => hollight.real => bool" + "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" + +lemma DEF_real_le: "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)" + by (import hollight DEF_real_le) + +constdefs + real_inv :: "hollight.real => hollight.real" + "real_inv == +%x::hollight.real. + mk_real + (%u::hreal * hreal. + EX xa::hreal * hreal. 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))" + by (import hollight DEF_real_inv) + +constdefs + real_sub :: "hollight.real => hollight.real => hollight.real" + "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))" + by (import hollight DEF_real_sub) + +constdefs + real_lt :: "hollight.real => hollight.real => bool" + "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)" + 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)" + 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)" + by (import hollight DEF_real_gt) + +constdefs + real_abs :: "hollight.real => hollight.real" + "real_abs == +%u::hollight.real. COND (real_le (real_of_num (0::nat)) u) u (real_neg u)" + +lemma DEF_real_abs: "real_abs = +(%u::hollight.real. COND (real_le (real_of_num (0::nat)) u) u (real_neg u))" + by (import hollight DEF_real_abs) + +constdefs + real_pow :: "hollight.real => nat => hollight.real" + "real_pow == +SOME real_pow::hollight.real => nat => hollight.real. + (ALL x::hollight.real. + real_pow x (0::nat) = real_of_num (NUMERAL_BIT1 (0::nat))) & + (ALL (x::hollight.real) n::nat. + 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::nat) = real_of_num (NUMERAL_BIT1 (0::nat))) & + (ALL (x::hollight.real) n::nat. + real_pow x (Suc n) = real_mul x (real_pow x n)))" + by (import hollight DEF_real_pow) + +constdefs + real_div :: "hollight.real => hollight.real => hollight.real" + "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))" + by (import hollight DEF_real_div) + +constdefs + real_max :: "hollight.real => hollight.real => hollight.real" + "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)" + by (import hollight DEF_real_max) + +constdefs + real_min :: "hollight.real => hollight.real => hollight.real" + "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)" + 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::nat)) xa = (EX y::hreal. xa = x y)) & + (ALL (y::hreal) z::hreal. 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::nat)) xa --> r (x xa) = xa) & + (ALL x::hreal. real_le (real_of_num (0::nat)) (r x)) & + (ALL (x::hreal) y::hreal. 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::nat)) 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'))" + 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'))" + 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) & +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::nat)" + 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)" + 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)" + by (import hollight REAL_EQ_ADD_RCANCEL) + +lemma REAL_NEG_NEG: "ALL x::hollight.real. 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)" + 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)" + by (import hollight REAL_MUL_LNEG) + +lemma REAL_ADD_RID: "ALL x::hollight.real. real_add x (real_of_num (0::nat)) = x" + by (import hollight REAL_ADD_RID) + +lemma REAL_LE_LNEG: "ALL (x::hollight.real) y::hollight.real. + real_le (real_neg x) y = real_le (real_of_num (0::nat)) (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" + 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::nat))" + 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)" + 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))" + by (import hollight REAL_POW_NEG) + +lemma REAL_ABS_NUM: "ALL x::nat. 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" + 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" + 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" + by (import hollight REAL_LET_TOTAL) + +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" + 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" + by (import hollight REAL_LT_TRANS) + +lemma REAL_LE_ADD: "ALL (x::hollight.real) y::hollight.real. + real_le (real_of_num (0::nat)) x & real_le (real_of_num (0::nat)) y --> + real_le (real_of_num (0::nat)) (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)" + by (import hollight REAL_LTE_ANTISYM) + +lemma REAL_LT_REFL: "ALL x::hollight.real. ~ real_lt x x" + by (import hollight REAL_LT_REFL) + +lemma REAL_LET_ADD: "ALL (x::hollight.real) y::hollight.real. + real_le (real_of_num (0::nat)) x & real_lt (real_of_num (0::nat)) y --> + real_lt (real_of_num (0::nat)) (real_add x y)" + by (import hollight REAL_LET_ADD) + +lemma REAL_ENTIRE: "ALL (x::hollight.real) y::hollight.real. + (real_mul x y = real_of_num (0::nat)) = + (x = real_of_num (0::nat) | y = real_of_num (0::nat))" + by (import hollight REAL_ENTIRE) + +lemma REAL_POW_2: "ALL x::hollight.real. + real_pow x (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))) = 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::nat)) 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::nat))) x = x) & +(ALL x::hollight.real. + real_mul (real_of_num (0::nat)) x = real_of_num (0::nat)) & +(ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real. + 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::nat) = real_of_num (NUMERAL_BIT1 (0::nat))) & +(ALL (x::hollight.real) xa::nat. + 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::nat)))) 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::nat)))) xa))" + by (import hollight REAL_POLY_NEG_CLAUSES) + +lemma REAL_OF_NUM_LT: "ALL (x::nat) xa::nat. 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" + 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" + 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::nat))) = + 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)" + 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) & +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)" + 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)" + by (import hollight REAL_LT_LADD_IMP) + +lemma REAL_LT_MUL: "ALL (x::hollight.real) y::hollight.real. + real_lt (real_of_num (0::nat)) x & real_lt (real_of_num (0::nat)) y --> + real_lt (real_of_num (0::nat)) (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::nat))" + 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::nat))" + by (import hollight REAL_EQ_ADD_RCANCEL_0) + +lemma REAL_NOT_EQ: "ALL (x::hollight.real) y::hollight.real. + (x ~= y) = (real_lt x y | real_lt y x)" + by (import hollight REAL_NOT_EQ) + +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)" + by (import hollight REAL_LET_ANTISYM) + +lemma REAL_LT_TOTAL: "ALL (x::hollight.real) y::hollight.real. x = y | real_lt x y | real_lt y x" + by (import hollight REAL_LT_TOTAL) + +lemma REAL_LE_01: "real_le (real_of_num (0::nat)) (real_of_num (NUMERAL_BIT1 (0::nat)))" + by (import hollight REAL_LE_01) + +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)" + 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::nat)) (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::nat))" + by (import hollight REAL_LT_RNEG) + +lemma REAL_NEG_EQ_0: "ALL x::hollight.real. + (real_neg x = real_of_num (0::nat)) = (x = real_of_num (0::nat))" + 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" + by (import hollight REAL_ADD_SUB) + +lemma REAL_LE_ADDR: "ALL (x::hollight.real) y::hollight.real. + real_le x (real_add x y) = real_le (real_of_num (0::nat)) 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::nat)) 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::nat)) 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::nat)) x" + by (import hollight REAL_LT_ADDL) + +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)" + by (import hollight REAL_ADD2_SUB2) + +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)" + by (import hollight REAL_LET_ADD2) + +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)" + 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)" + by (import hollight REAL_EQ_SUB_RADD) + +lemma REAL_ADD_SUB2: "ALL (x::hollight.real) y::hollight.real. + 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" + by (import hollight REAL_EQ_IMP_LE) + +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)" + by (import hollight REAL_DIFFSQ) + +lemma REAL_EQ_NEG2: "ALL (x::hollight.real) y::hollight.real. (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" + by (import hollight REAL_LT_NEG2) + +lemma REAL_ABS_ZERO: "ALL x::hollight.real. + (real_abs x = real_of_num (0::nat)) = (x = real_of_num (0::nat))" + by (import hollight REAL_ABS_ZERO) + +lemma REAL_ABS_0: "real_abs (real_of_num (0::nat)) = real_of_num (0::nat)" + by (import hollight REAL_ABS_0) + +lemma REAL_ABS_1: "real_abs (real_of_num (NUMERAL_BIT1 (0::nat))) = +real_of_num (NUMERAL_BIT1 (0::nat))" + 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))" + 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" + 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" + by (import hollight REAL_ABS_TRIANGLE_LT) + +lemma REAL_ABS_POS: "ALL x::hollight.real. real_le (real_of_num (0::nat)) (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)" + by (import hollight REAL_ABS_SUB) + +lemma REAL_ABS_NZ: "ALL x::hollight.real. + (x ~= real_of_num (0::nat)) = real_lt (real_of_num (0::nat)) (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" + by (import hollight REAL_ABS_ABS) + +lemma REAL_ABS_LE: "ALL x::hollight.real. 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::nat)) 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::nat)) 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)" + 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::nat)" + by (import hollight REAL_ABS_STILLNZ) + +lemma REAL_ABS_CASES: "ALL x::hollight.real. + x = real_of_num (0::nat) | real_lt (real_of_num (0::nat)) (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" + 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::nat)) 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::nat))" + 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)" + by (import hollight REAL_ABS_CIRCLE) + +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))" + 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::nat)))) + (real_abs (real_sub x x0))) + (real_sub y0 x0) & + real_lt + (real_mul (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) + (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)" + 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))" + 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))" + 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)" + 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" + 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" + 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" + 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)" + 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)" + 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)" + 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)" + 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)" + 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)" + 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)" + 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)" + 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" + 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" + 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) & +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) & +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)" + by (import hollight REAL_ABS_MUL) + +lemma REAL_POW_LE: "ALL (x::hollight.real) n::nat. + real_le (real_of_num (0::nat)) x --> + real_le (real_of_num (0::nat)) (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::nat)) x --> + real_lt (real_of_num (0::nat)) (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" + 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::nat)) 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::nat)) 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::nat)) 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::nat)) 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::nat) | 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::nat))" + 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::nat)) --> 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::nat)) --> 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" + by (import hollight REAL_INV_INV) + +lemma REAL_INV_EQ_0: "ALL x::hollight.real. + (real_inv x = real_of_num (0::nat)) = (x = real_of_num (0::nat))" + by (import hollight REAL_INV_EQ_0) + +lemma REAL_LT_INV: "ALL x::hollight.real. + real_lt (real_of_num (0::nat)) x --> + real_lt (real_of_num (0::nat)) (real_inv x)" + by (import hollight REAL_LT_INV) + +lemma REAL_LT_INV_EQ: "ALL x::hollight.real. + real_lt (real_of_num (0::nat)) (real_inv x) = + real_lt (real_of_num (0::nat)) 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)" + by (import hollight REAL_INV_NEG) + +lemma REAL_LE_INV_EQ: "ALL x::hollight.real. + real_le (real_of_num (0::nat)) (real_inv x) = + real_le (real_of_num (0::nat)) x" + by (import hollight REAL_LE_INV_EQ) + +lemma REAL_LE_INV: "ALL x::hollight.real. + real_le (real_of_num (0::nat)) x --> + real_le (real_of_num (0::nat)) (real_inv x)" + by (import hollight REAL_LE_INV) + +lemma REAL_INV_1: "real_inv (real_of_num (NUMERAL_BIT1 (0::nat))) = +real_of_num (NUMERAL_BIT1 (0::nat))" + by (import hollight REAL_INV_1) + +lemma REAL_DIV_1: "ALL x::hollight.real. real_div x (real_of_num (NUMERAL_BIT1 (0::nat))) = x" + by (import hollight REAL_DIV_1) + +lemma REAL_DIV_REFL: "ALL x::hollight.real. + x ~= real_of_num (0::nat) --> + real_div x x = real_of_num (NUMERAL_BIT1 (0::nat))" + by (import hollight REAL_DIV_REFL) + +lemma REAL_DIV_RMUL: "ALL (x::hollight.real) xa::hollight.real. + xa ~= real_of_num (0::nat) --> 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::nat) --> 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)" + 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)" + 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)" + 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" + 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)" + 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)" + 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)" + 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)" + by (import hollight REAL_POW_ADD) + +lemma REAL_POW_NZ: "ALL (x::hollight.real) n::nat. + x ~= real_of_num (0::nat) --> real_pow x n ~= real_of_num (0::nat)" + by (import hollight REAL_POW_NZ) + +lemma REAL_POW_SUB: "ALL (x::hollight.real) (m::nat) n::nat. + x ~= real_of_num (0::nat) & <= 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::nat)) x --> x ~= real_of_num (0::nat)" + 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::nat)) 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::nat)) 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::nat)) 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::nat)) 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_LMUL_EQ: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. + real_lt (real_of_num (0::nat)) z --> + real_le (real_mul z x) (real_mul z y) = real_le x y" + by (import hollight REAL_LE_LMUL_EQ) + +lemma REAL_LE_RDIV_EQ: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. + real_lt (real_of_num (0::nat)) 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::nat)) 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::nat)) 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::nat)) 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::nat)) 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::nat)) 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::nat)) 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::nat)) 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::nat)))) 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::nat)) = + (x = real_of_num (0::nat) & n ~= (0::nat))" + 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::nat)) w & + real_le w x & real_le (real_of_num (0::nat)) 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::nat)) w & + real_lt w x & real_le (real_of_num (0::nat)) 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::nat)) (real_mul x x) = + (x ~= real_of_num (0::nat))" + by (import hollight REAL_LT_SQUARE) + +lemma REAL_INV_LE_1: "ALL x::hollight.real. + real_le (real_of_num (NUMERAL_BIT1 (0::nat))) x --> + real_le (real_inv x) (real_of_num (NUMERAL_BIT1 (0::nat)))" + 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::nat))) x --> + real_le (real_of_num (NUMERAL_BIT1 (0::nat))) (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::nat)) x & + real_le x (real_of_num (NUMERAL_BIT1 (0::nat))) --> + real_le (real_pow x n) (real_of_num (NUMERAL_BIT1 (0::nat)))" + by (import hollight REAL_POW_1_LE) + +lemma REAL_POW_1: "ALL x::hollight.real. real_pow x (NUMERAL_BIT1 (0::nat)) = x" + by (import hollight REAL_POW_1) + +lemma REAL_POW_ONE: "ALL n::nat. + real_pow (real_of_num (NUMERAL_BIT1 (0::nat))) n = + real_of_num (NUMERAL_BIT1 (0::nat))" + by (import hollight REAL_POW_ONE) + +lemma REAL_LT_INV2: "ALL (x::hollight.real) y::hollight.real. + real_lt (real_of_num (0::nat)) 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::nat)) 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::nat)) x & + real_le x (real_of_num (NUMERAL_BIT1 (0::nat))) --> + real_le (real_of_num (NUMERAL_BIT1 (0::nat))) (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::nat) & xa ~= real_of_num (0::nat) --> + 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::nat)) d --> + (EX x::hollight.real. real_lt (real_of_num (0::nat)) 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::nat)) d1 & real_lt (real_of_num (0::nat)) d2 --> + (EX e::hollight.real. + real_lt (real_of_num (0::nat)) 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::nat)) 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::nat))) 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::nat) & real_le (real_of_num (0::nat)) 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::nat))) 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)" + 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::nat) & 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::nat) & 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::nat)) x & real_lt (real_of_num (0::nat)) xa --> + real_lt (real_of_num (0::nat)) (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::nat)) x & real_le (real_of_num (0::nat)) xa --> + real_le (real_of_num (0::nat)) (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::nat) --> + 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)))" + by (import hollight REAL_DIV_POW2) + +lemma REAL_DIV_POW2_ALT: "ALL (x::hollight.real) (m::nat) n::nat. + x ~= real_of_num (0::nat) --> + 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)))" + by (import hollight REAL_DIV_POW2_ALT) + +lemma REAL_LT_POW2: "ALL x::nat. + real_lt (real_of_num (0::nat)) + (real_pow (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) x)" + by (import hollight REAL_LT_POW2) + +lemma REAL_LE_POW2: "ALL n::nat. + real_le (real_of_num (NUMERAL_BIT1 (0::nat))) + (real_pow (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) 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::nat))) = + real_pow x (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))" + 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::nat)))) + (real_pow y (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))" + by (import hollight REAL_LE_SQUARE_ABS) + +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))" + 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))" + by (import hollight REAL_WLOG_LT) + +constdefs + mod_real :: "hollight.real => hollight.real => hollight.real => bool" + "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) + +constdefs + DECIMAL :: "nat => nat => hollight.real" + "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))" + by (import hollight DEF_DECIMAL) + +lemma RAT_LEMMA1: "(y1::hollight.real) ~= real_of_num (0::nat) & +(y2::hollight.real) ~= real_of_num (0::nat) --> +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))" + by (import hollight RAT_LEMMA1) + +lemma RAT_LEMMA2: "real_lt (real_of_num (0::nat)) (y1::hollight.real) & +real_lt (real_of_num (0::nat)) (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))" + by (import hollight RAT_LEMMA2) + +lemma RAT_LEMMA3: "real_lt (real_of_num (0::nat)) (y1::hollight.real) & +real_lt (real_of_num (0::nat)) (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))" + by (import hollight RAT_LEMMA3) + +lemma RAT_LEMMA4: "real_lt (real_of_num (0::nat)) (y1::hollight.real) & +real_lt (real_of_num (0::nat)) (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)" + by (import hollight RAT_LEMMA4) + +lemma RAT_LEMMA5: "real_lt (real_of_num (0::nat)) (y1::hollight.real) & +real_lt (real_of_num (0::nat)) (y2::hollight.real) --> +(real_div (x1::hollight.real) y1 = real_div (x2::hollight.real) y2) = +(real_mul x1 y2 = real_mul x2 y1)" + by (import hollight RAT_LEMMA5) + +constdefs + is_int :: "hollight.real => bool" + "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::nat))"]) + by (import hollight TYDEF_int) + +syntax + dest_int :: _ + +syntax + mk_int :: _ + +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)" + by (import hollight dest_int_rep) + +constdefs + int_le :: "hollight.int => hollight.int => bool" + "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))" + by (import hollight DEF_int_le) + +constdefs + int_lt :: "hollight.int => hollight.int => bool" + "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))" + by (import hollight DEF_int_lt) + +constdefs + int_ge :: "hollight.int => hollight.int => bool" + "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))" + by (import hollight DEF_int_ge) + +constdefs + int_gt :: "hollight.int => hollight.int => bool" + "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))" + by (import hollight DEF_int_gt) + +constdefs + int_of_num :: "nat => hollight.int" + "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))" + 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" + by (import hollight int_of_num_th) + +constdefs + int_neg :: "hollight.int => hollight.int" + "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)))" + by (import hollight DEF_int_neg) + +lemma int_neg_th: "ALL x::hollight.int. dest_int (int_neg x) = real_neg (dest_int x)" + by (import hollight int_neg_th) + +constdefs + int_add :: "hollight.int => hollight.int => hollight.int" + "int_add == +%(u::hollight.int) ua::hollight.int. + mk_int (real_add (dest_int u) (dest_int ua))" + +lemma DEF_int_add: "int_add = +(%(u::hollight.int) ua::hollight.int. + mk_int (real_add (dest_int u) (dest_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)" + by (import hollight int_add_th) + +constdefs + int_sub :: "hollight.int => hollight.int => hollight.int" + "int_sub == +%(u::hollight.int) ua::hollight.int. + mk_int (real_sub (dest_int u) (dest_int ua))" + +lemma DEF_int_sub: "int_sub = +(%(u::hollight.int) ua::hollight.int. + mk_int (real_sub (dest_int u) (dest_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)" + by (import hollight int_sub_th) + +constdefs + int_mul :: "hollight.int => hollight.int => hollight.int" + "int_mul == +%(u::hollight.int) ua::hollight.int. + mk_int (real_mul (dest_int u) (dest_int ua))" + +lemma DEF_int_mul: "int_mul = +(%(u::hollight.int) ua::hollight.int. + mk_int (real_mul (dest_int u) (dest_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)" + by (import hollight int_mul_th) + +constdefs + int_abs :: "hollight.int => hollight.int" + "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)))" + by (import hollight DEF_int_abs) + +lemma int_abs_th: "ALL x::hollight.int. dest_int (int_abs x) = real_abs (dest_int x)" + by (import hollight int_abs_th) + +constdefs + int_max :: "hollight.int => hollight.int => hollight.int" + "int_max == +%(u::hollight.int) ua::hollight.int. + mk_int (real_max (dest_int u) (dest_int ua))" + +lemma DEF_int_max: "int_max = +(%(u::hollight.int) ua::hollight.int. + mk_int (real_max (dest_int u) (dest_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)" + by (import hollight int_max_th) + +constdefs + int_min :: "hollight.int => hollight.int => hollight.int" + "int_min == +%(u::hollight.int) ua::hollight.int. + mk_int (real_min (dest_int u) (dest_int ua))" + +lemma DEF_int_min: "int_min = +(%(u::hollight.int) ua::hollight.int. + mk_int (real_min (dest_int u) (dest_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)" + by (import hollight int_min_th) + +constdefs + int_pow :: "hollight.int => nat => hollight.int" + "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))" + 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" + 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))" + 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::nat)))) 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::nat))))" + 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::nat)) 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::nat))) = + (int_abs x = int_of_num (NUMERAL_BIT1 (0::nat)) & + int_abs y = int_of_num (NUMERAL_BIT1 (0::nat)))" + by (import hollight INT_ABS_MUL_1) + +lemma INT_POW: "int_pow (x::hollight.int) (0::nat) = int_of_num (NUMERAL_BIT1 (0::nat)) & +(ALL xa::nat. 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::nat)) x) x (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" + by (import hollight INT_GE) + +lemma INT_GT: "ALL (x::hollight.int) xa::hollight.int. 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)" + by (import hollight INT_LT) + +lemma INT_ARCH: "ALL (x::hollight.int) d::hollight.int. + d ~= int_of_num (0::nat) --> (EX c::hollight.int. int_lt x (int_mul c d))" + by (import hollight INT_ARCH) + +constdefs + mod_int :: "hollight.int => hollight.int => hollight.int => bool" + "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) + +constdefs + IN :: "'A::type => ('A::type => bool) => bool" + "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) + +constdefs + GSPEC :: "('A::type => bool) => 'A::type => bool" + "GSPEC == %u::'A::type => bool. u" + +lemma DEF_GSPEC: "GSPEC = (%u::'A::type => bool. u)" + by (import hollight DEF_GSPEC) + +constdefs + SETSPEC :: "'q_36941::type => bool => 'q_36941::type => bool" + "SETSPEC == %(u::'q_36941::type) (ua::bool) ub::'q_36941::type. ua & u = ub" + +lemma DEF_SETSPEC: "SETSPEC = (%(u::'q_36941::type) (ua::bool) ub::'q_36941::type. ua & u = ub)" + by (import hollight DEF_SETSPEC) + +lemma IN_ELIM_THM: "(ALL (P::(bool => 'q_36974::type => bool) => bool) x::'q_36974::type. + IN x (GSPEC (%v::'q_36974::type. P (SETSPEC v))) = + P (%(p::bool) t::'q_36974::type. p & x = t)) & +(ALL (p::'q_37005::type => bool) x::'q_37005::type. + IN x + (GSPEC (%v::'q_37005::type. EX y::'q_37005::type. SETSPEC v (p y) y)) = + p x) & +(ALL (P::(bool => 'q_37033::type => bool) => bool) x::'q_37033::type. + GSPEC (%v::'q_37033::type. P (SETSPEC v)) x = + P (%(p::bool) t::'q_37033::type. p & x = t)) & +(ALL (p::'q_37062::type => bool) x::'q_37062::type. + GSPEC (%v::'q_37062::type. EX y::'q_37062::type. SETSPEC v (p y) y) x = + p x) & +(ALL (p::'q_37079::type => bool) x::'q_37079::type. IN x p = p x)" + by (import hollight IN_ELIM_THM) + +constdefs + EMPTY :: "'A::type => bool" + "EMPTY == %x::'A::type. False" + +lemma DEF_EMPTY: "EMPTY = (%x::'A::type. False)" + by (import hollight DEF_EMPTY) + +constdefs + INSERT :: "'A::type => ('A::type => bool) => 'A::type => bool" + "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) + +consts + UNIV :: "'A::type => bool" + +defs + UNIV_def: "hollight.UNIV == %x::'A::type. True" + +lemma DEF_UNIV: "hollight.UNIV = (%x::'A::type. True)" + by (import hollight DEF_UNIV) + +consts + UNION :: "('A::type => bool) => ('A::type => bool) => 'A::type => 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) + +constdefs + UNIONS :: "(('A::type => bool) => bool) => 'A::type => bool" + "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) + +consts + INTER :: "('A::type => bool) => ('A::type => bool) => 'A::type => bool" + +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) + +constdefs + INTERS :: "(('A::type => bool) => bool) => 'A::type => bool" + "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) + +constdefs + DIFF :: "('A::type => bool) => ('A::type => bool) => 'A::type => bool" + "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)" + by (import hollight INSERT) + +constdefs + DELETE :: "('A::type => bool) => 'A::type => 'A::type => bool" + "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) + +constdefs + SUBSET :: "('A::type => bool) => ('A::type => bool) => bool" + "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) + +constdefs + PSUBSET :: "('A::type => bool) => ('A::type => bool) => bool" + "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) + +constdefs + DISJOINT :: "('A::type => bool) => ('A::type => bool) => bool" + "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) + +constdefs + SING :: "('A::type => bool) => bool" + "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)" + by (import hollight DEF_SING) + +constdefs + FINITE :: "('A::type => bool) => bool" + "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) + +constdefs + INFINITE :: "('A::type => bool) => bool" + "INFINITE == %u::'A::type => bool. ~ FINITE u" + +lemma DEF_INFINITE: "INFINITE = (%u::'A::type => bool. ~ FINITE u)" + by (import hollight DEF_INFINITE) + +constdefs + IMAGE :: "('A::type => 'B::type) => ('A::type => bool) => 'B::type => bool" + "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) + +constdefs + INJ :: "('A::type => 'B::type) => ('A::type => bool) => ('B::type => bool) => bool" + "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)" + +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))" + by (import hollight DEF_INJ) + +constdefs + SURJ :: "('A::type => 'B::type) => ('A::type => bool) => ('B::type => bool) => bool" + "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))" + +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)))" + by (import hollight DEF_SURJ) + +constdefs + BIJ :: "('A::type => 'B::type) => ('A::type => bool) => ('B::type => bool) => bool" + "BIJ == +%(u::'A::type => 'B::type) (ua::'A::type => bool) ub::'B::type => 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)" + by (import hollight DEF_BIJ) + +constdefs + CHOICE :: "('A::type => bool) => 'A::type" + "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) + +constdefs + REST :: "('A::type => bool) => 'A::type => bool" + "REST == %u::'A::type => bool. DELETE u (CHOICE u)" + +lemma DEF_REST: "REST = (%u::'A::type => bool. DELETE u (CHOICE u))" + by (import hollight DEF_REST) + +constdefs + CARD_GE :: "('q_37578::type => bool) => ('q_37575::type => bool) => bool" + "CARD_GE == +%(u::'q_37578::type => bool) ua::'q_37575::type => bool. + EX f::'q_37578::type => 'q_37575::type. + ALL y::'q_37575::type. + IN y ua --> (EX x::'q_37578::type. IN x u & y = f x)" + +lemma DEF_CARD_GE: "CARD_GE = +(%(u::'q_37578::type => bool) ua::'q_37575::type => bool. + EX f::'q_37578::type => 'q_37575::type. + ALL y::'q_37575::type. + IN y ua --> (EX x::'q_37578::type. IN x u & y = f x))" + by (import hollight DEF_CARD_GE) + +constdefs + CARD_LE :: "('q_37587::type => bool) => ('q_37586::type => bool) => bool" + "CARD_LE == +%(u::'q_37587::type => bool) ua::'q_37586::type => bool. CARD_GE ua u" + +lemma DEF_CARD_LE: "CARD_LE = +(%(u::'q_37587::type => bool) ua::'q_37586::type => bool. CARD_GE ua u)" + by (import hollight DEF_CARD_LE) + +constdefs + CARD_EQ :: "('q_37597::type => bool) => ('q_37598::type => bool) => bool" + "CARD_EQ == +%(u::'q_37597::type => bool) ua::'q_37598::type => bool. + CARD_LE u ua & CARD_LE ua u" + +lemma DEF_CARD_EQ: "CARD_EQ = +(%(u::'q_37597::type => bool) ua::'q_37598::type => bool. + CARD_LE u ua & CARD_LE ua u)" + by (import hollight DEF_CARD_EQ) + +constdefs + CARD_GT :: "('q_37612::type => bool) => ('q_37613::type => bool) => bool" + "CARD_GT == +%(u::'q_37612::type => bool) ua::'q_37613::type => bool. + CARD_GE u ua & ~ CARD_GE ua u" + +lemma DEF_CARD_GT: "CARD_GT = +(%(u::'q_37612::type => bool) ua::'q_37613::type => bool. + CARD_GE u ua & ~ CARD_GE ua u)" + by (import hollight DEF_CARD_GT) + +constdefs + CARD_LT :: "('q_37628::type => bool) => ('q_37629::type => bool) => bool" + "CARD_LT == +%(u::'q_37628::type => bool) ua::'q_37629::type => bool. + CARD_LE u ua & ~ CARD_LE ua u" + +lemma DEF_CARD_LT: "CARD_LT = +(%(u::'q_37628::type => bool) ua::'q_37629::type => bool. + CARD_LE u ua & ~ CARD_LE ua u)" + by (import hollight DEF_CARD_LT) + +constdefs + COUNTABLE :: "('q_37642::type => bool) => bool" + "(op ==::(('q_37642::type => bool) => bool) + => (('q_37642::type => bool) => bool) => prop) + (COUNTABLE::('q_37642::type => bool) => bool) + ((CARD_GE::(nat => bool) => ('q_37642::type => bool) => bool) + (hollight.UNIV::nat => bool))" + +lemma DEF_COUNTABLE: "(op =::(('q_37642::type => bool) => bool) + => (('q_37642::type => bool) => bool) => bool) + (COUNTABLE::('q_37642::type => bool) => bool) + ((CARD_GE::(nat => bool) => ('q_37642::type => bool) => bool) + (hollight.UNIV::nat => bool))" + by (import hollight DEF_COUNTABLE) + +lemma NOT_IN_EMPTY: "ALL x::'A::type. ~ IN x EMPTY" + 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)" + 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)" + 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)" + 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)" + 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)" + by (import hollight IN_REST) + +lemma CHOICE_DEF: "ALL x::'A::type => bool. x ~= EMPTY --> IN (CHOICE 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))" + 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))" + by (import hollight EMPTY_NOT_UNIV) + +lemma EQ_UNIV: "(ALL x::'A::type. IN x (s::'A::type => bool)) = (s = hollight.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)" + 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)" + 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))" + 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)" + 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)" + 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_38479::type => bool) (xa::'q_38479::type => bool) + xb::'q_38479::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)" + 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)" + 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)" + 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))" + by (import hollight IN_DISJOINT) + +lemma DISJOINT_SYM: "ALL (x::'A::type => bool) xa::'A::type => bool. + DISJOINT x xa = DISJOINT xa x" + by (import hollight DISJOINT_SYM) + +lemma DISJOINT_EMPTY: "ALL x::'A::type => bool. DISJOINT EMPTY x & DISJOINT x EMPTY" + by (import hollight DISJOINT_EMPTY) + +lemma DISJOINT_EMPTY_REFL: "ALL x::'A::type => bool. (x = EMPTY) = DISJOINT x 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)" + 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_38897::type => bool) xa::'q_38897::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)" + 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)" + by (import hollight SET_CASES) + +lemma ABSORPTION: "ALL (x::'A::type) xa::'A::type => bool. IN x xa = (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" + 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))" + 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)" + 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_39353::type) + (INSERT (y::'q_39353::type) (s::'q_39353::type => 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_39420::type => bool) (q::'q_39420::type => bool) = +hollight.INTER q p & +hollight.INTER (hollight.INTER p q) (r::'q_39420::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" + by (import hollight INTER_ACI) + +lemma UNION_ACI: "hollight.UNION (p::'q_39486::type => bool) (q::'q_39486::type => bool) = +hollight.UNION q p & +hollight.UNION (hollight.UNION p q) (r::'q_39486::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" + by (import hollight UNION_ACI) + +lemma DELETE_NON_ELEMENT: "ALL (x::'A::type) xa::'A::type => bool. (~ IN x xa) = (DELETE 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))" + by (import hollight IN_DELETE_EQ) + +lemma EMPTY_DELETE: "ALL x::'A::type. DELETE EMPTY x = EMPTY" + by (import hollight EMPTY_DELETE) + +lemma DELETE_DELETE: "ALL (x::'A::type) xa::'A::type => bool. DELETE (DELETE xa x) x = DELETE 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" + by (import hollight DELETE_COMM) + +lemma DELETE_SUBSET: "ALL (x::'A::type) xa::'A::type => bool. SUBSET (DELETE xa x) 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)" + 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" + 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)" + 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))" + 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))" + 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" + 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" + by (import hollight DISJOINT_DELETE_SYM) + +lemma UNIONS_0: "(op =::('q_39893::type => bool) => ('q_39893::type => bool) => bool) + ((UNIONS::(('q_39893::type => bool) => bool) => 'q_39893::type => bool) + (EMPTY::('q_39893::type => bool) => bool)) + (EMPTY::'q_39893::type => bool)" + by (import hollight UNIONS_0) + +lemma UNIONS_1: "UNIONS (INSERT (s::'q_39899::type => bool) EMPTY) = s" + by (import hollight UNIONS_1) + +lemma UNIONS_2: "UNIONS + (INSERT (s::'q_39919::type => bool) + (INSERT (t::'q_39919::type => bool) EMPTY)) = +hollight.UNION s t" + by (import hollight UNIONS_2) + +lemma UNIONS_INSERT: "UNIONS + (INSERT (s::'q_39933::type => bool) + (u::('q_39933::type => bool) => bool)) = +hollight.UNION s (UNIONS u)" + by (import hollight UNIONS_INSERT) + +lemma FORALL_IN_UNIONS: "ALL (x::'q_39975::type => bool) xa::('q_39975::type => bool) => bool. + (ALL xb::'q_39975::type. IN xb (UNIONS xa) --> x xb) = + (ALL (t::'q_39975::type => bool) xb::'q_39975::type. + IN t xa & IN xb t --> x xb)" + by (import hollight FORALL_IN_UNIONS) + +lemma EMPTY_UNIONS: "ALL x::('q_40001::type => bool) => bool. + (UNIONS x = EMPTY) = + (ALL xa::'q_40001::type => bool. IN xa x --> xa = EMPTY)" + by (import hollight EMPTY_UNIONS) + +lemma IMAGE_CLAUSES: "IMAGE (f::'q_40027::type => 'q_40031::type) EMPTY = EMPTY & +IMAGE f (INSERT (x::'q_40027::type) (s::'q_40027::type => bool)) = +INSERT (f x) (IMAGE f s)" + by (import hollight IMAGE_CLAUSES) + +lemma IMAGE_UNION: "ALL (x::'q_40054::type => 'q_40065::type) (xa::'q_40054::type => bool) + xb::'q_40054::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_40098::type => 'q_40094::type) + (xa::'q_40089::type => 'q_40098::type) xb::'q_40089::type => bool. + IMAGE (x o xa) xb = IMAGE x (IMAGE xa xb)" + by (import hollight IMAGE_o) + +lemma IMAGE_SUBSET: "ALL (x::'q_40116::type => 'q_40127::type) (xa::'q_40116::type => bool) + xb::'q_40116::type => bool. + SUBSET xa xb --> SUBSET (IMAGE x xa) (IMAGE x xb)" + by (import hollight IMAGE_SUBSET) + +lemma IMAGE_DIFF_INJ: "(ALL (x::'q_40158::type) y::'q_40158::type. + (f::'q_40158::type => 'q_40169::type) x = f y --> x = y) --> +IMAGE f (DIFF (s::'q_40158::type => bool) (t::'q_40158::type => bool)) = +DIFF (IMAGE f s) (IMAGE f t)" + by (import hollight IMAGE_DIFF_INJ) + +lemma IMAGE_DELETE_INJ: "(ALL x::'q_40204::type. + (f::'q_40204::type => 'q_40203::type) x = f (a::'q_40204::type) --> + x = a) --> +IMAGE f (DELETE (s::'q_40204::type => bool) a) = DELETE (IMAGE f s) (f a)" + by (import hollight IMAGE_DELETE_INJ) + +lemma IMAGE_EQ_EMPTY: "ALL (x::'q_40227::type => 'q_40223::type) xa::'q_40227::type => bool. + (IMAGE x xa = EMPTY) = (xa = EMPTY)" + by (import hollight IMAGE_EQ_EMPTY) + +lemma FORALL_IN_IMAGE: "ALL (x::'q_40263::type => 'q_40262::type) xa::'q_40263::type => bool. + (ALL xb::'q_40262::type. + IN xb (IMAGE x xa) --> (P::'q_40262::type => bool) xb) = + (ALL xb::'q_40263::type. IN xb xa --> P (x xb))" + by (import hollight FORALL_IN_IMAGE) + +lemma EXISTS_IN_IMAGE: "ALL (x::'q_40299::type => 'q_40298::type) xa::'q_40299::type => bool. + (EX xb::'q_40298::type. + IN xb (IMAGE x xa) & (P::'q_40298::type => bool) xb) = + (EX xb::'q_40299::type. IN 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_40385::type => bool) c::'q_40390::type. + IMAGE (%x::'q_40385::type. c) s = COND (s = EMPTY) EMPTY (INSERT c EMPTY)" + by (import hollight IMAGE_CONST) + +lemma SIMPLE_IMAGE: "ALL (x::'q_40418::type => 'q_40422::type) xa::'q_40418::type => bool. + GSPEC + (%u::'q_40422::type. + EX xb::'q_40418::type. SETSPEC u (IN xb xa) (x xb)) = + IMAGE x xa" + by (import hollight SIMPLE_IMAGE) + +lemma EMPTY_GSPEC: "GSPEC (%u::'q_40439::type. Ex (SETSPEC u False)) = EMPTY" + by (import hollight EMPTY_GSPEC) + +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)" + by (import hollight FINITE_DELETE_IMP) + +lemma FINITE_DELETE: "ALL (s::'A::type => bool) x::'A::type. FINITE (DELETE s x) = FINITE s" + by (import hollight FINITE_DELETE) + +lemma FINITE_UNIONS: "ALL s::('q_40774::type => bool) => bool. + FINITE s --> + FINITE (UNIONS s) = (ALL t::'q_40774::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))" + 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))" + 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))" + 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))" + by (import hollight INFINITE_IMAGE_INJ) + +lemma INFINITE_NONEMPTY: "ALL s::'q_41257::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)" + 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'))" + 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_41555::type => bool) t::'q_41555::type => bool. + FINITE s --> FINITE (DIFF s t)" + by (import hollight FINITE_DIFF) + +constdefs + FINREC :: "('q_41615::type => 'q_41614::type => 'q_41614::type) +=> 'q_41614::type + => ('q_41615::type => bool) => 'q_41614::type => nat => bool" + "FINREC == +SOME FINREC::('q_41615::type => 'q_41614::type => 'q_41614::type) + => 'q_41614::type + => ('q_41615::type => bool) + => 'q_41614::type => nat => bool. + (ALL (f::'q_41615::type => 'q_41614::type => 'q_41614::type) + (s::'q_41615::type => bool) (a::'q_41614::type) b::'q_41614::type. + FINREC f b s a (0::nat) = (s = EMPTY & a = b)) & + (ALL (b::'q_41614::type) (s::'q_41615::type => bool) (n::nat) + (a::'q_41614::type) + f::'q_41615::type => 'q_41614::type => 'q_41614::type. + FINREC f b s a (Suc n) = + (EX (x::'q_41615::type) c::'q_41614::type. + IN x s & FINREC f b (DELETE s x) c n & a = f x c))" + +lemma DEF_FINREC: "FINREC = +(SOME FINREC::('q_41615::type => 'q_41614::type => 'q_41614::type) + => 'q_41614::type + => ('q_41615::type => bool) + => 'q_41614::type => nat => bool. + (ALL (f::'q_41615::type => 'q_41614::type => 'q_41614::type) + (s::'q_41615::type => bool) (a::'q_41614::type) b::'q_41614::type. + FINREC f b s a (0::nat) = (s = EMPTY & a = b)) & + (ALL (b::'q_41614::type) (s::'q_41615::type => bool) (n::nat) + (a::'q_41614::type) + f::'q_41615::type => 'q_41614::type => 'q_41614::type. + FINREC f b s a (Suc n) = + (EX (x::'q_41615::type) c::'q_41614::type. + IN x s & FINREC f b (DELETE s x) c n & a = f x c)))" + by (import hollight DEF_FINREC) + +lemma FINREC_1_LEMMA: "ALL (x::'q_41660::type => 'q_41659::type => 'q_41659::type) + (xa::'q_41659::type) (xb::'q_41660::type => bool) xc::'q_41659::type. + FINREC x xa xb xc (Suc (0::nat)) = + (EX xd::'q_41660::type. xb = INSERT xd EMPTY & 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)))" + 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)" + 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))" + 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))" + 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))))" + 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))))" + by (import hollight SET_RECURSION_LEMMA) + +constdefs + ITSET :: "('q_42316::type => 'q_42315::type => 'q_42315::type) +=> ('q_42316::type => bool) => 'q_42315::type => 'q_42315::type" + "ITSET == +%(u::'q_42316::type => 'q_42315::type => 'q_42315::type) + (ua::'q_42316::type => bool) ub::'q_42315::type. + (SOME g::('q_42316::type => bool) => 'q_42315::type. + g EMPTY = ub & + (ALL (x::'q_42316::type) s::'q_42316::type => bool. + FINITE s --> g (INSERT x s) = COND (IN x s) (g s) (u x (g s)))) + ua" + +lemma DEF_ITSET: "ITSET = +(%(u::'q_42316::type => 'q_42315::type => 'q_42315::type) + (ua::'q_42316::type => bool) ub::'q_42315::type. + (SOME g::('q_42316::type => bool) => 'q_42315::type. + g EMPTY = ub & + (ALL (x::'q_42316::type) s::'q_42316::type => bool. + FINITE s --> g (INSERT x s) = COND (IN x s) (g s) (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)))" + 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))" + by (import hollight FINITE_RECURSION_DELETE) + +lemma ITSET_EQ: "ALL (x::'q_42621::type => bool) + (xa::'q_42621::type => 'q_42622::type => 'q_42622::type) + (xb::'q_42621::type => 'q_42622::type => 'q_42622::type) + xc::'q_42622::type. + FINITE x & + (ALL xc::'q_42621::type. IN xc x --> xa xc = xb xc) & + (ALL (x::'q_42621::type) (y::'q_42621::type) s::'q_42622::type. + x ~= y --> xa x (xa y s) = xa y (xa x s)) & + (ALL (x::'q_42621::type) (y::'q_42621::type) s::'q_42622::type. + x ~= y --> xb x (xb y s) = xb y (xb x s)) --> + ITSET xa x xc = ITSET xb x xc" + by (import hollight ITSET_EQ) + +lemma SUBSET_RESTRICT: "ALL (x::'q_42655::type => bool) xa::'q_42655::type => bool. + SUBSET + (GSPEC + (%u::'q_42655::type. + EX xb::'q_42655::type. SETSPEC u (IN xb x & xa xb) xb)) + x" + by (import hollight SUBSET_RESTRICT) + +lemma FINITE_RESTRICT: "ALL (s::'A::type => bool) p::'q_42673::type. + FINITE s --> + FINITE + (GSPEC + (%u::'A::type. + EX x::'A::type. SETSPEC u (IN x s & (P::'A::type => bool) x) x))" + by (import hollight FINITE_RESTRICT) + +constdefs + CARD :: "('q_42709::type => bool) => nat" + "CARD == +%u::'q_42709::type => bool. ITSET (%x::'q_42709::type. Suc) u (0::nat)" + +lemma DEF_CARD: "CARD = +(%u::'q_42709::type => bool. ITSET (%x::'q_42709::type. 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)))))))" + 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" + 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::nat)) (CARD s)" + by (import hollight CARD_DELETE) + +lemma CARD_UNION_EQ: "ALL (s::'q_42954::type => bool) (t::'q_42954::type => bool) + u::'q_42954::type => bool. + FINITE u & hollight.INTER s t = EMPTY & hollight.UNION s t = u --> + CARD s + CARD t = CARD u" + by (import hollight CARD_UNION_EQ) + +constdefs + HAS_SIZE :: "('q_42990::type => bool) => nat => bool" + "HAS_SIZE == %(u::'q_42990::type => bool) ua::nat. FINITE u & CARD u = ua" + +lemma DEF_HAS_SIZE: "HAS_SIZE = (%(u::'q_42990::type => bool) ua::nat. FINITE u & CARD u = ua)" + by (import hollight DEF_HAS_SIZE) + +lemma HAS_SIZE_CARD: "ALL (x::'q_43009::type => bool) xa::nat. HAS_SIZE x xa --> CARD x = xa" + by (import hollight HAS_SIZE_CARD) + +lemma HAS_SIZE_0: "ALL (s::'A::type => bool) n::'q_43025::type. + HAS_SIZE s (0::nat) = (s = EMPTY)" + 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))" + by (import hollight HAS_SIZE_SUC) + +lemma HAS_SIZE_UNION: "ALL (x::'q_43147::type => bool) (xa::'q_43147::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)" + 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)" + by (import hollight HAS_SIZE_UNIONS) + +lemma HAS_SIZE_CLAUSES: "HAS_SIZE (s::'q_43395::type => bool) (0::nat) = (s = EMPTY) & +HAS_SIZE s (Suc (n::nat)) = +(EX (a::'q_43395::type) t::'q_43395::type => bool. + HAS_SIZE t n & ~ IN 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" + 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)" + 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" + by (import hollight CARD_SUBSET_LE) + +lemma CARD_EQ_0: "ALL s::'q_43711::type => bool. + FINITE s --> (CARD s = (0::nat)) = (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)" + 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)" + 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)" + 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" + 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" + 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)" + 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))" + 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)" + 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)))" + 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)))" + 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" + 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)" + 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)" + by (import hollight HAS_SIZE_FUNSPACE) + +lemma CARD_FUNSPACE: "ALL (s::'q_45066::type => bool) t::'q_45063::type => bool. + FINITE s & FINITE t --> + CARD + (GSPEC + (%u::'q_45066::type => 'q_45063::type. + EX f::'q_45066::type => 'q_45063::type. + SETSPEC u + ((ALL x::'q_45066::type. IN x s --> IN (f x) t) & + (ALL x::'q_45066::type. + ~ IN x s --> f x = (d::'q_45063::type))) + f)) = + EXP (CARD t) (CARD s)" + by (import hollight CARD_FUNSPACE) + +lemma FINITE_FUNSPACE: "ALL (s::'q_45132::type => bool) t::'q_45129::type => bool. + FINITE s & FINITE t --> + FINITE + (GSPEC + (%u::'q_45132::type => 'q_45129::type. + EX f::'q_45132::type => 'q_45129::type. + SETSPEC u + ((ALL x::'q_45132::type. IN x s --> IN (f x) t) & + (ALL x::'q_45132::type. + ~ IN x s --> f x = (d::'q_45129::type))) + 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::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::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))" + 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" + by (import hollight HAS_SIZE_NUMSEG_LT) + +lemma CARD_NUMSEG_LT: "ALL x::nat. CARD (GSPEC (%u::nat. EX m::nat. SETSPEC u (< m x) m)) = x" + by (import hollight CARD_NUMSEG_LT) + +lemma FINITE_NUMSEG_LT: "ALL x::nat. FINITE (GSPEC (%u::nat. EX m::nat. SETSPEC u (< m x) m))" + by (import hollight FINITE_NUMSEG_LT) + +lemma HAS_SIZE_NUMSEG_LE: "ALL x::nat. + HAS_SIZE (GSPEC (%xa::nat. EX xb::nat. SETSPEC xa (<= xb x) xb)) + (x + NUMERAL_BIT1 (0::nat))" + by (import hollight HAS_SIZE_NUMSEG_LE) + +lemma FINITE_NUMSEG_LE: "ALL x::nat. FINITE (GSPEC (%u::nat. EX m::nat. SETSPEC u (<= m x) m))" + by (import hollight FINITE_NUMSEG_LE) + +lemma CARD_NUMSEG_LE: "ALL x::nat. + CARD (GSPEC (%u::nat. EX m::nat. SETSPEC u (<= m x) m)) = + x + NUMERAL_BIT1 (0::nat)" + by (import hollight CARD_NUMSEG_LE) + +lemma num_FINITE: "ALL s::nat => bool. FINITE s = (EX a::nat. ALL x::nat. IN 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)" + 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)))" + by (import hollight HAS_SIZE_INDEX) + +constdefs + set_of_list :: "'q_45759::type hollight.list => 'q_45759::type => bool" + "set_of_list == +SOME set_of_list::'q_45759::type hollight.list => 'q_45759::type => bool. + set_of_list NIL = EMPTY & + (ALL (h::'q_45759::type) t::'q_45759::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_45759::type hollight.list => 'q_45759::type => bool. + set_of_list NIL = EMPTY & + (ALL (h::'q_45759::type) t::'q_45759::type hollight.list. + set_of_list (CONS h t) = INSERT h (set_of_list t)))" + by (import hollight DEF_set_of_list) + +constdefs + list_of_set :: "('q_45777::type => bool) => 'q_45777::type hollight.list" + "list_of_set == +%u::'q_45777::type => bool. + SOME l::'q_45777::type hollight.list. + set_of_list l = u & LENGTH l = CARD u" + +lemma DEF_list_of_set: "list_of_set = +(%u::'q_45777::type => bool. + SOME l::'q_45777::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_45826::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_45842::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_45887::type hollight.list. FINITE (set_of_list l)" + by (import hollight FINITE_SET_OF_LIST) + +lemma IN_SET_OF_LIST: "ALL (x::'q_45905::type) l::'q_45905::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_45930::type hollight.list) xa::'q_45930::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) + +constdefs + pairwise :: "('q_45989::type => 'q_45989::type => bool) +=> ('q_45989::type => bool) => bool" + "pairwise == +%(u::'q_45989::type => 'q_45989::type => bool) ua::'q_45989::type => bool. + ALL (x::'q_45989::type) y::'q_45989::type. + IN x ua & IN y ua & x ~= y --> u x y" + +lemma DEF_pairwise: "pairwise = +(%(u::'q_45989::type => 'q_45989::type => bool) ua::'q_45989::type => bool. + ALL (x::'q_45989::type) y::'q_45989::type. + IN x ua & IN y ua & x ~= y --> u x y)" + by (import hollight DEF_pairwise) + +constdefs + PAIRWISE :: "('q_46011::type => 'q_46011::type => bool) +=> 'q_46011::type hollight.list => bool" + "PAIRWISE == +SOME PAIRWISE::('q_46011::type => 'q_46011::type => bool) + => 'q_46011::type hollight.list => bool. + (ALL r::'q_46011::type => 'q_46011::type => bool. + PAIRWISE r NIL = True) & + (ALL (h::'q_46011::type) (r::'q_46011::type => 'q_46011::type => bool) + t::'q_46011::type hollight.list. + PAIRWISE r (CONS h t) = (ALL_list (r h) t & PAIRWISE r t))" + +lemma DEF_PAIRWISE: "PAIRWISE = +(SOME PAIRWISE::('q_46011::type => 'q_46011::type => bool) + => 'q_46011::type hollight.list => bool. + (ALL r::'q_46011::type => 'q_46011::type => bool. + PAIRWISE r NIL = True) & + (ALL (h::'q_46011::type) (r::'q_46011::type => 'q_46011::type => bool) + t::'q_46011::type hollight.list. + PAIRWISE r (CONS h t) = (ALL_list (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)))"]) + by (import hollight TYDEF_finite_image) + +syntax + dest_finite_image :: _ + +syntax + mk_finite_image :: _ + +lemmas "TYDEF_finite_image_@intern" = typedef_hol2hollight + [where a="a :: 'A::type finite_image" and r=r , + OF type_definition_finite_image] + +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))))" + by (import hollight FINITE_IMAGE_IMAGE) + +constdefs + dimindex :: "('A::type => bool) => nat" + "(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))" + 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))" + by (import hollight CARD_FINITE_IMAGE) + +lemma FINITE_FINITE_IMAGE: "(FINITE::('A::type finite_image => bool) => bool) + (hollight.UNIV::'A::type finite_image => bool)" + by (import hollight FINITE_FINITE_IMAGE) + +lemma DIMINDEX_NONZERO: "ALL s::'A::type => bool. dimindex s ~= (0::nat)" + by (import hollight DIMINDEX_NONZERO) + +lemma DIMINDEX_GE_1: "ALL x::'A::type => bool. <= (NUMERAL_BIT1 (0::nat)) (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" + by (import hollight DIMINDEX_FINITE_IMAGE) + +constdefs + finite_index :: "nat => 'A::type" + "(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))))" + 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))))" + by (import hollight FINITE_INDEX_INJ) + +lemma FORALL_FINITE_INDEX: "(op =::bool => bool => bool) + ((All::('N::type finite_image => bool) => bool) + (P::'N::type finite_image => 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::('N::type => bool) => nat) + (hollight.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)"]) + by (import hollight TYDEF_cart) + +syntax + dest_cart :: _ + +syntax + mk_cart :: _ + +lemmas "TYDEF_cart_@intern" = typedef_hol2hollight + [where a="a :: ('A::type, 'B::type) cart" and r=r , + OF type_definition_cart] + +consts + "$" :: "('q_46418::type, 'q_46425::type) cart => nat => 'q_46418::type" ("$") + +defs + "$_def": "$ == +%(u::('q_46418::type, 'q_46425::type) cart) ua::nat. + dest_cart u (finite_index ua)" + +lemma "DEF_$": "$ = +(%(u::('q_46418::type, 'q_46425::type) 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))))))" + by (import hollight CART_EQ) + +constdefs + lambda :: "(nat => 'A::type) => ('A::type, 'B::type) cart" + "(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) + (%u::nat => 'A::type. + (Eps::(('A::type, 'B::type) cart => bool) => ('A::type, 'B::type) cart) + (%f::('A::type, 'B::type) cart. + (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) f i) + (u i)))))" + +lemma DEF_lambda: "(op =::((nat => 'A::type) => ('A::type, 'B::type) cart) + => ((nat => 'A::type) => ('A::type, 'B::type) cart) => bool) + (lambda::(nat => 'A::type) => ('A::type, 'B::type) cart) + (%u::nat => 'A::type. + (Eps::(('A::type, 'B::type) cart => bool) => ('A::type, 'B::type) cart) + (%f::('A::type, 'B::type) cart. + (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) 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)))" + 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)))" + by (import hollight LAMBDA_UNIQUE) + +lemma LAMBDA_ETA: "ALL x::('q_46616::type, 'q_46620::type) cart. lambda ($ x) = x" + by (import hollight LAMBDA_ETA) + +typedef (open) ('A, 'B) finite_sum = "(Collect::(('A::type finite_image, 'B::type finite_image) sum => bool) + => ('A::type finite_image, 'B::type finite_image) sum set) + (%x::('A::type finite_image, 'B::type finite_image) sum. 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) sum => bool) + => ('A::type finite_image, 'B::type finite_image) sum) + (%x::('A::type finite_image, 'B::type finite_image) sum. True::bool)"]) + by (import hollight TYDEF_finite_sum) + +syntax + dest_finite_sum :: _ + +syntax + mk_finite_sum :: _ + +lemmas "TYDEF_finite_sum_@intern" = typedef_hol2hollight + [where a="a :: ('A::type, 'B::type) finite_sum" and r=r , + OF type_definition_finite_sum] + +constdefs + pastecart :: "('A::type, 'M::type) cart +=> ('A::type, 'N::type) cart + => ('A::type, ('M::type, 'N::type) finite_sum) cart" + "(op ==::(('A::type, 'M::type) cart + => ('A::type, 'N::type) cart + => ('A::type, ('M::type, 'N::type) finite_sum) cart) + => (('A::type, 'M::type) cart + => ('A::type, 'N::type) cart + => ('A::type, ('M::type, 'N::type) finite_sum) cart) + => prop) + (pastecart::('A::type, 'M::type) cart + => ('A::type, 'N::type) cart + => ('A::type, ('M::type, 'N::type) finite_sum) cart) + (%(u::('A::type, 'M::type) cart) ua::('A::type, 'N::type) cart. + (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 + ((dimindex::('M::type => bool) => nat) + (hollight.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))))))" + +lemma DEF_pastecart: "(op =::(('A::type, 'M::type) cart + => ('A::type, 'N::type) cart + => ('A::type, ('M::type, 'N::type) finite_sum) cart) + => (('A::type, 'M::type) cart + => ('A::type, 'N::type) cart + => ('A::type, ('M::type, 'N::type) finite_sum) cart) + => bool) + (pastecart::('A::type, 'M::type) cart + => ('A::type, 'N::type) cart + => ('A::type, ('M::type, 'N::type) finite_sum) cart) + (%(u::('A::type, 'M::type) cart) ua::('A::type, 'N::type) cart. + (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 + ((dimindex::('M::type => bool) => nat) + (hollight.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))))))" + by (import hollight DEF_pastecart) + +constdefs + fstcart :: "('A::type, ('M::type, 'N::type) finite_sum) cart +=> ('A::type, 'M::type) cart" + "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))" + by (import hollight DEF_fstcart) + +constdefs + sndcart :: "('A::type, ('M::type, 'N::type) finite_sum) cart +=> ('A::type, 'N::type) cart" + "(op ==::(('A::type, ('M::type, 'N::type) finite_sum) cart + => ('A::type, 'N::type) cart) + => (('A::type, ('M::type, 'N::type) finite_sum) cart + => ('A::type, 'N::type) cart) + => prop) + (sndcart::('A::type, ('M::type, 'N::type) finite_sum) cart + => ('A::type, 'N::type) cart) + (%u::('A::type, ('M::type, 'N::type) finite_sum) cart. + (lambda::(nat => 'A::type) => ('A::type, 'N::type) cart) + (%i::nat. + ($::('A::type, ('M::type, 'N::type) finite_sum) cart + => nat => 'A::type) + u ((op +::nat => nat => nat) i + ((dimindex::('M::type => bool) => nat) + (hollight.UNIV::'M::type => bool)))))" + +lemma DEF_sndcart: "(op =::(('A::type, ('M::type, 'N::type) finite_sum) cart + => ('A::type, 'N::type) cart) + => (('A::type, ('M::type, 'N::type) finite_sum) cart + => ('A::type, 'N::type) cart) + => bool) + (sndcart::('A::type, ('M::type, 'N::type) finite_sum) cart + => ('A::type, 'N::type) cart) + (%u::('A::type, ('M::type, 'N::type) finite_sum) cart. + (lambda::(nat => 'A::type) => ('A::type, 'N::type) cart) + (%i::nat. + ($::('A::type, ('M::type, 'N::type) finite_sum) cart + => nat => 'A::type) + u ((op +::nat => nat => nat) i + ((dimindex::('M::type => bool) => nat) + (hollight.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) + ((op +::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)) + ((op +::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_46940::type, ('q_46937::type, 'q_46935::type) finite_sum) cart. + pastecart (fstcart x) (sndcart x) = x" + by (import hollight PASTECART_FST_SND) + +lemma PASTECART_EQ: "ALL (x::('q_46978::type, ('q_46968::type, 'q_46979::type) finite_sum) cart) + y::('q_46978::type, ('q_46968::type, 'q_46979::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_46999::type, ('q_47000::type, 'q_47001::type) finite_sum) cart + => bool) = +(ALL (x::('q_46999::type, 'q_47000::type) cart) + y::('q_46999::type, 'q_47001::type) cart. P (pastecart x y))" + by (import hollight FORALL_PASTECART) + +lemma EXISTS_PASTECART: "Ex (P::('q_47021::type, ('q_47022::type, 'q_47023::type) finite_sum) cart + => bool) = +(EX (x::('q_47021::type, 'q_47022::type) cart) + y::('q_47021::type, 'q_47023::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_47348::type => bool) f::'q_47348::type => 'q_47348::type. + FINITE s & IMAGE f s = s --> + (ALL (x::'q_47348::type) y::'q_47348::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_47454::type => bool. + (ALL xa::'q_47454::type. IN xa EMPTY --> x xa) = True) & +(ALL (x::'q_47494::type => bool) (xa::'q_47494::type) + xb::'q_47494::type => bool. + (ALL xc::'q_47494::type. IN xc (INSERT xa xb) --> x xc) = + (x xa & (ALL xa::'q_47494::type. IN xa xb --> x xa)))" + by (import hollight FORALL_IN_CLAUSES) + +lemma EXISTS_IN_CLAUSES: "(ALL x::'q_47514::type => bool. + (EX xa::'q_47514::type. IN xa EMPTY & x xa) = False) & +(ALL (x::'q_47554::type => bool) (xa::'q_47554::type) + xb::'q_47554::type => bool. + (EX xc::'q_47554::type. IN xc (INSERT xa xb) & x xc) = + (x xa | (EX xa::'q_47554::type. IN xa xb & x xa)))" + by (import hollight EXISTS_IN_CLAUSES) + +lemma SURJECTIVE_ON_RIGHT_INVERSE: "ALL (x::'q_47610::type => 'q_47611::type) xa::'q_47611::type => bool. + (ALL xb::'q_47611::type. + IN xb xa --> + (EX xa::'q_47610::type. + IN xa (s::'q_47610::type => bool) & x xa = xb)) = + (EX g::'q_47611::type => 'q_47610::type. + ALL y::'q_47611::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_47704::type => 'q_47707::type) xa::'q_47704::type => bool. + (ALL (xb::'q_47704::type) y::'q_47704::type. + IN xb xa & IN y xa & x xb = x y --> xb = y) = + (EX xb::'q_47707::type => 'q_47704::type. + ALL xc::'q_47704::type. IN xc xa --> xb (x xc) = xc)" + by (import hollight INJECTIVE_ON_LEFT_INVERSE) + +lemma SURJECTIVE_RIGHT_INVERSE: "(ALL y::'q_47732::type. + EX x::'q_47735::type. (f::'q_47735::type => 'q_47732::type) x = y) = +(EX g::'q_47732::type => 'q_47735::type. ALL y::'q_47732::type. f (g y) = y)" + by (import hollight SURJECTIVE_RIGHT_INVERSE) + +lemma INJECTIVE_LEFT_INVERSE: "(ALL (x::'q_47769::type) xa::'q_47769::type. + (f::'q_47769::type => 'q_47772::type) x = f xa --> x = xa) = +(EX g::'q_47772::type => 'q_47769::type. ALL x::'q_47769::type. g (f x) = x)" + by (import hollight INJECTIVE_LEFT_INVERSE) + +lemma FUNCTION_FACTORS_RIGHT: "ALL (x::'q_47808::type => 'q_47809::type) + xa::'q_47796::type => 'q_47809::type. + (ALL xb::'q_47808::type. EX y::'q_47796::type. xa y = x xb) = + (EX xb::'q_47808::type => 'q_47796::type. x = xa o xb)" + by (import hollight FUNCTION_FACTORS_RIGHT) + +lemma FUNCTION_FACTORS_LEFT: "ALL (x::'q_47881::type => 'q_47882::type) + xa::'q_47881::type => 'q_47861::type. + (ALL (xb::'q_47881::type) y::'q_47881::type. + xa xb = xa y --> x xb = x y) = + (EX xb::'q_47861::type => 'q_47882::type. x = xb o xa)" + by (import hollight FUNCTION_FACTORS_LEFT) + +constdefs + dotdot :: "nat => nat => nat => bool" + "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_47957::type) (p::nat) m::nat. + <= m p & <= p (n::nat) --> + hollight.UNION (dotdot m p) (dotdot (p + NUMERAL_BIT1 (0::nat)) n) = + dotdot m n" + by (import hollight NUMSEG_COMBINE_R) + +lemma NUMSEG_COMBINE_L: "ALL (x::'q_47995::type) (p::nat) m::nat. + <= m p & <= p (n::nat) --> + hollight.UNION (dotdot m (p - NUMERAL_BIT1 (0::nat))) (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::nat)) 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::nat))) = 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::nat)" + by (import hollight CARD_NUMSEG_LEMMA) + +lemma CARD_NUMSEG: "ALL (m::nat) n::nat. CARD (dotdot m n) = n + NUMERAL_BIT1 (0::nat) - m" + by (import hollight CARD_NUMSEG) + +lemma HAS_SIZE_NUMSEG: "ALL (x::nat) xa::nat. + HAS_SIZE (dotdot x xa) (xa + NUMERAL_BIT1 (0::nat) - x)" + by (import hollight HAS_SIZE_NUMSEG) + +lemma CARD_NUMSEG_1: "ALL x::nat. CARD (dotdot (NUMERAL_BIT1 (0::nat)) x) = x" + by (import hollight CARD_NUMSEG_1) + +lemma HAS_SIZE_NUMSEG_1: "ALL x::nat. HAS_SIZE (dotdot (NUMERAL_BIT1 (0::nat)) x) x" + by (import hollight HAS_SIZE_NUMSEG_1) + +lemma NUMSEG_CLAUSES: "(ALL m::nat. + dotdot m (0::nat) = COND (m = (0::nat)) (INSERT (0::nat) 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::nat)) (CARD s)) & + IN j (dotdot (NUMERAL_BIT1 (0::nat)) (CARD s)) & f i = f j --> + i = j) & + s = IMAGE f (dotdot (NUMERAL_BIT1 (0::nat)) (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::nat)) --> + dotdot x (xa + xb) = + hollight.UNION (dotdot x xa) + (dotdot (xa + NUMERAL_BIT1 (0::nat)) (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) + +constdefs + neutral :: "('q_48776::type => 'q_48776::type => 'q_48776::type) => 'q_48776::type" + "neutral == +%u::'q_48776::type => 'q_48776::type => 'q_48776::type. + SOME x::'q_48776::type. ALL y::'q_48776::type. u x y = y & u y x = y" + +lemma DEF_neutral: "neutral = +(%u::'q_48776::type => 'q_48776::type => 'q_48776::type. + SOME x::'q_48776::type. ALL y::'q_48776::type. u x y = y & u y x = y)" + by (import hollight DEF_neutral) + +constdefs + monoidal :: "('A::type => 'A::type => 'A::type) => bool" + "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) + +constdefs + support :: "('B::type => 'B::type => 'B::type) +=> ('A::type => 'B::type) => ('A::type => bool) => 'A::type => bool" + "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) + +constdefs + iterate :: "('q_48881::type => 'q_48881::type => 'q_48881::type) +=> ('A::type => bool) => ('A::type => 'q_48881::type) => 'q_48881::type" + "iterate == +%(u::'q_48881::type => 'q_48881::type => 'q_48881::type) + (ua::'A::type => bool) ub::'A::type => 'q_48881::type. + ITSET (%x::'A::type. u (ub x)) (support u ub ua) (neutral u)" + +lemma DEF_iterate: "iterate = +(%(u::'q_48881::type => 'q_48881::type => 'q_48881::type) + (ua::'A::type => bool) ub::'A::type => 'q_48881::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_48924::type => 'q_48924::type => 'q_48924::type) + (xa::'q_48927::type => 'q_48924::type) (xb::'q_48927::type) + xc::'q_48927::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_48949::type => 'q_48949::type => 'q_48949::type) + (xa::'q_48960::type => 'q_48949::type) xb::'q_48960::type => bool. + support x xa (support x xa xb) = support x xa xb" + by (import hollight SUPPORT_SUPPORT) + +lemma SUPPORT_EMPTY: "ALL (x::'q_48985::type => 'q_48985::type => 'q_48985::type) + (xa::'q_48999::type => 'q_48985::type) xb::'q_48999::type => bool. + (ALL xc::'q_48999::type. IN xc xb --> xa xc = neutral x) = + (support x xa xb = EMPTY)" + by (import hollight SUPPORT_EMPTY) + +lemma SUPPORT_SUBSET: "ALL (x::'q_49019::type => 'q_49019::type => 'q_49019::type) + (xa::'q_49020::type => 'q_49019::type) xb::'q_49020::type => bool. + SUBSET (support x xa xb) xb" + by (import hollight SUPPORT_SUBSET) + +lemma FINITE_SUPPORT: "ALL (u::'q_49043::type => 'q_49043::type => 'q_49043::type) + (f::'q_49037::type => 'q_49043::type) s::'q_49037::type => bool. + FINITE s --> FINITE (support u f s)" + by (import hollight FINITE_SUPPORT) + +lemma SUPPORT_CLAUSES: "(ALL x::'q_49061::type => 'q_49091::type. + support (u_4215::'q_49091::type => 'q_49091::type => 'q_49091::type) x + EMPTY = + EMPTY) & +(ALL (x::'q_49109::type => 'q_49091::type) (xa::'q_49109::type) + xb::'q_49109::type => bool. + support u_4215 x (INSERT xa xb) = + COND (x xa = neutral u_4215) (support u_4215 x xb) + (INSERT xa (support u_4215 x xb))) & +(ALL (x::'q_49142::type => 'q_49091::type) (xa::'q_49142::type) + xb::'q_49142::type => bool. + support u_4215 x (DELETE xb xa) = DELETE (support u_4215 x xb) xa) & +(ALL (x::'q_49180::type => 'q_49091::type) (xa::'q_49180::type => bool) + xb::'q_49180::type => bool. + support u_4215 x (hollight.UNION xa xb) = + hollight.UNION (support u_4215 x xa) (support u_4215 x xb)) & +(ALL (x::'q_49218::type => 'q_49091::type) (xa::'q_49218::type => bool) + xb::'q_49218::type => bool. + support u_4215 x (hollight.INTER xa xb) = + hollight.INTER (support u_4215 x xa) (support u_4215 x xb)) & +(ALL (x::'q_49254::type => 'q_49091::type) (xa::'q_49254::type => bool) + xb::'q_49254::type => bool. + support u_4215 x (DIFF xa xb) = + DIFF (support u_4215 x xa) (support u_4215 x xb))" + by (import hollight SUPPORT_CLAUSES) + +lemma ITERATE_SUPPORT: "ALL (x::'q_49275::type => 'q_49275::type => 'q_49275::type) + (xa::'q_49276::type => 'q_49275::type) xb::'q_49276::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_49320::type => 'q_49320::type => 'q_49320::type) + (xa::'q_49348::type => bool) (xb::'q_49348::type => 'q_49320::type) + xc::'q_49348::type. + support x (%xa::'q_49348::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_49369::type => 'q_49369::type => 'q_49369::type) + (xa::'q_49378::type => 'q_49369::type) xb::'q_49378::type. + FINITE + (support x (%xc::'q_49378::type. COND (xc = xb) (xa xc) (neutral x)) + (s::'q_49378::type => bool))" + by (import hollight FINITE_SUPPORT_DELTA) + +lemma ITERATE_CLAUSES_GEN: "ALL u_4215::'B::type => 'B::type => 'B::type. + monoidal u_4215 --> + (ALL f::'A::type => 'B::type. iterate u_4215 EMPTY f = neutral u_4215) & + (ALL (f::'A::type => 'B::type) (x::'A::type) s::'A::type => bool. + monoidal u_4215 & FINITE (support u_4215 f s) --> + iterate u_4215 (INSERT x s) f = + COND (IN x s) (iterate u_4215 s f) + (u_4215 (f x) (iterate u_4215 s f)))" + by (import hollight ITERATE_CLAUSES_GEN) + +lemma ITERATE_CLAUSES: "ALL x::'q_49546::type => 'q_49546::type => 'q_49546::type. + monoidal x --> + (ALL f::'q_49504::type => 'q_49546::type. + iterate x EMPTY f = neutral x) & + (ALL (f::'q_49548::type => 'q_49546::type) (xa::'q_49548::type) + s::'q_49548::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_4215::'q_49634::type => 'q_49634::type => 'q_49634::type. + monoidal u_4215 --> + (ALL (f::'q_49619::type => 'q_49634::type) (s::'q_49619::type => bool) + x::'q_49619::type => bool. + FINITE s & FINITE x & DISJOINT s x --> + iterate u_4215 (hollight.UNION s x) f = + u_4215 (iterate u_4215 s f) (iterate u_4215 x f))" + by (import hollight ITERATE_UNION) + +lemma ITERATE_UNION_GEN: "ALL u_4215::'B::type => 'B::type => 'B::type. + monoidal u_4215 --> + (ALL (f::'A::type => 'B::type) (s::'A::type => bool) t::'A::type => bool. + FINITE (support u_4215 f s) & + FINITE (support u_4215 f t) & + DISJOINT (support u_4215 f s) (support u_4215 f t) --> + iterate u_4215 (hollight.UNION s t) f = + u_4215 (iterate u_4215 s f) (iterate u_4215 t f))" + by (import hollight ITERATE_UNION_GEN) + +lemma ITERATE_DIFF: "ALL u::'q_49777::type => 'q_49777::type => 'q_49777::type. + monoidal u --> + (ALL (f::'q_49773::type => 'q_49777::type) (s::'q_49773::type => bool) + t::'q_49773::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_4215::'B::type => 'B::type => 'B::type. + monoidal u_4215 --> + (ALL (f::'A::type => 'B::type) (s::'A::type => bool) t::'A::type => bool. + FINITE (support u_4215 f s) & + SUBSET (support u_4215 f t) (support u_4215 f s) --> + u_4215 (iterate u_4215 (DIFF s t) f) (iterate u_4215 t f) = + iterate u_4215 s f)" + by (import hollight ITERATE_DIFF_GEN) + +lemma ITERATE_CLOSED: "ALL u_4215::'B::type => 'B::type => 'B::type. + monoidal u_4215 --> + (ALL P::'B::type => bool. + P (neutral u_4215) & + (ALL (x::'B::type) y::'B::type. P x & P y --> P (u_4215 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_4215 x f)))" + by (import hollight ITERATE_CLOSED) + +lemma ITERATE_CLOSED_GEN: "ALL u_4215::'B::type => 'B::type => 'B::type. + monoidal u_4215 --> + (ALL P::'B::type => bool. + P (neutral u_4215) & + (ALL (x::'B::type) y::'B::type. P x & P y --> P (u_4215 x y)) --> + (ALL (f::'A::type => 'B::type) s::'A::type => bool. + FINITE (support u_4215 f s) & + (ALL x::'A::type. IN x s & f x ~= neutral u_4215 --> P (f x)) --> + P (iterate u_4215 s f)))" + by (import hollight ITERATE_CLOSED_GEN) + +lemma ITERATE_RELATED: "ALL u_4215::'B::type => 'B::type => 'B::type. + monoidal u_4215 --> + (ALL R::'B::type => 'B::type => bool. + R (neutral u_4215) (neutral u_4215) & + (ALL (x1::'B::type) (y1::'B::type) (x2::'B::type) y2::'B::type. + R x1 x2 & R y1 y2 --> R (u_4215 x1 y1) (u_4215 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_4215 x f) (iterate u_4215 x g)))" + by (import hollight ITERATE_RELATED) + +lemma ITERATE_EQ_NEUTRAL: "ALL u_4215::'B::type => 'B::type => 'B::type. + monoidal u_4215 --> + (ALL (f::'A::type => 'B::type) s::'A::type => bool. + (ALL x::'A::type. IN x s --> f x = neutral u_4215) --> + iterate u_4215 s f = neutral u_4215)" + 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_4215::'q_50229::type => 'q_50229::type => 'q_50229::type. + monoidal u_4215 --> + (ALL (x::'q_50248::type => 'q_50229::type) (xa::'q_50248::type) + xb::'q_50248::type => bool. + iterate u_4215 xb + (%xb::'q_50248::type. COND (xb = xa) (x xb) (neutral u_4215)) = + COND (IN xa xb) (x xa) (neutral u_4215))" + by (import hollight ITERATE_DELTA) + +lemma ITERATE_IMAGE: "ALL u_4215::'q_50327::type => 'q_50327::type => 'q_50327::type. + monoidal u_4215 --> + (ALL (f::'q_50326::type => 'q_50298::type) + (g::'q_50298::type => 'q_50327::type) x::'q_50326::type => bool. + FINITE x & + (ALL (xa::'q_50326::type) y::'q_50326::type. + IN xa x & IN y x & f xa = f y --> xa = y) --> + iterate u_4215 (IMAGE f x) g = iterate u_4215 x (g o f))" + by (import hollight ITERATE_IMAGE) + +constdefs + nsum :: "('q_50348::type => bool) => ('q_50348::type => nat) => nat" + "(op ==::(('q_50348::type => bool) => ('q_50348::type => nat) => nat) + => (('q_50348::type => bool) => ('q_50348::type => nat) => nat) + => prop) + (nsum::('q_50348::type => bool) => ('q_50348::type => nat) => nat) + ((iterate::(nat => nat => nat) + => ('q_50348::type => bool) => ('q_50348::type => nat) => nat) + (op +::nat => nat => nat))" + +lemma DEF_nsum: "(op =::(('q_50348::type => bool) => ('q_50348::type => nat) => nat) + => (('q_50348::type => bool) => ('q_50348::type => nat) => nat) + => bool) + (nsum::('q_50348::type => bool) => ('q_50348::type => nat) => nat) + ((iterate::(nat => nat => nat) + => ('q_50348::type => bool) => ('q_50348::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 * = NUMERAL_BIT1 (0::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_50386::type => nat. nsum EMPTY x = (0::nat)) & +(ALL (x::'q_50425::type) (xa::'q_50425::type => nat) + xb::'q_50425::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_50451::type => nat) (xa::'q_50451::type => bool) + xb::'q_50451::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_50506::type => nat) (s::'q_50506::type => bool) + t::'q_50506::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_50545::type => nat) xa::'q_50545::type => bool. + FINITE (support op + x xa) --> nsum (support op + x xa) x = nsum xa x" + by (import hollight NSUM_SUPPORT) + +lemma NSUM_ADD: "ALL (f::'q_50591::type => nat) (g::'q_50591::type => nat) + s::'q_50591::type => bool. + FINITE s --> nsum s (%x::'q_50591::type. f x + g x) = nsum s f + nsum s g" + by (import hollight NSUM_ADD) + +lemma NSUM_CMUL: "ALL (f::'q_50629::type => nat) (c::nat) s::'q_50629::type => bool. + FINITE s --> nsum s (%x::'q_50629::type. c * f x) = c * nsum s f" + by (import hollight NSUM_CMUL) + +lemma NSUM_LE: "ALL (x::'q_50668::type => nat) (xa::'q_50668::type => nat) + xb::'q_50668::type => bool. + FINITE xb & (ALL xc::'q_50668::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_50790::type => nat) (g::'q_50790::type => nat) + s::'q_50790::type => bool. + FINITE s & + s ~= EMPTY & (ALL x::'q_50790::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_50832::type => nat) (xa::'q_50832::type => nat) + xb::'q_50832::type => bool. + FINITE xb & (ALL xc::'q_50832::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_50862::type => bool. + FINITE s --> nsum s (%n::'q_50862::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::nat)) --> nsum xa x = (0::nat)" + by (import hollight NSUM_EQ_0) + +lemma NSUM_0: "ALL x::'A::type => bool. nsum x (%n::'A::type. 0::nat) = (0::nat)" + by (import hollight NSUM_0) + +lemma NSUM_POS_LE: "ALL (x::'q_50941::type => nat) xa::'q_50941::type => bool. + FINITE xa & (ALL xb::'q_50941::type. IN xb xa --> <= (0::nat) (x xb)) --> + <= (0::nat) (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::nat) (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_51076::type => nat) xa::'q_51076::type => bool. + FINITE xa & + (ALL xb::'q_51076::type. IN xb xa --> <= (0::nat) (x xb)) & + nsum xa x = (0::nat) --> + (ALL xb::'q_51076::type. IN xb xa --> x xb = (0::nat))" + by (import hollight NSUM_POS_EQ_0) + +lemma NSUM_SING: "ALL (x::'q_51096::type => nat) xa::'q_51096::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::nat)) = + COND (IN xa x) b (0::nat)" + 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_51236::type => 'q_51212::type) (xa::'q_51212::type => nat) + xb::'q_51236::type => bool. + FINITE xb & + (ALL (xa::'q_51236::type) y::'q_51236::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::nat)) --> + 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::nat)) --> + 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::nat)) --> + nsum (hollight.UNION u v) f = nsum v f" + by (import hollight NSUM_UNION_LZERO) + +lemma NSUM_RESTRICT: "ALL (f::'q_51457::type => nat) s::'q_51457::type => bool. + FINITE s --> + nsum s (%x::'q_51457::type. COND (IN x s) (f x) (0::nat)) = 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_51532::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_51675::type => bool) (f::'q_51675::type => nat) b::nat. + FINITE s & s ~= EMPTY & (ALL x::'q_51675::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_51717::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_51776::type => bool) (t::'q_51776::type => bool) + u::'q_51776::type => bool. + FINITE u & hollight.INTER s t = EMPTY & hollight.UNION s t = u --> + nsum s (f::'q_51776::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::nat)) --> + 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_51887::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::nat))" + by (import hollight NSUM_RESTRICT_SET) + +lemma NSUM_NSUM_RESTRICT: "ALL (R::'q_52016::type => 'q_52015::type => bool) + (f::'q_52016::type => 'q_52015::type => nat) (s::'q_52016::type => bool) + t::'q_52015::type => bool. + FINITE s & FINITE t --> + nsum s + (%x::'q_52016::type. + nsum + (GSPEC + (%u::'q_52015::type. + EX y::'q_52015::type. SETSPEC u (IN y t & R x y) y)) + (f x)) = + nsum t + (%y::'q_52015::type. + nsum + (GSPEC + (%u::'q_52016::type. + EX x::'q_52016::type. SETSPEC u (IN x s & R x y) x)) + (%x::'q_52016::type. f x y))" + by (import hollight NSUM_NSUM_RESTRICT) + +lemma CARD_EQ_NSUM: "ALL x::'q_52035::type => bool. + FINITE x --> CARD x = nsum x (%x::'q_52035::type. NUMERAL_BIT1 (0::nat))" + 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::nat)) --> + <= (nsum u f) (nsum v f)" + by (import hollight NSUM_SUBSET) + +lemma NSUM_SUBSET_SIMPLE: "ALL (u::'q_52495::type => bool) (v::'q_52495::type => bool) + f::'q_52495::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::nat) - xa) * x" + by (import hollight NSUM_CONST_NUMSEG) + +lemma NSUM_EQ_0_NUMSEG: "ALL (x::nat => nat) xa::'q_52734::type. + (ALL i::nat. <= (m::nat) i & <= i (n::nat) --> x i = (0::nat)) --> + nsum (dotdot m n) x = (0::nat)" + 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::nat)" + 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::nat) (xb p)) --> + <= (0::nat) (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::nat) (f p)) & + nsum (dotdot m n) f = (0::nat) --> + (ALL p::nat. <= m p & <= p n --> f p = (0::nat))" + 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::nat)) (f::nat => nat) = + COND (x = (0::nat)) (f (0::nat)) (0::nat)) & +(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::nat)) --> + nsum (dotdot xa (xb + xc)) x = + nsum (dotdot xa xb) x + + nsum (dotdot (xb + NUMERAL_BIT1 (0::nat)) (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::nat) (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::nat)) xb) x" + by (import hollight NSUM_CLAUSES_LEFT) + +lemma NSUM_CLAUSES_RIGHT: "ALL (f::nat => nat) (m::nat) n::nat. + < (0::nat) n & <= m n --> + nsum (dotdot m n) f = nsum (dotdot m (n - NUMERAL_BIT1 (0::nat))) f + f n" + by (import hollight NSUM_CLAUSES_RIGHT) + +consts + sum :: "('q_53311::type => bool) +=> ('q_53311::type => hollight.real) => hollight.real" + +defs + sum_def: "(op ==::(('q_53311::type => bool) + => ('q_53311::type => hollight.real) => hollight.real) + => (('q_53311::type => bool) + => ('q_53311::type => hollight.real) => hollight.real) + => prop) + (hollight.sum::('q_53311::type => bool) + => ('q_53311::type => hollight.real) => hollight.real) + ((iterate::(hollight.real => hollight.real => hollight.real) + => ('q_53311::type => bool) + => ('q_53311::type => hollight.real) => hollight.real) + (real_add::hollight.real => hollight.real => hollight.real))" + +lemma DEF_sum: "(op =::(('q_53311::type => bool) + => ('q_53311::type => hollight.real) => hollight.real) + => (('q_53311::type => bool) + => ('q_53311::type => hollight.real) => hollight.real) + => bool) + (hollight.sum::('q_53311::type => bool) + => ('q_53311::type => hollight.real) => hollight.real) + ((iterate::(hollight.real => hollight.real => hollight.real) + => ('q_53311::type => bool) + => ('q_53311::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::nat)" + by (import hollight NEUTRAL_REAL_ADD) + +lemma NEUTRAL_REAL_MUL: "neutral real_mul = real_of_num (NUMERAL_BIT1 (0::nat))" + 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_53353::type => hollight.real. + hollight.sum EMPTY x = real_of_num (0::nat)) & +(ALL (x::'q_53394::type) (xa::'q_53394::type => hollight.real) + xb::'q_53394::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_53420::type => hollight.real) (xa::'q_53420::type => bool) + xb::'q_53420::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_SUPPORT: "ALL (x::'q_53499::type => hollight.real) xa::'q_53499::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_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_53825::type => hollight.real) + (g::'q_53825::type => hollight.real) s::'q_53825::type => bool. + FINITE s & + s ~= EMPTY & (ALL x::'q_53825::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_POS_LE: "ALL (x::'q_54040::type => hollight.real) xa::'q_54040::type => bool. + FINITE xa & + (ALL xb::'q_54040::type. + IN xb xa --> real_le (real_of_num (0::nat)) (x xb)) --> + real_le (real_of_num (0::nat)) (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::nat)) (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_54187::type => hollight.real) xa::'q_54187::type => bool. + FINITE xa & + (ALL xb::'q_54187::type. + IN xb xa --> real_le (real_of_num (0::nat)) (x xb)) & + hollight.sum xa x = real_of_num (0::nat) --> + (ALL xb::'q_54187::type. IN xb xa --> x xb = real_of_num (0::nat))" + by (import hollight SUM_POS_EQ_0) + +lemma SUM_SING: "ALL (x::'q_54209::type => hollight.real) xa::'q_54209::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::nat))) = + COND (IN xa x) b (real_of_num (0::nat))" + by (import hollight SUM_DELTA) + +lemma SUM_IMAGE: "ALL (x::'q_54353::type => 'q_54329::type) + (xa::'q_54329::type => hollight.real) xb::'q_54353::type => bool. + FINITE xb & + (ALL (xa::'q_54353::type) y::'q_54353::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::nat)) --> + 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::nat)) --> + 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::nat)) --> + hollight.sum (hollight.UNION u v) f = hollight.sum v f" + by (import hollight SUM_UNION_LZERO) + +lemma SUM_RESTRICT: "ALL (f::'q_54580::type => hollight.real) s::'q_54580::type => bool. + FINITE s --> + hollight.sum s + (%x::'q_54580::type. COND (IN x s) (f x) (real_of_num (0::nat))) = + hollight.sum s f" + by (import hollight SUM_RESTRICT) + +lemma SUM_BOUND_GEN: "ALL (s::'A::type => bool) (t::'q_54639::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_54844::type => bool) (f::'q_54844::type => hollight.real) + b::hollight.real. + FINITE s & + s ~= EMPTY & (ALL x::'q_54844::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_54866::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_54927::type => bool) (t::'q_54927::type => bool) + u::'q_54927::type => bool. + FINITE u & hollight.INTER s t = EMPTY & hollight.UNION s t = u --> + real_add (hollight.sum s (f::'q_54927::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::nat)) --> + 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_55040::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::nat)))" + by (import hollight SUM_RESTRICT_SET) + +lemma SUM_SUM_RESTRICT: "ALL (R::'q_55171::type => 'q_55170::type => bool) + (f::'q_55171::type => 'q_55170::type => hollight.real) + (s::'q_55171::type => bool) t::'q_55170::type => bool. + FINITE s & FINITE t --> + hollight.sum s + (%x::'q_55171::type. + hollight.sum + (GSPEC + (%u::'q_55170::type. + EX y::'q_55170::type. SETSPEC u (IN y t & R x y) y)) + (f x)) = + hollight.sum t + (%y::'q_55170::type. + hollight.sum + (GSPEC + (%u::'q_55171::type. + EX x::'q_55171::type. SETSPEC u (IN x s & R x y) x)) + (%x::'q_55171::type. f x y))" + by (import hollight SUM_SUM_RESTRICT) + +lemma CARD_EQ_SUM: "ALL x::'q_55192::type => bool. + FINITE x --> + real_of_num (CARD x) = + hollight.sum x (%x::'q_55192::type. real_of_num (NUMERAL_BIT1 (0::nat)))" + 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_55589::type => nat) s::'q_55589::type => bool. + FINITE s --> + real_of_num (nsum s f) = + hollight.sum s (%x::'q_55589::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::nat))) & + (ALL x::'A::type. + IN x (DIFF 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: "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::nat)) (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::nat) - xa)) x" + by (import hollight SUM_CONST_NUMSEG) + +lemma SUM_EQ_0_NUMSEG: "ALL (x::nat => hollight.real) xa::'q_56115::type. + (ALL i::nat. + <= (m::nat) i & <= i (n::nat) --> x i = real_of_num (0::nat)) --> + hollight.sum (dotdot m n) x = real_of_num (0::nat)" + 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::nat)" + 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::nat)) (xb p)) --> + real_le (real_of_num (0::nat)) (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::nat)) (f p)) & + hollight.sum (dotdot m n) f = real_of_num (0::nat) --> + (ALL p::nat. <= m p & <= p n --> f p = real_of_num (0::nat))" + 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::nat)) (f::nat => hollight.real) = + COND (x = (0::nat)) (f (0::nat)) (real_of_num (0::nat))) & +(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::nat)) --> + hollight.sum (dotdot xa (xb + xc)) x = + real_add (hollight.sum (dotdot xa xb) x) + (hollight.sum (dotdot (xb + NUMERAL_BIT1 (0::nat)) (xb + xc)) x)" + by (import hollight SUM_ADD_SPLIT) + +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::nat) (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::nat)) xb) x)" + by (import hollight SUM_CLAUSES_LEFT) + +lemma SUM_CLAUSES_RIGHT: "ALL (f::nat => hollight.real) (m::nat) n::nat. + < (0::nat) n & <= m n --> + hollight.sum (dotdot m n) f = + real_add (hollight.sum (dotdot m (n - NUMERAL_BIT1 (0::nat))) 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) + +constdefs + CASEWISE :: "(('q_56787::type => 'q_56791::type) * + ('q_56792::type => 'q_56787::type => 'q_56751::type)) hollight.list +=> 'q_56792::type => 'q_56791::type => 'q_56751::type" + "CASEWISE == +SOME CASEWISE::(('q_56787::type => 'q_56791::type) * + ('q_56792::type + => 'q_56787::type => 'q_56751::type)) hollight.list + => 'q_56792::type => 'q_56791::type => 'q_56751::type. + (ALL (f::'q_56792::type) x::'q_56791::type. + CASEWISE NIL f x = (SOME y::'q_56751::type. True)) & + (ALL (h::('q_56787::type => 'q_56791::type) * + ('q_56792::type => 'q_56787::type => 'q_56751::type)) + (t::(('q_56787::type => 'q_56791::type) * + ('q_56792::type + => 'q_56787::type => 'q_56751::type)) hollight.list) + (f::'q_56792::type) x::'q_56791::type. + CASEWISE (CONS h t) f x = + COND (EX y::'q_56787::type. fst h y = x) + (snd h f (SOME y::'q_56787::type. fst h y = x)) (CASEWISE t f x))" + +lemma DEF_CASEWISE: "CASEWISE = +(SOME CASEWISE::(('q_56787::type => 'q_56791::type) * + ('q_56792::type + => 'q_56787::type => 'q_56751::type)) hollight.list + => 'q_56792::type => 'q_56791::type => 'q_56751::type. + (ALL (f::'q_56792::type) x::'q_56791::type. + CASEWISE NIL f x = (SOME y::'q_56751::type. True)) & + (ALL (h::('q_56787::type => 'q_56791::type) * + ('q_56792::type => 'q_56787::type => 'q_56751::type)) + (t::(('q_56787::type => 'q_56791::type) * + ('q_56792::type + => 'q_56787::type => 'q_56751::type)) hollight.list) + (f::'q_56792::type) x::'q_56791::type. + CASEWISE (CONS h t) f x = + COND (EX y::'q_56787::type. fst h y = x) + (snd h f (SOME y::'q_56787::type. fst h y = x)) (CASEWISE t f x)))" + by (import hollight DEF_CASEWISE) + +lemma CASEWISE: "(op &::bool => bool => bool) + ((op =::'q_56811::type => 'q_56811::type => bool) + ((CASEWISE::(('q_56803::type => 'q_56851::type) * + ('q_56852::type + => 'q_56803::type => 'q_56811::type)) hollight.list + => 'q_56852::type => 'q_56851::type => 'q_56811::type) + (NIL::(('q_56803::type => 'q_56851::type) * + ('q_56852::type + => 'q_56803::type => 'q_56811::type)) hollight.list) + (f::'q_56852::type) (x::'q_56851::type)) + ((Eps::('q_56811::type => bool) => 'q_56811::type) + (%y::'q_56811::type. True::bool))) + ((op =::'q_56812::type => 'q_56812::type => bool) + ((CASEWISE::(('q_56854::type => 'q_56851::type) * + ('q_56852::type + => 'q_56854::type => 'q_56812::type)) hollight.list + => 'q_56852::type => 'q_56851::type => 'q_56812::type) + ((CONS::('q_56854::type => 'q_56851::type) * + ('q_56852::type => 'q_56854::type => 'q_56812::type) + => (('q_56854::type => 'q_56851::type) * + ('q_56852::type + => 'q_56854::type => 'q_56812::type)) hollight.list + => (('q_56854::type => 'q_56851::type) * + ('q_56852::type + => 'q_56854::type => 'q_56812::type)) hollight.list) + ((Pair::('q_56854::type => 'q_56851::type) + => ('q_56852::type => 'q_56854::type => 'q_56812::type) + => ('q_56854::type => 'q_56851::type) * + ('q_56852::type => 'q_56854::type => 'q_56812::type)) + (s::'q_56854::type => 'q_56851::type) + (t::'q_56852::type => 'q_56854::type => 'q_56812::type)) + (clauses::(('q_56854::type => 'q_56851::type) * + ('q_56852::type + => 'q_56854::type => 'q_56812::type)) hollight.list)) + f x) + ((COND::bool => 'q_56812::type => 'q_56812::type => 'q_56812::type) + ((Ex::('q_56854::type => bool) => bool) + (%y::'q_56854::type. + (op =::'q_56851::type => 'q_56851::type => bool) (s y) x)) + (t f ((Eps::('q_56854::type => bool) => 'q_56854::type) + (%y::'q_56854::type. + (op =::'q_56851::type => 'q_56851::type => bool) (s y) x))) + ((CASEWISE::(('q_56854::type => 'q_56851::type) * + ('q_56852::type + => 'q_56854::type => 'q_56812::type)) hollight.list + => 'q_56852::type => 'q_56851::type => 'q_56812::type) + clauses f x)))" + by (import hollight CASEWISE) + +lemma CASEWISE_CASES: "ALL (clauses::(('q_56946::type => 'q_56943::type) * + ('q_56944::type + => 'q_56946::type => 'q_56953::type)) hollight.list) + (c::'q_56944::type) x::'q_56943::type. + (EX (s::'q_56946::type => 'q_56943::type) + (t::'q_56944::type => 'q_56946::type => 'q_56953::type) + a::'q_56946::type. + MEM (s, t) clauses & s a = x & CASEWISE clauses c x = t c a) | + ~ (EX (s::'q_56946::type => 'q_56943::type) + (t::'q_56944::type => 'q_56946::type => 'q_56953::type) + a::'q_56946::type. MEM (s, t) clauses & s a = x) & + CASEWISE clauses c x = (SOME y::'q_56953::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) + +constdefs + admissible :: "('q_57089::type => 'q_57082::type => bool) +=> (('q_57089::type => 'q_57085::type) => 'q_57095::type => bool) + => ('q_57095::type => 'q_57082::type) + => (('q_57089::type => 'q_57085::type) + => 'q_57095::type => 'q_57090::type) + => bool" + "admissible == +%(u::'q_57089::type => 'q_57082::type => bool) + (ua::('q_57089::type => 'q_57085::type) => 'q_57095::type => bool) + (ub::'q_57095::type => 'q_57082::type) + uc::('q_57089::type => 'q_57085::type) + => 'q_57095::type => 'q_57090::type. + ALL (f::'q_57089::type => 'q_57085::type) + (g::'q_57089::type => 'q_57085::type) a::'q_57095::type. + ua f a & + ua g a & (ALL z::'q_57089::type. u z (ub a) --> f z = g z) --> + uc f a = uc g a" + +lemma DEF_admissible: "admissible = +(%(u::'q_57089::type => 'q_57082::type => bool) + (ua::('q_57089::type => 'q_57085::type) => 'q_57095::type => bool) + (ub::'q_57095::type => 'q_57082::type) + uc::('q_57089::type => 'q_57085::type) + => 'q_57095::type => 'q_57090::type. + ALL (f::'q_57089::type => 'q_57085::type) + (g::'q_57089::type => 'q_57085::type) a::'q_57095::type. + ua f a & + ua g a & (ALL z::'q_57089::type. u z (ub a) --> f z = g z) --> + uc f a = uc g a)" + by (import hollight DEF_admissible) + +constdefs + tailadmissible :: "('A::type => 'A::type => bool) +=> (('A::type => 'B::type) => 'P::type => bool) + => ('P::type => 'A::type) + => (('A::type => 'B::type) => 'P::type => 'B::type) => bool" + "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. + 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) --> + 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))" + +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. + 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) --> + 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)))" + by (import hollight DEF_tailadmissible) + +constdefs + superadmissible :: "('q_57239::type => 'q_57239::type => bool) +=> (('q_57239::type => 'q_57241::type) => 'q_57247::type => bool) + => ('q_57247::type => 'q_57239::type) + => (('q_57239::type => 'q_57241::type) + => 'q_57247::type => 'q_57241::type) + => bool" + "superadmissible == +%(u::'q_57239::type => 'q_57239::type => bool) + (ua::('q_57239::type => 'q_57241::type) => 'q_57247::type => bool) + (ub::'q_57247::type => 'q_57239::type) + uc::('q_57239::type => 'q_57241::type) + => 'q_57247::type => 'q_57241::type. + admissible u + (%(f::'q_57239::type => 'q_57241::type) a::'q_57247::type. True) ub + ua --> + tailadmissible u ua ub uc" + +lemma DEF_superadmissible: "superadmissible = +(%(u::'q_57239::type => 'q_57239::type => bool) + (ua::('q_57239::type => 'q_57241::type) => 'q_57247::type => bool) + (ub::'q_57247::type => 'q_57239::type) + uc::('q_57239::type => 'q_57241::type) + => 'q_57247::type => 'q_57241::type. + admissible u + (%(f::'q_57239::type => 'q_57241::type) a::'q_57247::type. 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_354::'A::type => 'q_57627::type => bool) + (p::('A::type => 'B::type) => 'P::type => bool) + (P::('A::type => 'B::type) => 'P::type => bool) + (s::'P::type => 'q_57627::type) + (h::('A::type => 'B::type) => 'P::type => 'B::type) + k::('A::type => 'B::type) => 'P::type => 'B::type. + admissible u_354 p s P & + admissible u_354 (%(f::'A::type => 'B::type) x::'P::type. p f x & P f x) + s h & + admissible u_354 + (%(f::'A::type => 'B::type) x::'P::type. p f x & ~ P f x) s k --> + admissible u_354 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_354::'q_57702::type => 'q_57701::type => bool) + (p::('q_57702::type => 'q_57703::type) => 'q_57704::type => bool) + (s::'q_57704::type => 'q_57701::type) + (%f::'q_57702::type => 'q_57703::type. c::'q_57704::type => 'q_57705::type)" + 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))" + 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))" + 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 WF_REC_CASES: "ALL (u_354::'A::type => 'A::type => bool) + clauses::(('P::type => 'A::type) * + (('A::type => 'B::type) + => 'P::type => 'B::type)) hollight.list. + WF u_354 & + 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_354 y (G f a) --> u_354 y (s a)) & + (ALL (f::'A::type => 'B::type) (g::'A::type => 'B::type) + a::'P::type. + (ALL z::'A::type. u_354 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_58634::type) + (p2::'q_58645::type => 'A::type => 'q_58639::type) + (p1'::'A::type => 'q_58634::type) + p2'::'q_58645::type => 'A::type => 'q_58639::type. + (ALL (c::'q_58645::type) (x::'A::type) y::'A::type. + p1 x = p1' y --> p2 c x = p2' c y) --> + (ALL (c::'q_58645::type) (x::'A::type) y::'A::type. + p1' x = p1 y --> p2' c x = p2 c y)" + by (import hollight cth) + +lemma RECURSION_CASEWISE_PAIRWISE: "ALL x::(('q_58682::type => 'q_58662::type) * + (('q_58662::type => 'q_58678::type) + => 'q_58682::type => 'q_58678::type)) hollight.list. + (EX u::'q_58662::type => 'q_58662::type => bool. + WF u & + ALL_list + (GABS + (%f::('q_58682::type => 'q_58662::type) * + (('q_58662::type => 'q_58678::type) + => 'q_58682::type => 'q_58678::type) + => bool. + ALL (s::'q_58682::type => 'q_58662::type) + t::('q_58662::type => 'q_58678::type) + => 'q_58682::type => 'q_58678::type. + GEQ (f (s, t)) + (tailadmissible u + (%(f::'q_58662::type => 'q_58678::type) + a::'q_58682::type. True) + s t))) + x) & + ALL_list + (GABS + (%f::('q_58682::type => 'q_58662::type) * + (('q_58662::type => 'q_58678::type) + => 'q_58682::type => 'q_58678::type) + => bool. + ALL (a::'q_58682::type => 'q_58662::type) + b::('q_58662::type => 'q_58678::type) + => 'q_58682::type => 'q_58678::type. + GEQ (f (a, b)) + (ALL (c::'q_58662::type => 'q_58678::type) (x::'q_58682::type) + y::'q_58682::type. a x = a y --> b c x = b c y))) + x & + PAIRWISE + (GABS + (%f::('q_58682::type => 'q_58662::type) * + (('q_58662::type => 'q_58678::type) + => 'q_58682::type => 'q_58678::type) + => ('q_58682::type => 'q_58662::type) * + (('q_58662::type => 'q_58678::type) + => 'q_58682::type => 'q_58678::type) + => bool. + ALL (s::'q_58682::type => 'q_58662::type) + t::('q_58662::type => 'q_58678::type) + => 'q_58682::type => 'q_58678::type. + GEQ (f (s, t)) + (GABS + (%f::('q_58682::type => 'q_58662::type) * + (('q_58662::type => 'q_58678::type) + => 'q_58682::type => 'q_58678::type) + => bool. + ALL (s'::'q_58682::type => 'q_58662::type) + t'::('q_58662::type => 'q_58678::type) + => 'q_58682::type => 'q_58678::type. + GEQ (f (s', t')) + (ALL (c::'q_58662::type => 'q_58678::type) + (x::'q_58682::type) y::'q_58682::type. + s x = s' y --> t c x = t' c y))))) + x --> + (EX f::'q_58662::type => 'q_58678::type. + ALL_list + (GABS + (%fa::('q_58682::type => 'q_58662::type) * + (('q_58662::type => 'q_58678::type) + => 'q_58682::type => 'q_58678::type) + => bool. + ALL (s::'q_58682::type => 'q_58662::type) + t::('q_58662::type => 'q_58678::type) + => 'q_58682::type => 'q_58678::type. + GEQ (fa (s, t)) (ALL x::'q_58682::type. f (s x) = t f x))) + x)" + by (import hollight RECURSION_CASEWISE_PAIRWISE) + +lemma SUPERADMISSIBLE_T: "superadmissible (u_354::'q_58792::type => 'q_58792::type => bool) + (%(f::'q_58792::type => 'q_58794::type) x::'q_58798::type. True) + (s::'q_58798::type => 'q_58792::type) + (t::('q_58792::type => 'q_58794::type) + => 'q_58798::type => 'q_58794::type) = +tailadmissible u_354 + (%(f::'q_58792::type => 'q_58794::type) x::'q_58798::type. True) s t" + by (import hollight SUPERADMISSIBLE_T) + +lemma SUB_SUB: "ALL (x::nat) xa::nat. <= xa x --> (ALL a::nat. a - (x - xa) = a + xa - x)" + by (import hollight SUB_SUB) + +lemma SUB_OLD: "(ALL m::nat. (0::nat) - m = (0::nat)) & +(ALL (m::nat) n::nat. Suc m - n = COND (< m n) (0::nat) (Suc (m - n)))" + by (import hollight SUB_OLD) + +lemma real_le: "ALL (x::hollight.real) xa::hollight.real. real_le x xa = (~ real_lt xa x)" + by (import hollight real_le) + +lemma REAL_MUL_RID: "ALL x::hollight.real. real_mul x (real_of_num (NUMERAL_BIT1 (0::nat))) = x" + by (import hollight REAL_MUL_RID) + +lemma REAL_MUL_RINV: "ALL x::hollight.real. + x ~= real_of_num (0::nat) --> + real_mul x (real_inv x) = real_of_num (NUMERAL_BIT1 (0::nat))" + by (import hollight REAL_MUL_RINV) + +lemma REAL_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)" + by (import hollight REAL_RDISTRIB) + +lemma REAL_EQ_LADD: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. + (real_add x y = real_add x z) = (y = z)" + by (import hollight REAL_EQ_LADD) + +lemma REAL_EQ_RADD: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. + (real_add x z = real_add y z) = (x = y)" + by (import hollight REAL_EQ_RADD) + +lemma REAL_ADD_LID_UNIQ: "ALL (x::hollight.real) y::hollight.real. + (real_add x y = y) = (x = real_of_num (0::nat))" + by (import hollight REAL_ADD_LID_UNIQ) + +lemma REAL_ADD_RID_UNIQ: "ALL (x::hollight.real) y::hollight.real. + (real_add x y = x) = (y = real_of_num (0::nat))" + by (import hollight REAL_ADD_RID_UNIQ) + +lemma REAL_LNEG_UNIQ: "ALL (x::hollight.real) y::hollight.real. + (real_add x y = real_of_num (0::nat)) = (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::nat)) = (y = real_neg x)" + by (import hollight REAL_RNEG_UNIQ) + +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)" + by (import hollight REAL_NEG_ADD) + +lemma REAL_MUL_LZERO: "ALL x::hollight.real. + real_mul (real_of_num (0::nat)) x = real_of_num (0::nat)" + by (import hollight REAL_MUL_LZERO) + +lemma REAL_MUL_RZERO: "ALL x::hollight.real. + real_mul x (real_of_num (0::nat)) = real_of_num (0::nat)" + by (import hollight REAL_MUL_RZERO) + +lemma REAL_NEG_LMUL: "ALL (x::hollight.real) y::hollight.real. + 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)" + by (import hollight REAL_NEG_RMUL) + +lemma REAL_NEGNEG: "ALL x::hollight.real. real_neg (real_neg x) = x" + by (import hollight REAL_NEGNEG) + +lemma REAL_NEG_MUL2: "ALL (x::hollight.real) xa::hollight.real. + real_mul (real_neg x) (real_neg xa) = real_mul x xa" + 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" + 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" + by (import hollight REAL_LT_RADD) + +lemma REAL_NOT_LT: "ALL (x::hollight.real) y::hollight.real. (~ real_lt x y) = real_le y x" + by (import hollight REAL_NOT_LT) + +lemma REAL_LT_ANTISYM: "ALL (x::hollight.real) y::hollight.real. ~ (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" + by (import hollight REAL_LT_GT) + +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_LT: "ALL (x::hollight.real) y::hollight.real. real_le x y = (real_lt x y | x = y)" + by (import hollight REAL_LE_LT) + +lemma REAL_LT_LE: "ALL (x::hollight.real) y::hollight.real. + real_lt x y = (real_le x y & x ~= y)" + by (import hollight REAL_LT_LE) + +lemma REAL_LT_IMP_LE: "ALL (x::hollight.real) y::hollight.real. 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" + by (import hollight REAL_LTE_TRANS) + +lemma REAL_LE_TRANS: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. + real_le x y & real_le y z --> real_le x z" + by (import hollight REAL_LE_TRANS) + +lemma REAL_NEG_LT0: "ALL x::hollight.real. + real_lt (real_neg x) (real_of_num (0::nat)) = + real_lt (real_of_num (0::nat)) x" + by (import hollight REAL_NEG_LT0) + +lemma REAL_NEG_GT0: "ALL x::hollight.real. + real_lt (real_of_num (0::nat)) (real_neg x) = + real_lt x (real_of_num (0::nat))" + by (import hollight REAL_NEG_GT0) + +lemma REAL_NEG_LE0: "ALL x::hollight.real. + real_le (real_neg x) (real_of_num (0::nat)) = + real_le (real_of_num (0::nat)) x" + by (import hollight REAL_NEG_LE0) + +lemma REAL_NEG_GE0: "ALL x::hollight.real. + real_le (real_of_num (0::nat)) (real_neg x) = + real_le x (real_of_num (0::nat))" + by (import hollight REAL_NEG_GE0) + +lemma REAL_LT_NEGTOTAL: "ALL x::hollight.real. + x = real_of_num (0::nat) | + real_lt (real_of_num (0::nat)) x | + real_lt (real_of_num (0::nat)) (real_neg x)" + by (import hollight REAL_LT_NEGTOTAL) + +lemma REAL_LE_NEGTOTAL: "ALL x::hollight.real. + real_le (real_of_num (0::nat)) x | + real_le (real_of_num (0::nat)) (real_neg x)" + by (import hollight REAL_LE_NEGTOTAL) + +lemma REAL_LE_MUL: "ALL (x::hollight.real) y::hollight.real. + real_le (real_of_num (0::nat)) x & real_le (real_of_num (0::nat)) y --> + real_le (real_of_num (0::nat)) (real_mul x y)" + by (import hollight REAL_LE_MUL) + +lemma REAL_LE_SQUARE: "ALL x::hollight.real. real_le (real_of_num (0::nat)) (real_mul x x)" + by (import hollight REAL_LE_SQUARE) + +lemma REAL_LT_01: "real_lt (real_of_num (0::nat)) (real_of_num (NUMERAL_BIT1 (0::nat)))" + 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" + 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" + 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)" + by (import hollight REAL_LT_ADD2) + +lemma REAL_LT_ADD: "ALL (x::hollight.real) y::hollight.real. + real_lt (real_of_num (0::nat)) x & real_lt (real_of_num (0::nat)) y --> + real_lt (real_of_num (0::nat)) (real_add x y)" + by (import hollight REAL_LT_ADD) + +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" + 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)" + 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::nat))))" + 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" + 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" + by (import hollight REAL_SUB_ADD2) + +lemma REAL_SUB_REFL: "ALL x::hollight.real. real_sub x x = real_of_num (0::nat)" + by (import hollight REAL_SUB_REFL) + +lemma REAL_SUB_0: "ALL (x::hollight.real) y::hollight.real. + (real_sub x y = real_of_num (0::nat)) = (x = y)" + by (import hollight REAL_SUB_0) + +lemma REAL_LE_DOUBLE: "ALL x::hollight.real. + real_le (real_of_num (0::nat)) (real_add x x) = + real_le (real_of_num (0::nat)) 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::nat)) 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::nat))" + by (import hollight REAL_LE_NEGR) + +lemma REAL_NEG_EQ0: "ALL x::hollight.real. + (real_neg x = real_of_num (0::nat)) = (x = real_of_num (0::nat))" + by (import hollight REAL_NEG_EQ0) + +lemma REAL_NEG_0: "real_neg (real_of_num (0::nat)) = real_of_num (0::nat)" + by (import hollight REAL_NEG_0) + +lemma REAL_NEG_SUB: "ALL (x::hollight.real) y::hollight.real. + real_neg (real_sub x y) = real_sub y x" + by (import hollight REAL_NEG_SUB) + +lemma REAL_SUB_LT: "ALL (x::hollight.real) y::hollight.real. + real_lt (real_of_num (0::nat)) (real_sub x y) = real_lt y x" + by (import hollight REAL_SUB_LT) + +lemma REAL_SUB_LE: "ALL (x::hollight.real) y::hollight.real. + real_le (real_of_num (0::nat)) (real_sub x y) = real_le y x" + by (import hollight REAL_SUB_LE) + +lemma REAL_EQ_LMUL: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. + (real_mul x y = real_mul x z) = (x = real_of_num (0::nat) | y = z)" + by (import hollight REAL_EQ_LMUL) + +lemma REAL_EQ_RMUL: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. + (real_mul x z = real_mul y z) = (z = real_of_num (0::nat) | x = y)" + by (import hollight REAL_EQ_RMUL) + +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)" + 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)" + by (import hollight REAL_SUB_RDISTRIB) + +lemma REAL_NEG_EQ: "ALL (x::hollight.real) y::hollight.real. (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::nat)))) x" + by (import hollight REAL_NEG_MINUS1) + +lemma REAL_INV_NZ: "ALL x::hollight.real. + x ~= real_of_num (0::nat) --> real_inv x ~= real_of_num (0::nat)" + by (import hollight REAL_INV_NZ) + +lemma REAL_INVINV: "ALL x::hollight.real. + x ~= real_of_num (0::nat) --> real_inv (real_inv x) = x" + by (import hollight REAL_INVINV) + +lemma REAL_LT_IMP_NE: "ALL (x::hollight.real) y::hollight.real. real_lt x y --> x ~= y" + by (import hollight REAL_LT_IMP_NE) + +lemma REAL_INV_POS: "ALL x::hollight.real. + real_lt (real_of_num (0::nat)) x --> + real_lt (real_of_num (0::nat)) (real_inv x)" + by (import hollight REAL_INV_POS) + +lemma REAL_LT_LMUL_0: "ALL (x::hollight.real) y::hollight.real. + real_lt (real_of_num (0::nat)) x --> + real_lt (real_of_num (0::nat)) (real_mul x y) = + real_lt (real_of_num (0::nat)) y" + by (import hollight REAL_LT_LMUL_0) + +lemma REAL_LT_RMUL_0: "ALL (x::hollight.real) y::hollight.real. + real_lt (real_of_num (0::nat)) y --> + real_lt (real_of_num (0::nat)) (real_mul x y) = + real_lt (real_of_num (0::nat)) x" + by (import hollight REAL_LT_RMUL_0) + +lemma REAL_LT_LMUL_EQ: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. + real_lt (real_of_num (0::nat)) x --> + real_lt (real_mul x y) (real_mul x z) = real_lt y z" + by (import hollight REAL_LT_LMUL_EQ) + +lemma REAL_LT_RMUL_EQ: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. + real_lt (real_of_num (0::nat)) z --> + real_lt (real_mul x z) (real_mul y z) = real_lt x y" + by (import hollight REAL_LT_RMUL_EQ) + +lemma REAL_LT_RMUL_IMP: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. + real_lt x y & real_lt (real_of_num (0::nat)) z --> + real_lt (real_mul x z) (real_mul y z)" + by (import hollight REAL_LT_RMUL_IMP) + +lemma REAL_LT_LMUL_IMP: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. + real_lt y z & real_lt (real_of_num (0::nat)) x --> + real_lt (real_mul x y) (real_mul x z)" + by (import hollight REAL_LT_LMUL_IMP) + +lemma REAL_LINV_UNIQ: "ALL (x::hollight.real) y::hollight.real. + real_mul x y = real_of_num (NUMERAL_BIT1 (0::nat)) --> x = real_inv y" + by (import hollight REAL_LINV_UNIQ) + +lemma REAL_RINV_UNIQ: "ALL (x::hollight.real) y::hollight.real. + real_mul x y = real_of_num (NUMERAL_BIT1 (0::nat)) --> y = real_inv x" + by (import hollight REAL_RINV_UNIQ) + +lemma REAL_NEG_INV: "ALL x::hollight.real. + x ~= real_of_num (0::nat) --> + real_neg (real_inv x) = real_inv (real_neg x)" + by (import hollight REAL_NEG_INV) + +lemma REAL_INV_1OVER: "ALL x::hollight.real. + real_inv x = real_div (real_of_num (NUMERAL_BIT1 (0::nat))) x" + by (import hollight REAL_INV_1OVER) + +lemma REAL: "ALL x::nat. + real_of_num (Suc x) = + real_add (real_of_num x) (real_of_num (NUMERAL_BIT1 (0::nat)))" + by (import hollight REAL) + +lemma REAL_POS: "ALL n::nat. real_le (real_of_num (0::nat)) (real_of_num n)" + by (import hollight REAL_POS) + +lemma REAL_LE: "ALL (m::nat) n::nat. real_le (real_of_num m) (real_of_num n) = <= m n" + by (import hollight REAL_LE) + +lemma REAL_LT: "ALL (m::nat) n::nat. real_lt (real_of_num m) (real_of_num n) = < m n" + by (import hollight REAL_LT) + +lemma th: "((m::nat) = (n::nat)) = (<= m n & <= n m)" + by (import hollight th) + +lemma REAL_INJ: "ALL (m::nat) n::nat. (real_of_num m = real_of_num n) = (m = n)" + by (import hollight REAL_INJ) + +lemma REAL_ADD: "ALL (m::nat) n::nat. + real_add (real_of_num m) (real_of_num n) = real_of_num (m + n)" + by (import hollight REAL_ADD) + +lemma REAL_MUL: "ALL (m::nat) n::nat. + real_mul (real_of_num m) (real_of_num n) = real_of_num (m * n)" + by (import hollight REAL_MUL) + +lemma REAL_INV1: "real_inv (real_of_num (NUMERAL_BIT1 (0::nat))) = +real_of_num (NUMERAL_BIT1 (0::nat))" + by (import hollight REAL_INV1) + +lemma REAL_DIV_LZERO: "ALL x::hollight.real. + real_div (real_of_num (0::nat)) x = real_of_num (0::nat)" + by (import hollight REAL_DIV_LZERO) + +lemma REAL_LT_NZ: "ALL n::nat. + (real_of_num n ~= real_of_num (0::nat)) = + real_lt (real_of_num (0::nat)) (real_of_num n)" + by (import hollight REAL_LT_NZ) + +lemma REAL_NZ_IMP_LT: "ALL n::nat. n ~= (0::nat) --> real_lt (real_of_num (0::nat)) (real_of_num n)" + by (import hollight REAL_NZ_IMP_LT) + +lemma REAL_LT_RDIV_0: "ALL (y::hollight.real) z::hollight.real. + real_lt (real_of_num (0::nat)) z --> + real_lt (real_of_num (0::nat)) (real_div y z) = + real_lt (real_of_num (0::nat)) y" + by (import hollight REAL_LT_RDIV_0) + +lemma REAL_LT_RDIV: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. + real_lt (real_of_num (0::nat)) z --> + real_lt (real_div x z) (real_div y z) = real_lt x y" + by (import hollight REAL_LT_RDIV) + +lemma REAL_LT_FRACTION_0: "ALL (n::nat) d::hollight.real. + n ~= (0::nat) --> + real_lt (real_of_num (0::nat)) (real_div d (real_of_num n)) = + real_lt (real_of_num (0::nat)) d" + by (import hollight REAL_LT_FRACTION_0) + +lemma REAL_LT_MULTIPLE: "ALL (x::nat) xa::hollight.real. + < (NUMERAL_BIT1 (0::nat)) x --> + real_lt xa (real_mul (real_of_num x) xa) = + real_lt (real_of_num (0::nat)) xa" + by (import hollight REAL_LT_MULTIPLE) + +lemma REAL_LT_FRACTION: "ALL (n::nat) d::hollight.real. + < (NUMERAL_BIT1 (0::nat)) n --> + real_lt (real_div d (real_of_num n)) d = real_lt (real_of_num (0::nat)) d" + by (import hollight REAL_LT_FRACTION) + +lemma REAL_LT_HALF1: "ALL d::hollight.real. + real_lt (real_of_num (0::nat)) + (real_div d (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) = + real_lt (real_of_num (0::nat)) d" + by (import hollight REAL_LT_HALF1) + +lemma REAL_LT_HALF2: "ALL d::hollight.real. + real_lt (real_div d (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) + d = + real_lt (real_of_num (0::nat)) d" + by (import hollight REAL_LT_HALF2) + +lemma REAL_DOUBLE: "ALL x::hollight.real. + real_add x x = + real_mul (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) x" + by (import hollight REAL_DOUBLE) + +lemma REAL_HALF_DOUBLE: "ALL x::hollight.real. + real_add + (real_div x (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) + (real_div x (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) = + x" + by (import hollight REAL_HALF_DOUBLE) + +lemma REAL_SUB_SUB: "ALL (x::hollight.real) y::hollight.real. + 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)" + 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)" + 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" + 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" + 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)" + 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" + 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" + by (import hollight REAL_LE_NEG) + +lemma REAL_SUB_LZERO: "ALL x::hollight.real. real_sub (real_of_num (0::nat)) 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::nat)) = x" + by (import hollight REAL_SUB_RZERO) + +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)" + by (import hollight REAL_LTE_ADD2) + +lemma REAL_LTE_ADD: "ALL (x::hollight.real) y::hollight.real. + real_lt (real_of_num (0::nat)) x & real_le (real_of_num (0::nat)) y --> + real_lt (real_of_num (0::nat)) (real_add x y)" + by (import hollight REAL_LTE_ADD) + +lemma REAL_LT_MUL2_ALT: "ALL (x1::hollight.real) (x2::hollight.real) (y1::hollight.real) + y2::hollight.real. + real_le (real_of_num (0::nat)) x1 & + real_le (real_of_num (0::nat)) y1 & real_lt x1 x2 & real_lt y1 y2 --> + real_lt (real_mul x1 y1) (real_mul x2 y2)" + by (import hollight REAL_LT_MUL2_ALT) + +lemma REAL_SUB_LNEG: "ALL (x::hollight.real) y::hollight.real. + 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" + 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" + 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" + by (import hollight REAL_SUB_TRIANGLE) + +lemma REAL_INV_MUL_WEAK: "ALL (x::hollight.real) xa::hollight.real. + x ~= real_of_num (0::nat) & xa ~= real_of_num (0::nat) --> + real_inv (real_mul x xa) = real_mul (real_inv x) (real_inv xa)" + by (import hollight REAL_INV_MUL_WEAK) + +lemma REAL_LE_LMUL_LOCAL: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. + real_lt (real_of_num (0::nat)) x --> + real_le (real_mul x y) (real_mul x z) = real_le y z" + by (import hollight REAL_LE_LMUL_LOCAL) + +lemma REAL_LE_RMUL_EQ: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. + real_lt (real_of_num (0::nat)) z --> + real_le (real_mul x z) (real_mul y z) = real_le x y" + by (import hollight REAL_LE_RMUL_EQ) + +lemma REAL_SUB_INV2: "ALL (x::hollight.real) y::hollight.real. + x ~= real_of_num (0::nat) & y ~= real_of_num (0::nat) --> + real_sub (real_inv x) (real_inv y) = + real_div (real_sub y x) (real_mul x y)" + by (import hollight REAL_SUB_INV2) + +lemma REAL_SUB_SUB2: "ALL (x::hollight.real) y::hollight.real. real_sub x (real_sub x y) = y" + by (import hollight REAL_SUB_SUB2) + +lemma REAL_MEAN: "ALL (x::hollight.real) y::hollight.real. + real_lt x y --> (EX z::hollight.real. real_lt x z & real_lt z y)" + by (import hollight REAL_MEAN) + +lemma REAL_EQ_LMUL2: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. + x ~= real_of_num (0::nat) --> (y = z) = (real_mul x y = real_mul x z)" + by (import hollight REAL_EQ_LMUL2) + +lemma REAL_LE_MUL2V: "ALL (x1::hollight.real) (x2::hollight.real) (y1::hollight.real) + y2::hollight.real. + real_le (real_of_num (0::nat)) x1 & + real_le (real_of_num (0::nat)) y1 & real_le x1 x2 & real_le y1 y2 --> + real_le (real_mul x1 y1) (real_mul x2 y2)" + by (import hollight REAL_LE_MUL2V) + +lemma REAL_LE_LDIV: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. + real_lt (real_of_num (0::nat)) x & real_le y (real_mul z x) --> + real_le (real_div y x) z" + by (import hollight REAL_LE_LDIV) + +lemma REAL_LE_RDIV: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. + real_lt (real_of_num (0::nat)) x & real_le (real_mul y x) z --> + real_le y (real_div z x)" + by (import hollight REAL_LE_RDIV) + +lemma REAL_LT_1: "ALL (x::hollight.real) y::hollight.real. + real_le (real_of_num (0::nat)) x & real_lt x y --> + real_lt (real_div x y) (real_of_num (NUMERAL_BIT1 (0::nat)))" + by (import hollight REAL_LT_1) + +lemma REAL_LE_LMUL_IMP: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. + real_le (real_of_num (0::nat)) x & real_le y z --> + real_le (real_mul x y) (real_mul x z)" + by (import hollight REAL_LE_LMUL_IMP) + +lemma REAL_LE_RMUL_IMP: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real. + real_le (real_of_num (0::nat)) x & real_le xa xb --> + real_le (real_mul xa x) (real_mul xb x)" + by (import hollight REAL_LE_RMUL_IMP) + +lemma REAL_INV_LT1: "ALL x::hollight.real. + real_lt (real_of_num (0::nat)) x & + real_lt x (real_of_num (NUMERAL_BIT1 (0::nat))) --> + real_lt (real_of_num (NUMERAL_BIT1 (0::nat))) (real_inv x)" + by (import hollight REAL_INV_LT1) + +lemma REAL_POS_NZ: "ALL x::hollight.real. + real_lt (real_of_num (0::nat)) x --> x ~= real_of_num (0::nat)" + by (import hollight REAL_POS_NZ) + +lemma REAL_EQ_RMUL_IMP: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. + z ~= real_of_num (0::nat) & real_mul x z = real_mul y z --> x = y" + by (import hollight REAL_EQ_RMUL_IMP) + +lemma REAL_EQ_LMUL_IMP: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real. + x ~= real_of_num (0::nat) & real_mul x xa = real_mul x xb --> xa = xb" + by (import hollight REAL_EQ_LMUL_IMP) + +lemma REAL_FACT_NZ: "ALL n::nat. real_of_num (FACT n) ~= real_of_num (0::nat)" + by (import hollight REAL_FACT_NZ) + +lemma REAL_POSSQ: "ALL x::hollight.real. + real_lt (real_of_num (0::nat)) (real_mul x x) = + (x ~= real_of_num (0::nat))" + by (import hollight REAL_POSSQ) + +lemma REAL_SUMSQ: "ALL (x::hollight.real) y::hollight.real. + (real_add (real_mul x x) (real_mul y y) = real_of_num (0::nat)) = + (x = real_of_num (0::nat) & y = real_of_num (0::nat))" + by (import hollight REAL_SUMSQ) + +lemma REAL_EQ_NEG: "ALL (x::hollight.real) y::hollight.real. (real_neg x = real_neg y) = (x = y)" + by (import hollight REAL_EQ_NEG) + +lemma REAL_DIV_MUL2: "ALL (x::hollight.real) z::hollight.real. + x ~= real_of_num (0::nat) & z ~= real_of_num (0::nat) --> + (ALL y::hollight.real. + real_div y z = real_div (real_mul x y) (real_mul x z))" + by (import hollight REAL_DIV_MUL2) + +lemma REAL_MIDDLE1: "ALL (a::hollight.real) b::hollight.real. + real_le a b --> + real_le a + (real_div (real_add a b) + (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))))" + by (import hollight REAL_MIDDLE1) + +lemma REAL_MIDDLE2: "ALL (a::hollight.real) b::hollight.real. + real_le a b --> + real_le + (real_div (real_add a b) + (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) + b" + by (import hollight REAL_MIDDLE2) + +lemma ABS_ZERO: "ALL x::hollight.real. + (real_abs x = real_of_num (0::nat)) = (x = real_of_num (0::nat))" + by (import hollight ABS_ZERO) + +lemma ABS_0: "real_abs (real_of_num (0::nat)) = real_of_num (0::nat)" + by (import hollight ABS_0) + +lemma ABS_1: "real_abs (real_of_num (NUMERAL_BIT1 (0::nat))) = +real_of_num (NUMERAL_BIT1 (0::nat))" + by (import hollight ABS_1) + +lemma ABS_NEG: "ALL x::hollight.real. real_abs (real_neg x) = real_abs x" + by (import hollight ABS_NEG) + +lemma 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))" + by (import hollight ABS_TRIANGLE) + +lemma ABS_POS: "ALL x::hollight.real. real_le (real_of_num (0::nat)) (real_abs x)" + by (import hollight ABS_POS) + +lemma ABS_MUL: "ALL (x::hollight.real) y::hollight.real. + real_abs (real_mul x y) = real_mul (real_abs x) (real_abs y)" + by (import hollight ABS_MUL) + +lemma ABS_LT_MUL2: "ALL (w::hollight.real) (x::hollight.real) (y::hollight.real) + z::hollight.real. + real_lt (real_abs w) y & real_lt (real_abs x) z --> + real_lt (real_abs (real_mul w x)) (real_mul y z)" + by (import hollight ABS_LT_MUL2) + +lemma ABS_SUB: "ALL (x::hollight.real) y::hollight.real. + real_abs (real_sub x y) = real_abs (real_sub y x)" + by (import hollight ABS_SUB) + +lemma ABS_NZ: "ALL x::hollight.real. + (x ~= real_of_num (0::nat)) = real_lt (real_of_num (0::nat)) (real_abs x)" + by (import hollight ABS_NZ) + +lemma ABS_INV: "ALL x::hollight.real. + x ~= real_of_num (0::nat) --> + real_abs (real_inv x) = real_inv (real_abs x)" + by (import hollight ABS_INV) + +lemma ABS_ABS: "ALL x::hollight.real. real_abs (real_abs x) = real_abs x" + by (import hollight ABS_ABS) + +lemma ABS_LE: "ALL x::hollight.real. real_le x (real_abs x)" + by (import hollight ABS_LE) + +lemma ABS_REFL: "ALL x::hollight.real. (real_abs x = x) = real_le (real_of_num (0::nat)) x" + by (import hollight ABS_REFL) + +lemma ABS_N: "ALL n::nat. real_abs (real_of_num n) = real_of_num n" + by (import hollight ABS_N) + +lemma ABS_BETWEEN: "ALL (x::hollight.real) (y::hollight.real) d::hollight.real. + (real_lt (real_of_num (0::nat)) 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 ABS_BETWEEN) + +lemma 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)" + by (import hollight ABS_BOUND) + +lemma 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::nat)" + by (import hollight ABS_STILLNZ) + +lemma ABS_CASES: "ALL x::hollight.real. + x = real_of_num (0::nat) | real_lt (real_of_num (0::nat)) (real_abs x)" + by (import hollight ABS_CASES) + +lemma 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" + by (import hollight ABS_BETWEEN1) + +lemma ABS_SIGN: "ALL (x::hollight.real) y::hollight.real. + real_lt (real_abs (real_sub x y)) y --> real_lt (real_of_num (0::nat)) x" + by (import hollight ABS_SIGN) + +lemma 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::nat))" + by (import hollight ABS_SIGN2) + +lemma ABS_DIV: "ALL y::hollight.real. + y ~= real_of_num (0::nat) --> + (ALL x::hollight.real. + real_abs (real_div x y) = real_div (real_abs x) (real_abs y))" + by (import hollight ABS_DIV) + +lemma 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)" + by (import hollight 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))" + by (import hollight REAL_SUB_ABS) + +lemma 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))" + by (import hollight ABS_SUB_ABS) + +lemma ABS_BETWEEN2: "ALL (x0::hollight.real) (x::hollight.real) (y0::hollight.real) + y::hollight.real. + real_lt x0 y0 & + real_lt (real_abs (real_sub x x0)) + (real_div (real_sub y0 x0) + (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) & + real_lt (real_abs (real_sub y y0)) + (real_div (real_sub y0 x0) + (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) --> + real_lt x y" + by (import hollight ABS_BETWEEN2) + +lemma 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)" + by (import hollight ABS_BOUNDS) + +lemma POW_0: "ALL n::nat. real_pow (real_of_num (0::nat)) (Suc n) = real_of_num (0::nat)" + by (import hollight POW_0) + +lemma POW_NZ: "ALL (c::hollight.real) n::nat. + c ~= real_of_num (0::nat) --> real_pow c n ~= real_of_num (0::nat)" + by (import hollight POW_NZ) + +lemma POW_INV: "ALL (c::hollight.real) x::nat. + c ~= real_of_num (0::nat) --> + real_inv (real_pow c x) = real_pow (real_inv c) x" + by (import hollight POW_INV) + +lemma POW_ABS: "ALL (c::hollight.real) n::nat. + real_pow (real_abs c) n = real_abs (real_pow c n)" + by (import hollight POW_ABS) + +lemma POW_PLUS1: "ALL (e::hollight.real) x::nat. + real_lt (real_of_num (0::nat)) e --> + real_le + (real_add (real_of_num (NUMERAL_BIT1 (0::nat))) + (real_mul (real_of_num x) e)) + (real_pow (real_add (real_of_num (NUMERAL_BIT1 (0::nat))) e) x)" + by (import hollight POW_PLUS1) + +lemma POW_ADD: "ALL (c::hollight.real) (m::nat) n::nat. + real_pow c (m + n) = real_mul (real_pow c m) (real_pow c n)" + by (import hollight POW_ADD) + +lemma POW_1: "ALL x::hollight.real. real_pow x (NUMERAL_BIT1 (0::nat)) = x" + by (import hollight POW_1) + +lemma POW_2: "ALL x::hollight.real. + real_pow x (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))) = real_mul x x" + by (import hollight POW_2) + +lemma POW_POS: "ALL (x::hollight.real) xa::nat. + real_le (real_of_num (0::nat)) x --> + real_le (real_of_num (0::nat)) (real_pow x xa)" + by (import hollight POW_POS) + +lemma POW_LE: "ALL (n::nat) (x::hollight.real) y::hollight.real. + real_le (real_of_num (0::nat)) x & real_le x y --> + real_le (real_pow x n) (real_pow y n)" + by (import hollight POW_LE) + +lemma POW_M1: "ALL n::nat. + real_abs (real_pow (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) n) = + real_of_num (NUMERAL_BIT1 (0::nat))" + by (import hollight POW_M1) + +lemma POW_MUL: "ALL (n::nat) (x::hollight.real) y::hollight.real. + real_pow (real_mul x y) n = real_mul (real_pow x n) (real_pow y n)" + by (import hollight POW_MUL) + +lemma REAL_LE_SQUARE_POW: "ALL x::hollight.real. + real_le (real_of_num (0::nat)) + (real_pow x (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))" + by (import hollight REAL_LE_SQUARE_POW) + +lemma ABS_POW2: "ALL x::hollight.real. + real_abs (real_pow x (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) = + real_pow x (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))" + by (import hollight ABS_POW2) + +lemma REAL_LE1_POW2: "ALL x::hollight.real. + real_le (real_of_num (NUMERAL_BIT1 (0::nat))) x --> + real_le (real_of_num (NUMERAL_BIT1 (0::nat))) + (real_pow x (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))" + by (import hollight REAL_LE1_POW2) + +lemma REAL_LT1_POW2: "ALL x::hollight.real. + real_lt (real_of_num (NUMERAL_BIT1 (0::nat))) x --> + real_lt (real_of_num (NUMERAL_BIT1 (0::nat))) + (real_pow x (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))" + by (import hollight REAL_LT1_POW2) + +lemma POW_POS_LT: "ALL (x::hollight.real) n::nat. + real_lt (real_of_num (0::nat)) x --> + real_lt (real_of_num (0::nat)) (real_pow x (Suc n))" + by (import hollight POW_POS_LT) + +lemma POW_2_LE1: "ALL n::nat. + real_le (real_of_num (NUMERAL_BIT1 (0::nat))) + (real_pow (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) n)" + by (import hollight POW_2_LE1) + +lemma POW_2_LT: "ALL n::nat. + real_lt (real_of_num n) + (real_pow (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) n)" + by (import hollight POW_2_LT) + +lemma POW_MINUS1: "ALL n::nat. + real_pow (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) + (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)) * n) = + real_of_num (NUMERAL_BIT1 (0::nat))" + by (import hollight POW_MINUS1) + +lemma REAL_SUP_EXISTS: "ALL P::hollight.real => bool. + Ex P & + (EX z::hollight.real. ALL x::hollight.real. P x --> real_lt x z) --> + (EX s::hollight.real. + ALL y::hollight.real. + (EX x::hollight.real. P x & real_lt y x) = real_lt y s)" + by (import hollight REAL_SUP_EXISTS) + +constdefs + sup :: "(hollight.real => bool) => hollight.real" + "sup == +%u::hollight.real => bool. + SOME a::hollight.real. + (ALL x::hollight.real. IN x u --> real_le x a) & + (ALL b::hollight.real. + (ALL x::hollight.real. IN x u --> real_le x b) --> real_le a b)" + +lemma DEF_sup: "sup = +(%u::hollight.real => bool. + SOME a::hollight.real. + (ALL x::hollight.real. IN x u --> real_le x a) & + (ALL b::hollight.real. + (ALL x::hollight.real. IN x u --> real_le x b) --> real_le a b))" + by (import hollight DEF_sup) + +lemma sup: "sup (P::hollight.real => bool) = +(SOME s::hollight.real. + ALL y::hollight.real. + (EX x::hollight.real. P x & real_lt y x) = real_lt y s)" + by (import hollight sup) + +lemma REAL_SUP: "ALL P::hollight.real => bool. + Ex P & + (EX z::hollight.real. ALL x::hollight.real. P x --> real_lt x z) --> + (ALL y::hollight.real. + (EX x::hollight.real. P x & real_lt y x) = real_lt y (sup P))" + by (import hollight REAL_SUP) + +lemma REAL_SUP_UBOUND: "ALL P::hollight.real => bool. + Ex P & + (EX z::hollight.real. ALL x::hollight.real. P x --> real_lt x z) --> + (ALL y::hollight.real. P y --> real_le y (sup P))" + by (import hollight REAL_SUP_UBOUND) + +lemma SETOK_LE_LT: "ALL P::hollight.real => bool. + (Ex P & + (EX z::hollight.real. ALL x::hollight.real. P x --> real_le x z)) = + (Ex P & + (EX x::hollight.real. ALL xa::hollight.real. P xa --> real_lt xa x))" + by (import hollight SETOK_LE_LT) + +lemma REAL_SUP_LE: "ALL P::hollight.real => bool. + Ex P & + (EX z::hollight.real. ALL x::hollight.real. P x --> real_le x z) --> + (ALL y::hollight.real. + (EX x::hollight.real. P x & real_lt y x) = real_lt y (sup P))" + by (import hollight REAL_SUP_LE) + +lemma REAL_SUP_UBOUND_LE: "ALL P::hollight.real => bool. + Ex P & + (EX z::hollight.real. ALL x::hollight.real. P x --> real_le x z) --> + (ALL y::hollight.real. P y --> real_le y (sup P))" + by (import hollight REAL_SUP_UBOUND_LE) + +lemma REAL_ARCH_SIMPLE: "ALL x::hollight.real. EX n::nat. real_le x (real_of_num n)" + by (import hollight REAL_ARCH_SIMPLE) + +lemma REAL_ARCH: "ALL x::hollight.real. + real_lt (real_of_num (0::nat)) x --> + (ALL y::hollight.real. EX n::nat. real_lt y (real_mul (real_of_num n) x))" + by (import hollight REAL_ARCH) + +lemma REAL_ARCH_LEAST: "ALL y::hollight.real. + real_lt (real_of_num (0::nat)) y --> + (ALL x::hollight.real. + real_le (real_of_num (0::nat)) x --> + (EX n::nat. + real_le (real_mul (real_of_num n) y) x & + real_lt x (real_mul (real_of_num (Suc n)) y)))" + by (import hollight REAL_ARCH_LEAST) + +lemma sum_EXISTS: "EX x::nat * nat => (nat => hollight.real) => hollight.real. + (ALL (f::nat => hollight.real) n::nat. + x (n, 0::nat) f = real_of_num (0::nat)) & + (ALL (f::nat => hollight.real) (m::nat) n::nat. + x (n, Suc m) f = real_add (x (n, m) f) (f (n + m)))" + by (import hollight sum_EXISTS) + +constdefs + psum :: "nat * nat => (nat => hollight.real) => hollight.real" + "psum == +SOME sum::nat * nat => (nat => hollight.real) => hollight.real. + (ALL (f::nat => hollight.real) n::nat. + sum (n, 0::nat) f = real_of_num (0::nat)) & + (ALL (f::nat => hollight.real) (m::nat) n::nat. + sum (n, Suc m) f = real_add (sum (n, m) f) (f (n + m)))" + +lemma DEF_psum: "psum = +(SOME sum::nat * nat => (nat => hollight.real) => hollight.real. + (ALL (f::nat => hollight.real) n::nat. + sum (n, 0::nat) f = real_of_num (0::nat)) & + (ALL (f::nat => hollight.real) (m::nat) n::nat. + sum (n, Suc m) f = real_add (sum (n, m) f) (f (n + m))))" + by (import hollight DEF_psum) + +lemma sum: "psum (n::nat, 0::nat) (f::nat => hollight.real) = real_of_num (0::nat) & +psum (n, Suc (m::nat)) f = real_add (psum (n, m) f) (f (n + m))" + by (import hollight sum) + +lemma PSUM_SUM: "ALL (f::nat => hollight.real) (m::nat) n::nat. + psum (m, n) f = + hollight.sum + (GSPEC (%u::nat. EX i::nat. SETSPEC u (<= m i & < i (m + n)) i)) f" + by (import hollight PSUM_SUM) + +lemma PSUM_SUM_NUMSEG: "ALL (f::nat => hollight.real) (m::nat) n::nat. + ~ (m = (0::nat) & n = (0::nat)) --> + psum (m, n) f = hollight.sum (dotdot m (m + n - NUMERAL_BIT1 (0::nat))) f" + by (import hollight PSUM_SUM_NUMSEG) + +lemma SUM_TWO: "ALL (f::nat => hollight.real) (n::nat) p::nat. + real_add (psum (0::nat, n) f) (psum (n, p) f) = psum (0::nat, n + p) f" + by (import hollight SUM_TWO) + +lemma SUM_DIFF: "ALL (f::nat => hollight.real) (m::nat) n::nat. + psum (m, n) f = real_sub (psum (0::nat, m + n) f) (psum (0::nat, m) f)" + by (import hollight SUM_DIFF) + +lemma ABS_SUM: "ALL (f::nat => hollight.real) (m::nat) n::nat. + real_le (real_abs (psum (m, n) f)) + (psum (m, n) (%n::nat. real_abs (f n)))" + by (import hollight ABS_SUM) + +lemma SUM_LE: "ALL (f::nat => hollight.real) (g::nat => hollight.real) (m::nat) n::nat. + (ALL r::nat. <= m r & < r (n + m) --> real_le (f r) (g r)) --> + real_le (psum (m, n) f) (psum (m, n) g)" + by (import hollight SUM_LE) + +lemma SUM_EQ: "ALL (f::nat => hollight.real) (g::nat => hollight.real) (m::nat) n::nat. + (ALL r::nat. <= m r & < r (n + m) --> f r = g r) --> + psum (m, n) f = psum (m, n) g" + by (import hollight SUM_EQ) + +lemma SUM_POS: "ALL f::nat => hollight.real. + (ALL n::nat. real_le (real_of_num (0::nat)) (f n)) --> + (ALL (m::nat) n::nat. real_le (real_of_num (0::nat)) (psum (m, n) f))" + by (import hollight SUM_POS) + +lemma SUM_POS_GEN: "ALL (f::nat => hollight.real) (m::nat) n::nat. + (ALL n::nat. <= m n --> real_le (real_of_num (0::nat)) (f n)) --> + real_le (real_of_num (0::nat)) (psum (m, n) f)" + by (import hollight SUM_POS_GEN) + +lemma SUM_ABS: "ALL (f::nat => hollight.real) (m::nat) x::nat. + real_abs (psum (m, x) (%m::nat. real_abs (f m))) = + psum (m, x) (%m::nat. real_abs (f m))" + by (import hollight SUM_ABS) + +lemma SUM_ABS_LE: "ALL (f::nat => hollight.real) (m::nat) n::nat. + real_le (real_abs (psum (m, n) f)) + (psum (m, n) (%n::nat. real_abs (f n)))" + by (import hollight SUM_ABS_LE) + +lemma SUM_ZERO: "ALL (f::nat => hollight.real) N::nat. + (ALL n::nat. >= n N --> f n = real_of_num (0::nat)) --> + (ALL (m::nat) n::nat. >= m N --> psum (m, n) f = real_of_num (0::nat))" + by (import hollight SUM_ZERO) + +lemma SUM_ADD: "ALL (f::nat => hollight.real) (g::nat => hollight.real) (m::nat) n::nat. + psum (m, n) (%n::nat. real_add (f n) (g n)) = + real_add (psum (m, n) f) (psum (m, n) g)" + by (import hollight SUM_ADD) + +lemma SUM_CMUL: "ALL (f::nat => hollight.real) (c::hollight.real) (m::nat) n::nat. + psum (m, n) (%n::nat. real_mul c (f n)) = real_mul c (psum (m, n) f)" + by (import hollight SUM_CMUL) + +lemma SUM_NEG: "ALL (f::nat => hollight.real) (n::nat) d::nat. + psum (n, d) (%n::nat. real_neg (f n)) = real_neg (psum (n, d) f)" + by (import hollight SUM_NEG) + +lemma SUM_SUB: "ALL (f::nat => hollight.real) (g::nat => hollight.real) (m::nat) n::nat. + psum (m, n) (%x::nat. real_sub (f x) (g x)) = + real_sub (psum (m, n) f) (psum (m, n) g)" + by (import hollight SUM_SUB) + +lemma SUM_SUBST: "ALL (f::nat => hollight.real) (g::nat => hollight.real) (m::nat) n::nat. + (ALL p::nat. <= m p & < p (m + n) --> f p = g p) --> + psum (m, n) f = psum (m, n) g" + by (import hollight SUM_SUBST) + +lemma SUM_NSUB: "ALL (n::nat) (f::nat => hollight.real) c::hollight.real. + real_sub (psum (0::nat, n) f) (real_mul (real_of_num n) c) = + psum (0::nat, n) (%p::nat. real_sub (f p) c)" + by (import hollight SUM_NSUB) + +lemma SUM_BOUND: "ALL (f::nat => hollight.real) (K::hollight.real) (m::nat) n::nat. + (ALL p::nat. <= m p & < p (m + n) --> real_le (f p) K) --> + real_le (psum (m, n) f) (real_mul (real_of_num n) K)" + by (import hollight SUM_BOUND) + +lemma SUM_GROUP: "ALL (n::nat) (k::nat) f::nat => hollight.real. + psum (0::nat, n) (%m::nat. psum (m * k, k) f) = psum (0::nat, n * k) f" + by (import hollight SUM_GROUP) + +lemma SUM_1: "ALL (f::nat => hollight.real) n::nat. + psum (n, NUMERAL_BIT1 (0::nat)) f = f n" + by (import hollight SUM_1) + +lemma SUM_2: "ALL (f::nat => hollight.real) n::nat. + psum (n, NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))) f = + real_add (f n) (f (n + NUMERAL_BIT1 (0::nat)))" + by (import hollight SUM_2) + +lemma SUM_OFFSET: "ALL (f::nat => hollight.real) (n::nat) k::nat. + psum (0::nat, n) (%m::nat. f (m + k)) = + real_sub (psum (0::nat, n + k) f) (psum (0::nat, k) f)" + by (import hollight SUM_OFFSET) + +lemma SUM_REINDEX: "ALL (f::nat => hollight.real) (m::nat) (k::nat) n::nat. + psum (m + k, n) f = psum (m, n) (%r::nat. f (r + k))" + by (import hollight SUM_REINDEX) + +lemma SUM_0: "ALL (m::nat) n::nat. + psum (m, n) (%r::nat. real_of_num (0::nat)) = real_of_num (0::nat)" + by (import hollight SUM_0) + +lemma SUM_CANCEL: "ALL (f::nat => hollight.real) (n::nat) d::nat. + psum (n, d) (%n::nat. real_sub (f (Suc n)) (f n)) = + real_sub (f (n + d)) (f n)" + by (import hollight SUM_CANCEL) + +lemma SUM_HORNER: "ALL (f::nat => hollight.real) (n::nat) x::hollight.real. + psum (0::nat, Suc n) (%i::nat. real_mul (f i) (real_pow x i)) = + real_add (f (0::nat)) + (real_mul x + (psum (0::nat, n) (%i::nat. real_mul (f (Suc i)) (real_pow x i))))" + by (import hollight SUM_HORNER) + +lemma SUM_CONST: "ALL (c::hollight.real) n::nat. + psum (0::nat, n) (%m::nat. c) = real_mul (real_of_num n) c" + by (import hollight SUM_CONST) + +lemma SUM_SPLIT: "ALL (f::nat => hollight.real) (n::nat) p::nat. + real_add (psum (m::nat, n) f) (psum (m + n, p) f) = psum (m, n + p) f" + by (import hollight SUM_SPLIT) + +lemma SUM_SWAP: "ALL (f::nat => nat => hollight.real) (m1::nat) (n1::nat) (m2::nat) n2::nat. + psum (m1, n1) (%a::nat. psum (m2, n2) (f a)) = + psum (m2, n2) (%b::nat. psum (m1, n1) (%a::nat. f a b))" + by (import hollight SUM_SWAP) + +lemma SUM_EQ_0: "(ALL r::nat. + <= (m::nat) r & < r (m + (n::nat)) --> + (f::nat => hollight.real) r = real_of_num (0::nat)) --> +psum (m, n) f = real_of_num (0::nat)" + by (import hollight SUM_EQ_0) + +lemma SUM_MORETERMS_EQ: "ALL (m::nat) (n::nat) p::nat. + <= n p & + (ALL r::nat. + <= (m + n) r & < r (m + p) --> + (f::nat => hollight.real) r = real_of_num (0::nat)) --> + psum (m, p) f = psum (m, n) f" + by (import hollight SUM_MORETERMS_EQ) + +lemma SUM_DIFFERENCES_EQ: "ALL (x::nat) (xa::nat) xb::nat. + <= xa xb & + (ALL r::nat. + <= (x + xa) r & < r (x + xb) --> + (f::nat => hollight.real) r = (g::nat => hollight.real) r) --> + real_sub (psum (x, xb) f) (psum (x, xa) f) = + real_sub (psum (x, xb) g) (psum (x, xa) g)" + by (import hollight SUM_DIFFERENCES_EQ) + +constdefs + re_Union :: "(('A::type => bool) => bool) => 'A::type => bool" + "re_Union == +%(u::('A::type => bool) => bool) x::'A::type. + EX s::'A::type => bool. u s & s x" + +lemma DEF_re_Union: "re_Union = +(%(u::('A::type => bool) => bool) x::'A::type. + EX s::'A::type => bool. u s & s x)" + by (import hollight DEF_re_Union) + +constdefs + re_union :: "('A::type => bool) => ('A::type => bool) => 'A::type => bool" + "re_union == +%(u::'A::type => bool) (ua::'A::type => bool) x::'A::type. u x | ua x" + +lemma DEF_re_union: "re_union = +(%(u::'A::type => bool) (ua::'A::type => bool) x::'A::type. u x | ua x)" + by (import hollight DEF_re_union) + +constdefs + re_intersect :: "('A::type => bool) => ('A::type => bool) => 'A::type => bool" + "re_intersect == +%(u::'A::type => bool) (ua::'A::type => bool) x::'A::type. u x & ua x" + +lemma DEF_re_intersect: "re_intersect = +(%(u::'A::type => bool) (ua::'A::type => bool) x::'A::type. u x & ua x)" + by (import hollight DEF_re_intersect) + +constdefs + re_null :: "'A::type => bool" + "re_null == %x::'A::type. False" + +lemma DEF_re_null: "re_null = (%x::'A::type. False)" + by (import hollight DEF_re_null) + +constdefs + re_universe :: "'A::type => bool" + "re_universe == %x::'A::type. True" + +lemma DEF_re_universe: "re_universe = (%x::'A::type. True)" + by (import hollight DEF_re_universe) + +constdefs + re_subset :: "('A::type => bool) => ('A::type => bool) => bool" + "re_subset == +%(u::'A::type => bool) ua::'A::type => bool. ALL x::'A::type. u x --> ua x" + +lemma DEF_re_subset: "re_subset = +(%(u::'A::type => bool) ua::'A::type => bool. ALL x::'A::type. u x --> ua x)" + by (import hollight DEF_re_subset) + +constdefs + re_compl :: "('A::type => bool) => 'A::type => bool" + "re_compl == %(u::'A::type => bool) x::'A::type. ~ u x" + +lemma DEF_re_compl: "re_compl = (%(u::'A::type => bool) x::'A::type. ~ u x)" + by (import hollight DEF_re_compl) + +lemma SUBSETA_REFL: "ALL S::'A::type => bool. re_subset S S" + by (import hollight SUBSETA_REFL) + +lemma COMPL_MEM: "ALL (S::'A::type => bool) x::'A::type. S x = (~ re_compl S x)" + by (import hollight COMPL_MEM) + +lemma SUBSETA_ANTISYM: "ALL (P::'A::type => bool) Q::'A::type => bool. + (re_subset P Q & re_subset Q P) = (P = Q)" + by (import hollight SUBSETA_ANTISYM) + +lemma SUBSETA_TRANS: "ALL (P::'A::type => bool) (Q::'A::type => bool) R::'A::type => bool. + re_subset P Q & re_subset Q R --> re_subset P R" + by (import hollight SUBSETA_TRANS) + +constdefs + istopology :: "(('A::type => bool) => bool) => bool" + "istopology == +%u::('A::type => bool) => bool. + u re_null & + u re_universe & + (ALL (a::'A::type => bool) b::'A::type => bool. + u a & u b --> u (re_intersect a b)) & + (ALL P::('A::type => bool) => bool. re_subset P u --> u (re_Union P))" + +lemma DEF_istopology: "istopology = +(%u::('A::type => bool) => bool. + u re_null & + u re_universe & + (ALL (a::'A::type => bool) b::'A::type => bool. + u a & u b --> u (re_intersect a b)) & + (ALL P::('A::type => bool) => bool. re_subset P u --> u (re_Union P)))" + by (import hollight DEF_istopology) + +typedef (open) ('A) topology = "(Collect::((('A::type => bool) => bool) => bool) + => (('A::type => bool) => bool) set) + (istopology::(('A::type => bool) => bool) => bool)" morphisms "open" "topology" + apply (rule light_ex_imp_nonempty[where t="(Eps::((('A::type => bool) => bool) => bool) => ('A::type => bool) => bool) + (istopology::(('A::type => bool) => bool) => bool)"]) + by (import hollight TYDEF_topology) + +syntax + "open" :: _ + +syntax + topology :: _ + +lemmas "TYDEF_topology_@intern" = typedef_hol2hollight + [where a="a :: 'A::type topology" and r=r , + OF type_definition_topology] + +lemma TOPOLOGY: "ALL L::'A::type topology. + open L re_null & + open L re_universe & + (ALL (a::'A::type => bool) b::'A::type => bool. + open L a & open L b --> open L (re_intersect a b)) & + (ALL P::('A::type => bool) => bool. + re_subset P (open L) --> open L (re_Union P))" + by (import hollight TOPOLOGY) + +lemma TOPOLOGY_UNION: "ALL (x::'A::type topology) xa::('A::type => bool) => bool. + re_subset xa (open x) --> open x (re_Union xa)" + by (import hollight TOPOLOGY_UNION) + +constdefs + neigh :: "'A::type topology => ('A::type => bool) * 'A::type => bool" + "neigh == +%(u::'A::type topology) ua::('A::type => bool) * 'A::type. + EX P::'A::type => bool. open u P & re_subset P (fst ua) & P (snd ua)" + +lemma DEF_neigh: "neigh = +(%(u::'A::type topology) ua::('A::type => bool) * 'A::type. + EX P::'A::type => bool. open u P & re_subset P (fst ua) & P (snd ua))" + by (import hollight DEF_neigh) + +lemma OPEN_OWN_NEIGH: "ALL (S::'A::type => bool) (top::'A::type topology) x::'A::type. + open top S & S x --> neigh top (S, x)" + by (import hollight OPEN_OWN_NEIGH) + +lemma OPEN_UNOPEN: "ALL (S::'A::type => bool) top::'A::type topology. + open top S = + (re_Union (%P::'A::type => bool. open top P & re_subset P S) = S)" + by (import hollight OPEN_UNOPEN) + +lemma OPEN_SUBOPEN: "ALL (S::'A::type => bool) top::'A::type topology. + open top S = + (ALL x::'A::type. + S x --> + (EX xa::'A::type => bool. xa x & open top xa & re_subset xa S))" + by (import hollight OPEN_SUBOPEN) + +lemma OPEN_NEIGH: "ALL (S::'A::type => bool) top::'A::type topology. + open top S = + (ALL x::'A::type. + S x --> + (EX xa::'A::type => bool. neigh top (xa, x) & re_subset xa S))" + by (import hollight OPEN_NEIGH) + +constdefs + closed :: "'A::type topology => ('A::type => bool) => bool" + "closed == %(u::'A::type topology) ua::'A::type => bool. open u (re_compl ua)" + +lemma DEF_closed: "closed = +(%(u::'A::type topology) ua::'A::type => bool. open u (re_compl ua))" + by (import hollight DEF_closed) + +constdefs + limpt :: "'A::type topology => 'A::type => ('A::type => bool) => bool" + "limpt == +%(u::'A::type topology) (ua::'A::type) ub::'A::type => bool. + ALL N::'A::type => bool. + neigh u (N, ua) --> (EX y::'A::type. ua ~= y & ub y & N y)" + +lemma DEF_limpt: "limpt = +(%(u::'A::type topology) (ua::'A::type) ub::'A::type => bool. + ALL N::'A::type => bool. + neigh u (N, ua) --> (EX y::'A::type. ua ~= y & ub y & N y))" + by (import hollight DEF_limpt) + +lemma CLOSED_LIMPT: "ALL (top::'A::type topology) S::'A::type => bool. + closed top S = (ALL x::'A::type. limpt top x S --> S x)" + by (import hollight CLOSED_LIMPT) + +constdefs + ismet :: "('A::type * 'A::type => hollight.real) => bool" + "ismet == +%u::'A::type * 'A::type => hollight.real. + (ALL (x::'A::type) y::'A::type. + (u (x, y) = real_of_num (0::nat)) = (x = y)) & + (ALL (x::'A::type) (y::'A::type) z::'A::type. + real_le (u (y, z)) (real_add (u (x, y)) (u (x, z))))" + +lemma DEF_ismet: "ismet = +(%u::'A::type * 'A::type => hollight.real. + (ALL (x::'A::type) y::'A::type. + (u (x, y) = real_of_num (0::nat)) = (x = y)) & + (ALL (x::'A::type) (y::'A::type) z::'A::type. + real_le (u (y, z)) (real_add (u (x, y)) (u (x, z)))))" + by (import hollight DEF_ismet) + +typedef (open) ('A) metric = "(Collect::(('A::type * 'A::type => hollight.real) => bool) + => ('A::type * 'A::type => hollight.real) set) + (ismet::('A::type * 'A::type => hollight.real) => bool)" morphisms "mdist" "metric" + apply (rule light_ex_imp_nonempty[where t="(Eps::(('A::type * 'A::type => hollight.real) => bool) + => 'A::type * 'A::type => hollight.real) + (ismet::('A::type * 'A::type => hollight.real) => bool)"]) + by (import hollight TYDEF_metric) + +syntax + mdist :: _ + +syntax + metric :: _ + +lemmas "TYDEF_metric_@intern" = typedef_hol2hollight + [where a="a :: 'A::type metric" and r=r , + OF type_definition_metric] + +lemma METRIC_ISMET: "ALL m::'A::type metric. ismet (mdist m)" + by (import hollight METRIC_ISMET) + +lemma METRIC_ZERO: "ALL (m::'A::type metric) (x::'A::type) y::'A::type. + (mdist m (x, y) = real_of_num (0::nat)) = (x = y)" + by (import hollight METRIC_ZERO) + +lemma METRIC_SAME: "ALL (m::'A::type metric) x::'A::type. mdist m (x, x) = real_of_num (0::nat)" + by (import hollight METRIC_SAME) + +lemma METRIC_POS: "ALL (m::'A::type metric) (x::'A::type) y::'A::type. + real_le (real_of_num (0::nat)) (mdist m (x, y))" + by (import hollight METRIC_POS) + +lemma METRIC_SYM: "ALL (m::'A::type metric) (x::'A::type) y::'A::type. + mdist m (x, y) = mdist m (y, x)" + by (import hollight METRIC_SYM) + +lemma METRIC_TRIANGLE: "ALL (m::'A::type metric) (x::'A::type) (y::'A::type) z::'A::type. + real_le (mdist m (x, z)) (real_add (mdist m (x, y)) (mdist m (y, z)))" + by (import hollight METRIC_TRIANGLE) + +lemma METRIC_NZ: "ALL (m::'A::type metric) (x::'A::type) y::'A::type. + x ~= y --> real_lt (real_of_num (0::nat)) (mdist m (x, y))" + by (import hollight METRIC_NZ) + +constdefs + mtop :: "'A::type metric => 'A::type topology" + "mtop == +%u::'A::type metric. + topology + (%S::'A::type => bool. + ALL x::'A::type. + S x --> + (EX e::hollight.real. + real_lt (real_of_num (0::nat)) e & + (ALL y::'A::type. real_lt (mdist u (x, y)) e --> S y)))" + +lemma DEF_mtop: "mtop = +(%u::'A::type metric. + topology + (%S::'A::type => bool. + ALL x::'A::type. + S x --> + (EX e::hollight.real. + real_lt (real_of_num (0::nat)) e & + (ALL y::'A::type. real_lt (mdist u (x, y)) e --> S y))))" + by (import hollight DEF_mtop) + +lemma mtop_istopology: "ALL m::'A::type metric. + istopology + (%S::'A::type => bool. + ALL x::'A::type. + S x --> + (EX e::hollight.real. + real_lt (real_of_num (0::nat)) e & + (ALL y::'A::type. real_lt (mdist m (x, y)) e --> S y)))" + by (import hollight mtop_istopology) + +lemma MTOP_OPEN: "ALL m::'A::type metric. + open (mtop m) (S::'A::type => bool) = + (ALL x::'A::type. + S x --> + (EX e::hollight.real. + real_lt (real_of_num (0::nat)) e & + (ALL y::'A::type. real_lt (mdist m (x, y)) e --> S y)))" + by (import hollight MTOP_OPEN) + +constdefs + ball :: "'A::type metric => 'A::type * hollight.real => 'A::type => bool" + "ball == +%(u::'A::type metric) (ua::'A::type * hollight.real) y::'A::type. + real_lt (mdist u (fst ua, y)) (snd ua)" + +lemma DEF_ball: "ball = +(%(u::'A::type metric) (ua::'A::type * hollight.real) y::'A::type. + real_lt (mdist u (fst ua, y)) (snd ua))" + by (import hollight DEF_ball) + +lemma BALL_OPEN: "ALL (m::'A::type metric) (x::'A::type) e::hollight.real. + real_lt (real_of_num (0::nat)) e --> open (mtop m) (ball m (x, e))" + by (import hollight BALL_OPEN) + +lemma BALL_NEIGH: "ALL (m::'A::type metric) (x::'A::type) e::hollight.real. + real_lt (real_of_num (0::nat)) e --> neigh (mtop m) (ball m (x, e), x)" + by (import hollight BALL_NEIGH) + +lemma MTOP_LIMPT: "ALL (m::'A::type metric) (x::'A::type) S::'A::type => bool. + limpt (mtop m) x S = + (ALL e::hollight.real. + real_lt (real_of_num (0::nat)) e --> + (EX y::'A::type. x ~= y & S y & real_lt (mdist m (x, y)) e))" + by (import hollight MTOP_LIMPT) + +lemma ISMET_R1: "ismet + (GABS + (%f::hollight.real * hollight.real => hollight.real. + ALL (x::hollight.real) y::hollight.real. + GEQ (f (x, y)) (real_abs (real_sub y x))))" + by (import hollight ISMET_R1) + +constdefs + mr1 :: "hollight.real metric" + "mr1 == +metric + (GABS + (%f::hollight.real * hollight.real => hollight.real. + ALL (x::hollight.real) y::hollight.real. + GEQ (f (x, y)) (real_abs (real_sub y x))))" + +lemma DEF_mr1: "mr1 = +metric + (GABS + (%f::hollight.real * hollight.real => hollight.real. + ALL (x::hollight.real) y::hollight.real. + GEQ (f (x, y)) (real_abs (real_sub y x))))" + by (import hollight DEF_mr1) + +lemma MR1_DEF: "ALL (x::hollight.real) y::hollight.real. + mdist mr1 (x, y) = real_abs (real_sub y x)" + by (import hollight MR1_DEF) + +lemma MR1_ADD: "ALL (x::hollight.real) d::hollight.real. + mdist mr1 (x, real_add x d) = real_abs d" + by (import hollight MR1_ADD) + +lemma MR1_SUB: "ALL (x::hollight.real) d::hollight.real. + mdist mr1 (x, real_sub x d) = real_abs d" + by (import hollight MR1_SUB) + +lemma MR1_ADD_LE: "ALL (x::hollight.real) d::hollight.real. + real_le (real_of_num (0::nat)) d --> mdist mr1 (x, real_add x d) = d" + by (import hollight MR1_ADD_LE) + +lemma MR1_SUB_LE: "ALL (x::hollight.real) d::hollight.real. + real_le (real_of_num (0::nat)) d --> mdist mr1 (x, real_sub x d) = d" + by (import hollight MR1_SUB_LE) + +lemma MR1_ADD_LT: "ALL (x::hollight.real) d::hollight.real. + real_lt (real_of_num (0::nat)) d --> mdist mr1 (x, real_add x d) = d" + by (import hollight MR1_ADD_LT) + +lemma MR1_SUB_LT: "ALL (x::hollight.real) d::hollight.real. + real_lt (real_of_num (0::nat)) d --> mdist mr1 (x, real_sub x d) = d" + by (import hollight MR1_SUB_LT) + +lemma MR1_BETWEEN1: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. + real_lt x z & real_lt (mdist mr1 (x, y)) (real_sub z x) --> real_lt y z" + by (import hollight MR1_BETWEEN1) + +lemma MR1_LIMPT: "ALL x::hollight.real. limpt (mtop mr1) x re_universe" + by (import hollight MR1_LIMPT) + +constdefs + dorder :: "('A::type => 'A::type => bool) => bool" + "dorder == +%u::'A::type => 'A::type => bool. + ALL (x::'A::type) y::'A::type. + u x x & u y y --> + (EX z::'A::type. u z z & (ALL w::'A::type. u w z --> u w x & u w y))" + +lemma DEF_dorder: "dorder = +(%u::'A::type => 'A::type => bool. + ALL (x::'A::type) y::'A::type. + u x x & u y y --> + (EX z::'A::type. u z z & (ALL w::'A::type. u w z --> u w x & u w y)))" + by (import hollight DEF_dorder) + +constdefs + tends :: "('B::type => 'A::type) +=> 'A::type => 'A::type topology * ('B::type => 'B::type => bool) => bool" + "tends == +%(u::'B::type => 'A::type) (ua::'A::type) + ub::'A::type topology * ('B::type => 'B::type => bool). + ALL N::'A::type => bool. + neigh (fst ub) (N, ua) --> + (EX n::'B::type. + snd ub n n & (ALL m::'B::type. snd ub m n --> N (u m)))" + +lemma DEF_tends: "tends = +(%(u::'B::type => 'A::type) (ua::'A::type) + ub::'A::type topology * ('B::type => 'B::type => bool). + ALL N::'A::type => bool. + neigh (fst ub) (N, ua) --> + (EX n::'B::type. + snd ub n n & (ALL m::'B::type. snd ub m n --> N (u m))))" + by (import hollight DEF_tends) + +constdefs + bounded :: "'A::type metric * ('B::type => 'B::type => bool) +=> ('B::type => 'A::type) => bool" + "bounded == +%(u::'A::type metric * ('B::type => 'B::type => bool)) + ua::'B::type => 'A::type. + EX (k::hollight.real) (x::'A::type) N::'B::type. + snd u N N & + (ALL n::'B::type. snd u n N --> real_lt (mdist (fst u) (ua n, x)) k)" + +lemma DEF_bounded: "bounded = +(%(u::'A::type metric * ('B::type => 'B::type => bool)) + ua::'B::type => 'A::type. + EX (k::hollight.real) (x::'A::type) N::'B::type. + snd u N N & + (ALL n::'B::type. snd u n N --> real_lt (mdist (fst u) (ua n, x)) k))" + by (import hollight DEF_bounded) + +constdefs + tendsto :: "'A::type metric * 'A::type => 'A::type => 'A::type => bool" + "tendsto == +%(u::'A::type metric * 'A::type) (ua::'A::type) ub::'A::type. + real_lt (real_of_num (0::nat)) (mdist (fst u) (snd u, ua)) & + real_le (mdist (fst u) (snd u, ua)) (mdist (fst u) (snd u, ub))" + +lemma DEF_tendsto: "tendsto = +(%(u::'A::type metric * 'A::type) (ua::'A::type) ub::'A::type. + real_lt (real_of_num (0::nat)) (mdist (fst u) (snd u, ua)) & + real_le (mdist (fst u) (snd u, ua)) (mdist (fst u) (snd u, ub)))" + by (import hollight DEF_tendsto) + +lemma DORDER_LEMMA: "ALL g::'A::type => 'A::type => bool. + dorder g --> + (ALL (P::'A::type => bool) Q::'A::type => bool. + (EX n::'A::type. g n n & (ALL m::'A::type. g m n --> P m)) & + (EX n::'A::type. g n n & (ALL m::'A::type. g m n --> Q m)) --> + (EX n::'A::type. g n n & (ALL m::'A::type. g m n --> P m & Q m)))" + by (import hollight DORDER_LEMMA) + +lemma DORDER_NGE: "dorder >=" + by (import hollight DORDER_NGE) + +lemma DORDER_TENDSTO: "ALL (m::'A::type metric) x::'A::type. dorder (tendsto (m, x))" + by (import hollight DORDER_TENDSTO) + +lemma MTOP_TENDS: "ALL (d::'A::type metric) (g::'B::type => 'B::type => bool) + (x::'B::type => 'A::type) x0::'A::type. + tends x x0 (mtop d, g) = + (ALL e::hollight.real. + real_lt (real_of_num (0::nat)) e --> + (EX n::'B::type. + g n n & + (ALL m::'B::type. g m n --> real_lt (mdist d (x m, x0)) e)))" + by (import hollight MTOP_TENDS) + +lemma MTOP_TENDS_UNIQ: "ALL (g::'B::type => 'B::type => bool) d::'A::type metric. + dorder g --> + tends (x::'B::type => 'A::type) (x0::'A::type) (mtop d, g) & + tends x (x1::'A::type) (mtop d, g) --> + x0 = x1" + by (import hollight MTOP_TENDS_UNIQ) + +lemma SEQ_TENDS: "ALL (d::'A::type metric) (x::nat => 'A::type) x0::'A::type. + tends x x0 (mtop d, >=) = + (ALL xa::hollight.real. + real_lt (real_of_num (0::nat)) xa --> + (EX xb::nat. + ALL xc::nat. >= xc xb --> real_lt (mdist d (x xc, x0)) xa))" + by (import hollight SEQ_TENDS) + +lemma LIM_TENDS: "ALL (m1::'A::type metric) (m2::'B::type metric) (f::'A::type => 'B::type) + (x0::'A::type) y0::'B::type. + limpt (mtop m1) x0 re_universe --> + tends f y0 (mtop m2, tendsto (m1, x0)) = + (ALL e::hollight.real. + real_lt (real_of_num (0::nat)) e --> + (EX d::hollight.real. + real_lt (real_of_num (0::nat)) d & + (ALL x::'A::type. + real_lt (real_of_num (0::nat)) (mdist m1 (x, x0)) & + real_le (mdist m1 (x, x0)) d --> + real_lt (mdist m2 (f x, y0)) e)))" + by (import hollight LIM_TENDS) + +lemma LIM_TENDS2: "ALL (m1::'A::type metric) (m2::'B::type metric) (f::'A::type => 'B::type) + (x0::'A::type) y0::'B::type. + limpt (mtop m1) x0 re_universe --> + tends f y0 (mtop m2, tendsto (m1, x0)) = + (ALL e::hollight.real. + real_lt (real_of_num (0::nat)) e --> + (EX d::hollight.real. + real_lt (real_of_num (0::nat)) d & + (ALL x::'A::type. + real_lt (real_of_num (0::nat)) (mdist m1 (x, x0)) & + real_lt (mdist m1 (x, x0)) d --> + real_lt (mdist m2 (f x, y0)) e)))" + by (import hollight LIM_TENDS2) + +lemma MR1_BOUNDED: "ALL (g::'A::type => 'A::type => bool) f::'A::type => hollight.real. + bounded (mr1, g) f = + (EX (k::hollight.real) N::'A::type. + g N N & (ALL n::'A::type. g n N --> real_lt (real_abs (f n)) k))" + by (import hollight MR1_BOUNDED) + +lemma NET_NULL: "ALL (g::'A::type => 'A::type => bool) (x::'A::type => hollight.real) + x0::hollight.real. + tends x x0 (mtop mr1, g) = + tends (%n::'A::type. real_sub (x n) x0) (real_of_num (0::nat)) + (mtop mr1, g)" + by (import hollight NET_NULL) + +lemma NET_CONV_BOUNDED: "ALL (g::'A::type => 'A::type => bool) (x::'A::type => hollight.real) + x0::hollight.real. tends x x0 (mtop mr1, g) --> bounded (mr1, g) x" + by (import hollight NET_CONV_BOUNDED) + +lemma NET_CONV_NZ: "ALL (g::'A::type => 'A::type => bool) (x::'A::type => hollight.real) + x0::hollight.real. + tends x x0 (mtop mr1, g) & x0 ~= real_of_num (0::nat) --> + (EX N::'A::type. + g N N & (ALL n::'A::type. g n N --> x n ~= real_of_num (0::nat)))" + by (import hollight NET_CONV_NZ) + +lemma NET_CONV_IBOUNDED: "ALL (g::'A::type => 'A::type => bool) (x::'A::type => hollight.real) + x0::hollight.real. + tends x x0 (mtop mr1, g) & x0 ~= real_of_num (0::nat) --> + bounded (mr1, g) (%n::'A::type. real_inv (x n))" + by (import hollight NET_CONV_IBOUNDED) + +lemma NET_NULL_ADD: "ALL g::'A::type => 'A::type => bool. + dorder g --> + (ALL (x::'A::type => hollight.real) y::'A::type => hollight.real. + tends x (real_of_num (0::nat)) (mtop mr1, g) & + tends y (real_of_num (0::nat)) (mtop mr1, g) --> + tends (%n::'A::type. real_add (x n) (y n)) (real_of_num (0::nat)) + (mtop mr1, g))" + by (import hollight NET_NULL_ADD) + +lemma NET_NULL_MUL: "ALL g::'A::type => 'A::type => bool. + dorder g --> + (ALL (x::'A::type => hollight.real) y::'A::type => hollight.real. + bounded (mr1, g) x & tends y (real_of_num (0::nat)) (mtop mr1, g) --> + tends (%n::'A::type. real_mul (x n) (y n)) (real_of_num (0::nat)) + (mtop mr1, g))" + by (import hollight NET_NULL_MUL) + +lemma NET_NULL_CMUL: "ALL (g::'A::type => 'A::type => bool) (k::hollight.real) + x::'A::type => hollight.real. + tends x (real_of_num (0::nat)) (mtop mr1, g) --> + tends (%n::'A::type. real_mul k (x n)) (real_of_num (0::nat)) + (mtop mr1, g)" + by (import hollight NET_NULL_CMUL) + +lemma NET_ADD: "ALL (g::'A::type => 'A::type => bool) (x::'A::type => hollight.real) + (x0::hollight.real) (y::'A::type => hollight.real) y0::hollight.real. + dorder g --> + tends x x0 (mtop mr1, g) & tends y y0 (mtop mr1, g) --> + tends (%n::'A::type. real_add (x n) (y n)) (real_add x0 y0) (mtop mr1, g)" + by (import hollight NET_ADD) + +lemma NET_NEG: "ALL (g::'A::type => 'A::type => bool) (x::'A::type => hollight.real) + x0::hollight.real. + dorder g --> + tends x x0 (mtop mr1, g) = + tends (%n::'A::type. real_neg (x n)) (real_neg x0) (mtop mr1, g)" + by (import hollight NET_NEG) + +lemma NET_SUB: "ALL (g::'A::type => 'A::type => bool) (x::'A::type => hollight.real) + (x0::hollight.real) (y::'A::type => hollight.real) y0::hollight.real. + dorder g --> + tends x x0 (mtop mr1, g) & tends y y0 (mtop mr1, g) --> + tends (%xa::'A::type. real_sub (x xa) (y xa)) (real_sub x0 y0) + (mtop mr1, g)" + by (import hollight NET_SUB) + +lemma NET_MUL: "ALL (g::'A::type => 'A::type => bool) (x::'A::type => hollight.real) + (y::'A::type => hollight.real) (x0::hollight.real) y0::hollight.real. + dorder g --> + tends x x0 (mtop mr1, g) & tends y y0 (mtop mr1, g) --> + tends (%n::'A::type. real_mul (x n) (y n)) (real_mul x0 y0) (mtop mr1, g)" + by (import hollight NET_MUL) + +lemma NET_INV: "ALL (g::'A::type => 'A::type => bool) (x::'A::type => hollight.real) + x0::hollight.real. + dorder g --> + tends x x0 (mtop mr1, g) & x0 ~= real_of_num (0::nat) --> + tends (%n::'A::type. real_inv (x n)) (real_inv x0) (mtop mr1, g)" + by (import hollight NET_INV) + +lemma NET_DIV: "ALL (g::'A::type => 'A::type => bool) (x::'A::type => hollight.real) + (x0::hollight.real) (y::'A::type => hollight.real) y0::hollight.real. + dorder g --> + tends x x0 (mtop mr1, g) & + tends y y0 (mtop mr1, g) & y0 ~= real_of_num (0::nat) --> + tends (%xa::'A::type. real_div (x xa) (y xa)) (real_div x0 y0) + (mtop mr1, g)" + by (import hollight NET_DIV) + +lemma NET_ABS: "ALL (x::'A::type => hollight.real) x0::hollight.real. + tends x x0 (mtop mr1, g::'A::type => 'A::type => bool) --> + tends (%n::'A::type. real_abs (x n)) (real_abs x0) (mtop mr1, g)" + by (import hollight NET_ABS) + +lemma NET_SUM: "ALL g::'q_71813::type => 'q_71813::type => bool. + dorder g & + tends (%x::'q_71813::type. real_of_num (0::nat)) (real_of_num (0::nat)) + (mtop mr1, g) --> + (ALL (x::nat) n::nat. + (ALL r::nat. + <= x r & < r (x + n) --> + tends ((f::nat => 'q_71813::type => hollight.real) r) + ((l::nat => hollight.real) r) (mtop mr1, g)) --> + tends (%xa::'q_71813::type. psum (x, n) (%r::nat. f r xa)) + (psum (x, n) l) (mtop mr1, g))" + by (import hollight NET_SUM) + +lemma NET_LE: "ALL (g::'A::type => 'A::type => bool) (x::'A::type => hollight.real) + (x0::hollight.real) (y::'A::type => hollight.real) y0::hollight.real. + dorder g --> + tends x x0 (mtop mr1, g) & + tends y y0 (mtop mr1, g) & + (EX N::'A::type. + g N N & (ALL n::'A::type. g n N --> real_le (x n) (y n))) --> + real_le x0 y0" + by (import hollight NET_LE) + +constdefs + tends_num_real :: "(nat => hollight.real) => hollight.real => bool" + "tends_num_real == +%(u::nat => hollight.real) ua::hollight.real. tends u ua (mtop mr1, >=)" + +lemma DEF_tends_num_real: "tends_num_real = +(%(u::nat => hollight.real) ua::hollight.real. tends u ua (mtop mr1, >=))" + by (import hollight DEF_tends_num_real) + +lemma SEQ: "ALL (x::nat => hollight.real) x0::hollight.real. + tends_num_real x x0 = + (ALL e::hollight.real. + real_lt (real_of_num (0::nat)) e --> + (EX N::nat. + ALL n::nat. >= n N --> real_lt (real_abs (real_sub (x n) x0)) e))" + by (import hollight SEQ) + +lemma SEQ_CONST: "ALL k::hollight.real. tends_num_real (%x::nat. k) k" + by (import hollight SEQ_CONST) + +lemma SEQ_ADD: "ALL (x::nat => hollight.real) (x0::hollight.real) (y::nat => hollight.real) + y0::hollight.real. + tends_num_real x x0 & tends_num_real y y0 --> + tends_num_real (%n::nat. real_add (x n) (y n)) (real_add x0 y0)" + by (import hollight SEQ_ADD) + +lemma SEQ_MUL: "ALL (x::nat => hollight.real) (x0::hollight.real) (y::nat => hollight.real) + y0::hollight.real. + tends_num_real x x0 & tends_num_real y y0 --> + tends_num_real (%n::nat. real_mul (x n) (y n)) (real_mul x0 y0)" + by (import hollight SEQ_MUL) + +lemma SEQ_NEG: "ALL (x::nat => hollight.real) x0::hollight.real. + tends_num_real x x0 = + tends_num_real (%n::nat. real_neg (x n)) (real_neg x0)" + by (import hollight SEQ_NEG) + +lemma SEQ_INV: "ALL (x::nat => hollight.real) x0::hollight.real. + tends_num_real x x0 & x0 ~= real_of_num (0::nat) --> + tends_num_real (%n::nat. real_inv (x n)) (real_inv x0)" + by (import hollight SEQ_INV) + +lemma SEQ_SUB: "ALL (x::nat => hollight.real) (x0::hollight.real) (y::nat => hollight.real) + y0::hollight.real. + tends_num_real x x0 & tends_num_real y y0 --> + tends_num_real (%n::nat. real_sub (x n) (y n)) (real_sub x0 y0)" + by (import hollight SEQ_SUB) + +lemma SEQ_DIV: "ALL (x::nat => hollight.real) (x0::hollight.real) (y::nat => hollight.real) + y0::hollight.real. + tends_num_real x x0 & + tends_num_real y y0 & y0 ~= real_of_num (0::nat) --> + tends_num_real (%n::nat. real_div (x n) (y n)) (real_div x0 y0)" + by (import hollight SEQ_DIV) + +lemma SEQ_UNIQ: "ALL (x::nat => hollight.real) (x1::hollight.real) x2::hollight.real. + tends_num_real x x1 & tends_num_real x x2 --> x1 = x2" + by (import hollight SEQ_UNIQ) + +lemma SEQ_NULL: "ALL (s::nat => hollight.real) l::hollight.real. + tends_num_real s l = + tends_num_real (%n::nat. real_sub (s n) l) (real_of_num (0::nat))" + by (import hollight SEQ_NULL) + +lemma SEQ_SUM: "ALL (f::nat => nat => hollight.real) (l::nat => hollight.real) (m::nat) + n::nat. + (ALL x::nat. <= m x & < x (m + n) --> tends_num_real (f x) (l x)) --> + tends_num_real (%k::nat. psum (m, n) (%r::nat. f r k)) (psum (m, n) l)" + by (import hollight SEQ_SUM) + +lemma SEQ_TRANSFORM: "ALL (x::nat => hollight.real) (xa::nat => hollight.real) (xb::hollight.real) + xc::nat. + (ALL n::nat. <= xc n --> x n = xa n) & tends_num_real x xb --> + tends_num_real xa xb" + by (import hollight SEQ_TRANSFORM) + +constdefs + convergent :: "(nat => hollight.real) => bool" + "convergent == %u::nat => hollight.real. Ex (tends_num_real u)" + +lemma DEF_convergent: "convergent = (%u::nat => hollight.real. Ex (tends_num_real u))" + by (import hollight DEF_convergent) + +constdefs + cauchy :: "(nat => hollight.real) => bool" + "cauchy == +%u::nat => hollight.real. + ALL e::hollight.real. + real_lt (real_of_num (0::nat)) e --> + (EX N::nat. + ALL (m::nat) n::nat. + >= m N & >= n N --> + real_lt (real_abs (real_sub (u m) (u n))) e)" + +lemma DEF_cauchy: "cauchy = +(%u::nat => hollight.real. + ALL e::hollight.real. + real_lt (real_of_num (0::nat)) e --> + (EX N::nat. + ALL (m::nat) n::nat. + >= m N & >= n N --> + real_lt (real_abs (real_sub (u m) (u n))) e))" + by (import hollight DEF_cauchy) + +constdefs + lim :: "(nat => hollight.real) => hollight.real" + "lim == %u::nat => hollight.real. Eps (tends_num_real u)" + +lemma DEF_lim: "lim = (%u::nat => hollight.real. Eps (tends_num_real u))" + by (import hollight DEF_lim) + +lemma SEQ_LIM: "ALL f::nat => hollight.real. convergent f = tends_num_real f (lim f)" + by (import hollight SEQ_LIM) + +constdefs + subseq :: "(nat => nat) => bool" + "subseq == %u::nat => nat. ALL (m::nat) n::nat. < m n --> < (u m) (u n)" + +lemma DEF_subseq: "subseq = (%u::nat => nat. ALL (m::nat) n::nat. < m n --> < (u m) (u n))" + by (import hollight DEF_subseq) + +lemma SUBSEQ_SUC: "ALL f::nat => nat. subseq f = (ALL n::nat. < (f n) (f (Suc n)))" + by (import hollight SUBSEQ_SUC) + +consts + mono :: "(nat => hollight.real) => bool" + +defs + mono_def: "hollight.mono == +%u::nat => hollight.real. + (ALL (m::nat) n::nat. <= m n --> real_le (u m) (u n)) | + (ALL (m::nat) n::nat. <= m n --> hollight.real_ge (u m) (u n))" + +lemma DEF_mono: "hollight.mono = +(%u::nat => hollight.real. + (ALL (m::nat) n::nat. <= m n --> real_le (u m) (u n)) | + (ALL (m::nat) n::nat. <= m n --> hollight.real_ge (u m) (u n)))" + by (import hollight DEF_mono) + +lemma MONO_SUC: "ALL f::nat => hollight.real. + hollight.mono f = + ((ALL x::nat. hollight.real_ge (f (Suc x)) (f x)) | + (ALL n::nat. real_le (f (Suc n)) (f n)))" + by (import hollight MONO_SUC) + +lemma MAX_LEMMA: "ALL (s::nat => hollight.real) N::nat. + EX k::hollight.real. ALL n::nat. < n N --> real_lt (real_abs (s n)) k" + by (import hollight MAX_LEMMA) + +lemma SEQ_BOUNDED: "ALL s::nat => hollight.real. + bounded (mr1, >=) s = + (EX k::hollight.real. ALL n::nat. real_lt (real_abs (s n)) k)" + by (import hollight SEQ_BOUNDED) + +lemma SEQ_BOUNDED_2: "ALL (f::nat => hollight.real) (k::hollight.real) K::hollight.real. + (ALL n::nat. real_le k (f n) & real_le (f n) K) --> bounded (mr1, >=) f" + by (import hollight SEQ_BOUNDED_2) + +lemma SEQ_CBOUNDED: "ALL f::nat => hollight.real. cauchy f --> bounded (mr1, >=) f" + by (import hollight SEQ_CBOUNDED) + +lemma SEQ_ICONV: "ALL f::nat => hollight.real. + bounded (mr1, >=) f & + (ALL (m::nat) n::nat. >= m n --> hollight.real_ge (f m) (f n)) --> + convergent f" + by (import hollight SEQ_ICONV) + +lemma SEQ_NEG_CONV: "ALL f::nat => hollight.real. + convergent f = convergent (%n::nat. real_neg (f n))" + by (import hollight SEQ_NEG_CONV) + +lemma SEQ_NEG_BOUNDED: "ALL f::nat => hollight.real. + bounded (mr1, >=) (%n::nat. real_neg (f n)) = bounded (mr1, >=) f" + by (import hollight SEQ_NEG_BOUNDED) + +lemma SEQ_BCONV: "ALL f::nat => hollight.real. + bounded (mr1, >=) f & hollight.mono f --> convergent f" + by (import hollight SEQ_BCONV) + +lemma SEQ_MONOSUB: "ALL s::nat => hollight.real. + EX f::nat => nat. subseq f & hollight.mono (%n::nat. s (f n))" + by (import hollight SEQ_MONOSUB) + +lemma SEQ_SBOUNDED: "ALL (s::nat => hollight.real) f::nat => nat. + bounded (mr1, >=) s --> bounded (mr1, >=) (%n::nat. s (f n))" + by (import hollight SEQ_SBOUNDED) + +lemma SEQ_SUBLE: "ALL (f::nat => nat) x::nat. subseq f --> <= x (f x)" + by (import hollight SEQ_SUBLE) + +lemma SEQ_DIRECT: "ALL f::nat => nat. + subseq f --> (ALL (N1::nat) N2::nat. EX x::nat. >= x N1 & >= (f x) N2)" + by (import hollight SEQ_DIRECT) + +lemma SEQ_CAUCHY: "ALL f::nat => hollight.real. cauchy f = convergent f" + by (import hollight SEQ_CAUCHY) + +lemma SEQ_LE: "ALL (f::nat => hollight.real) (g::nat => hollight.real) (l::hollight.real) + m::hollight.real. + tends_num_real f l & + tends_num_real g m & + (EX N::nat. ALL n::nat. >= n N --> real_le (f n) (g n)) --> + real_le l m" + by (import hollight SEQ_LE) + +lemma SEQ_LE_0: "ALL (x::nat => hollight.real) xa::nat => hollight.real. + tends_num_real x (real_of_num (0::nat)) & + (EX xb::nat. + ALL xc::nat. + >= xc xb --> real_le (real_abs (xa xc)) (real_abs (x xc))) --> + tends_num_real xa (real_of_num (0::nat))" + by (import hollight SEQ_LE_0) + +lemma SEQ_SUC: "ALL (f::nat => hollight.real) l::hollight.real. + tends_num_real f l = tends_num_real (%n::nat. f (Suc n)) l" + by (import hollight SEQ_SUC) + +lemma SEQ_ABS: "ALL f::nat => hollight.real. + tends_num_real (%n::nat. real_abs (f n)) (real_of_num (0::nat)) = + tends_num_real f (real_of_num (0::nat))" + by (import hollight SEQ_ABS) + +lemma SEQ_ABS_IMP: "ALL (f::nat => hollight.real) l::hollight.real. + tends_num_real f l --> + tends_num_real (%n::nat. real_abs (f n)) (real_abs l)" + by (import hollight SEQ_ABS_IMP) + +lemma SEQ_INV0: "ALL f::nat => hollight.real. + (ALL y::hollight.real. + EX N::nat. ALL n::nat. >= n N --> hollight.real_gt (f n) y) --> + tends_num_real (%n::nat. real_inv (f n)) (real_of_num (0::nat))" + by (import hollight SEQ_INV0) + +lemma SEQ_POWER_ABS: "ALL c::hollight.real. + real_lt (real_abs c) (real_of_num (NUMERAL_BIT1 (0::nat))) --> + tends_num_real (real_pow (real_abs c)) (real_of_num (0::nat))" + by (import hollight SEQ_POWER_ABS) + +lemma SEQ_POWER: "ALL c::hollight.real. + real_lt (real_abs c) (real_of_num (NUMERAL_BIT1 (0::nat))) --> + tends_num_real (real_pow c) (real_of_num (0::nat))" + by (import hollight SEQ_POWER) + +lemma NEST_LEMMA: "ALL (f::nat => hollight.real) g::nat => hollight.real. + (ALL n::nat. hollight.real_ge (f (Suc n)) (f n)) & + (ALL n::nat. real_le (g (Suc n)) (g n)) & + (ALL n::nat. real_le (f n) (g n)) --> + (EX (l::hollight.real) m::hollight.real. + real_le l m & + ((ALL n::nat. real_le (f n) l) & tends_num_real f l) & + (ALL n::nat. real_le m (g n)) & tends_num_real g m)" + by (import hollight NEST_LEMMA) + +lemma NEST_LEMMA_UNIQ: "ALL (f::nat => hollight.real) g::nat => hollight.real. + (ALL n::nat. hollight.real_ge (f (Suc n)) (f n)) & + (ALL n::nat. real_le (g (Suc n)) (g n)) & + (ALL n::nat. real_le (f n) (g n)) & + tends_num_real (%n::nat. real_sub (f n) (g n)) (real_of_num (0::nat)) --> + (EX l::hollight.real. + ((ALL n::nat. real_le (f n) l) & tends_num_real f l) & + (ALL n::nat. real_le l (g n)) & tends_num_real g l)" + by (import hollight NEST_LEMMA_UNIQ) + +lemma BOLZANO_LEMMA: "ALL P::hollight.real * hollight.real => bool. + (ALL (a::hollight.real) (b::hollight.real) c::hollight.real. + real_le a b & real_le b c & P (a, b) & P (b, c) --> P (a, c)) & + (ALL x::hollight.real. + EX d::hollight.real. + real_lt (real_of_num (0::nat)) d & + (ALL (a::hollight.real) b::hollight.real. + real_le a x & real_le x b & real_lt (real_sub b a) d --> + P (a, b))) --> + (ALL (a::hollight.real) b::hollight.real. real_le a b --> P (a, b))" + by (import hollight BOLZANO_LEMMA) + +constdefs + sums :: "(nat => hollight.real) => hollight.real => bool" + "sums == +%u::nat => hollight.real. tends_num_real (%n::nat. psum (0::nat, n) u)" + +lemma DEF_sums: "sums = +(%u::nat => hollight.real. tends_num_real (%n::nat. psum (0::nat, n) u))" + by (import hollight DEF_sums) + +constdefs + summable :: "(nat => hollight.real) => bool" + "summable == %u::nat => hollight.real. Ex (sums u)" + +lemma DEF_summable: "summable = (%u::nat => hollight.real. Ex (sums u))" + by (import hollight DEF_summable) + +constdefs + suminf :: "(nat => hollight.real) => hollight.real" + "suminf == %u::nat => hollight.real. Eps (sums u)" + +lemma DEF_suminf: "suminf = (%u::nat => hollight.real. Eps (sums u))" + by (import hollight DEF_suminf) + +lemma SUM_SUMMABLE: "ALL (f::nat => hollight.real) l::hollight.real. sums f l --> summable f" + by (import hollight SUM_SUMMABLE) + +lemma SUMMABLE_SUM: "ALL f::nat => hollight.real. summable f --> sums f (suminf f)" + by (import hollight SUMMABLE_SUM) + +lemma SUM_UNIQ: "ALL (f::nat => hollight.real) x::hollight.real. sums f x --> x = suminf f" + by (import hollight SUM_UNIQ) + +lemma SER_UNIQ: "ALL (f::nat => hollight.real) (x::hollight.real) y::hollight.real. + sums f x & sums f y --> x = y" + by (import hollight SER_UNIQ) + +lemma SER_0: "ALL (f::nat => hollight.real) n::nat. + (ALL m::nat. <= n m --> f m = real_of_num (0::nat)) --> + sums f (psum (0::nat, n) f)" + by (import hollight SER_0) + +lemma SER_POS_LE: "ALL (f::nat => hollight.real) n::nat. + summable f & + (ALL m::nat. <= n m --> real_le (real_of_num (0::nat)) (f m)) --> + real_le (psum (0::nat, n) f) (suminf f)" + by (import hollight SER_POS_LE) + +lemma SER_POS_LT: "ALL (f::nat => hollight.real) n::nat. + summable f & + (ALL m::nat. <= n m --> real_lt (real_of_num (0::nat)) (f m)) --> + real_lt (psum (0::nat, n) f) (suminf f)" + by (import hollight SER_POS_LT) + +lemma SER_GROUP: "ALL (f::nat => hollight.real) k::nat. + summable f & < (0::nat) k --> + sums (%n::nat. psum (n * k, k) f) (suminf f)" + by (import hollight SER_GROUP) + +lemma SER_PAIR: "ALL f::nat => hollight.real. + summable f --> + sums + (%n::nat. + psum + (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)) * n, + NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))) + f) + (suminf f)" + by (import hollight SER_PAIR) + +lemma SER_OFFSET: "ALL f::nat => hollight.real. + summable f --> + (ALL k::nat. + sums (%n::nat. f (n + k)) (real_sub (suminf f) (psum (0::nat, k) f)))" + by (import hollight SER_OFFSET) + +lemma SER_OFFSET_REV: "ALL (f::nat => hollight.real) k::nat. + summable (%n::nat. f (n + k)) --> + sums f (real_add (psum (0::nat, k) f) (suminf (%n::nat. f (n + k))))" + by (import hollight SER_OFFSET_REV) + +lemma SER_POS_LT_PAIR: "ALL (f::nat => hollight.real) n::nat. + summable f & + (ALL d::nat. + real_lt (real_of_num (0::nat)) + (real_add (f (n + NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)) * d)) + (f (n + + (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)) * d + + NUMERAL_BIT1 (0::nat)))))) --> + real_lt (psum (0::nat, n) f) (suminf f)" + by (import hollight SER_POS_LT_PAIR) + +lemma SER_ADD: "ALL (x::nat => hollight.real) (x0::hollight.real) (y::nat => hollight.real) + y0::hollight.real. + sums x x0 & sums y y0 --> + sums (%n::nat. real_add (x n) (y n)) (real_add x0 y0)" + by (import hollight SER_ADD) + +lemma SER_CMUL: "ALL (x::nat => hollight.real) (x0::hollight.real) c::hollight.real. + sums x x0 --> sums (%n::nat. real_mul c (x n)) (real_mul c x0)" + by (import hollight SER_CMUL) + +lemma SER_NEG: "ALL (x::nat => hollight.real) x0::hollight.real. + sums x x0 --> sums (%xa::nat. real_neg (x xa)) (real_neg x0)" + by (import hollight SER_NEG) + +lemma SER_SUB: "ALL (x::nat => hollight.real) (x0::hollight.real) (y::nat => hollight.real) + y0::hollight.real. + sums x x0 & sums y y0 --> + sums (%n::nat. real_sub (x n) (y n)) (real_sub x0 y0)" + by (import hollight SER_SUB) + +lemma SER_CDIV: "ALL (x::nat => hollight.real) (x0::hollight.real) c::hollight.real. + sums x x0 --> sums (%xa::nat. real_div (x xa) c) (real_div x0 c)" + by (import hollight SER_CDIV) + +lemma SER_CAUCHY: "ALL f::nat => hollight.real. + summable f = + (ALL e::hollight.real. + real_lt (real_of_num (0::nat)) e --> + (EX N::nat. + ALL (m::nat) n::nat. + >= m N --> real_lt (real_abs (psum (m, n) f)) e))" + by (import hollight SER_CAUCHY) + +lemma SER_ZERO: "ALL f::nat => hollight.real. + summable f --> tends_num_real f (real_of_num (0::nat))" + by (import hollight SER_ZERO) + +lemma SER_COMPAR: "ALL (f::nat => hollight.real) g::nat => hollight.real. + (EX x::nat. ALL xa::nat. >= xa x --> real_le (real_abs (f xa)) (g xa)) & + summable g --> + summable f" + by (import hollight SER_COMPAR) + +lemma SER_COMPARA: "ALL (f::nat => hollight.real) g::nat => hollight.real. + (EX x::nat. ALL xa::nat. >= xa x --> real_le (real_abs (f xa)) (g xa)) & + summable g --> + summable (%k::nat. real_abs (f k))" + by (import hollight SER_COMPARA) + +lemma SER_LE: "ALL (f::nat => hollight.real) g::nat => hollight.real. + (ALL n::nat. real_le (f n) (g n)) & summable f & summable g --> + real_le (suminf f) (suminf g)" + by (import hollight SER_LE) + +lemma SER_LE2: "ALL (f::nat => hollight.real) g::nat => hollight.real. + (ALL n::nat. real_le (real_abs (f n)) (g n)) & summable g --> + summable f & real_le (suminf f) (suminf g)" + by (import hollight SER_LE2) + +lemma SER_ACONV: "ALL f::nat => hollight.real. + summable (%n::nat. real_abs (f n)) --> summable f" + by (import hollight SER_ACONV) + +lemma SER_ABS: "ALL f::nat => hollight.real. + summable (%n::nat. real_abs (f n)) --> + real_le (real_abs (suminf f)) (suminf (%n::nat. real_abs (f n)))" + by (import hollight SER_ABS) + +lemma GP_FINITE: "ALL x::hollight.real. + x ~= real_of_num (NUMERAL_BIT1 (0::nat)) --> + (ALL n::nat. + psum (0::nat, n) (real_pow x) = + real_div + (real_sub (real_pow x n) (real_of_num (NUMERAL_BIT1 (0::nat)))) + (real_sub x (real_of_num (NUMERAL_BIT1 (0::nat)))))" + by (import hollight GP_FINITE) + +lemma GP: "ALL x::hollight.real. + real_lt (real_abs x) (real_of_num (NUMERAL_BIT1 (0::nat))) --> + sums (real_pow x) + (real_inv (real_sub (real_of_num (NUMERAL_BIT1 (0::nat))) x))" + by (import hollight GP) + +lemma ABS_NEG_LEMMA: "ALL (c::hollight.real) (x::hollight.real) y::hollight.real. + real_le c (real_of_num (0::nat)) --> + real_le (real_abs x) (real_mul c (real_abs y)) --> + x = real_of_num (0::nat)" + by (import hollight ABS_NEG_LEMMA) + +lemma SER_RATIO: "ALL (f::nat => hollight.real) (c::hollight.real) N::nat. + real_lt c (real_of_num (NUMERAL_BIT1 (0::nat))) & + (ALL n::nat. + >= n N --> + real_le (real_abs (f (Suc n))) (real_mul c (real_abs (f n)))) --> + summable f" + by (import hollight SER_RATIO) + +lemma SEQ_TRUNCATION: "ALL (f::nat => hollight.real) (l::hollight.real) (n::nat) b::hollight.real. + sums f l & (ALL m::nat. real_le (real_abs (psum (n, m) f)) b) --> + real_le (real_abs (real_sub l (psum (0::nat, n) f))) b" + by (import hollight SEQ_TRUNCATION) + +constdefs + tends_real_real :: "(hollight.real => hollight.real) => hollight.real => hollight.real => bool" + "tends_real_real == +%(u::hollight.real => hollight.real) (ua::hollight.real) ub::hollight.real. + tends u ua (mtop mr1, tendsto (mr1, ub))" + +lemma DEF_tends_real_real: "tends_real_real = +(%(u::hollight.real => hollight.real) (ua::hollight.real) ub::hollight.real. + tends u ua (mtop mr1, tendsto (mr1, ub)))" + by (import hollight DEF_tends_real_real) + +lemma LIM: "ALL (f::hollight.real => hollight.real) (y0::hollight.real) + x0::hollight.real. + tends_real_real f y0 x0 = + (ALL e::hollight.real. + real_lt (real_of_num (0::nat)) e --> + (EX d::hollight.real. + real_lt (real_of_num (0::nat)) d & + (ALL x::hollight.real. + real_lt (real_of_num (0::nat)) (real_abs (real_sub x x0)) & + real_lt (real_abs (real_sub x x0)) d --> + real_lt (real_abs (real_sub (f x) y0)) e)))" + by (import hollight LIM) + +lemma LIM_CONST: "ALL k::hollight.real. All (tends_real_real (%x::hollight.real. k) k)" + by (import hollight LIM_CONST) + +lemma LIM_ADD: "ALL (f::hollight.real => hollight.real) (g::hollight.real => hollight.real) + (l::hollight.real) m::hollight.real. + tends_real_real f l (x::hollight.real) & tends_real_real g m x --> + tends_real_real (%x::hollight.real. real_add (f x) (g x)) (real_add l m) + x" + by (import hollight LIM_ADD) + +lemma LIM_MUL: "ALL (f::hollight.real => hollight.real) (g::hollight.real => hollight.real) + (l::hollight.real) m::hollight.real. + tends_real_real f l (x::hollight.real) & tends_real_real g m x --> + tends_real_real (%x::hollight.real. real_mul (f x) (g x)) (real_mul l m) + x" + by (import hollight LIM_MUL) + +lemma LIM_NEG: "ALL (f::hollight.real => hollight.real) l::hollight.real. + tends_real_real f l (x::hollight.real) = + tends_real_real (%x::hollight.real. real_neg (f x)) (real_neg l) x" + by (import hollight LIM_NEG) + +lemma LIM_INV: "ALL (f::hollight.real => hollight.real) l::hollight.real. + tends_real_real f l (x::hollight.real) & l ~= real_of_num (0::nat) --> + tends_real_real (%x::hollight.real. real_inv (f x)) (real_inv l) x" + by (import hollight LIM_INV) + +lemma LIM_SUB: "ALL (f::hollight.real => hollight.real) (g::hollight.real => hollight.real) + (l::hollight.real) m::hollight.real. + tends_real_real f l (x::hollight.real) & tends_real_real g m x --> + tends_real_real (%x::hollight.real. real_sub (f x) (g x)) (real_sub l m) + x" + by (import hollight LIM_SUB) + +lemma LIM_DIV: "ALL (f::hollight.real => hollight.real) (g::hollight.real => hollight.real) + (l::hollight.real) m::hollight.real. + tends_real_real f l (x::hollight.real) & + tends_real_real g m x & m ~= real_of_num (0::nat) --> + tends_real_real (%x::hollight.real. real_div (f x) (g x)) (real_div l m) + x" + by (import hollight LIM_DIV) + +lemma LIM_NULL: "ALL (f::hollight.real => hollight.real) (l::hollight.real) x::hollight.real. + tends_real_real f l x = + tends_real_real (%x::hollight.real. real_sub (f x) l) + (real_of_num (0::nat)) x" + by (import hollight LIM_NULL) + +lemma LIM_SUM: "ALL (f::nat => hollight.real => hollight.real) (l::nat => hollight.real) + (m::nat) (n::nat) x::hollight.real. + (ALL xa::nat. + <= m xa & < xa (m + n) --> tends_real_real (f xa) (l xa) x) --> + tends_real_real (%x::hollight.real. psum (m, n) (%r::nat. f r x)) + (psum (m, n) l) x" + by (import hollight LIM_SUM) + +lemma LIM_X: "ALL x0::hollight.real. tends_real_real (%x::hollight.real. x) x0 x0" + by (import hollight LIM_X) + +lemma LIM_UNIQ: "ALL (f::hollight.real => hollight.real) (l::hollight.real) + (m::hollight.real) x::hollight.real. + tends_real_real f l x & tends_real_real f m x --> l = m" + by (import hollight LIM_UNIQ) + +lemma LIM_EQUAL: "ALL (f::hollight.real => hollight.real) (g::hollight.real => hollight.real) + (l::hollight.real) x0::hollight.real. + (ALL x::hollight.real. x ~= x0 --> f x = g x) --> + tends_real_real f l x0 = tends_real_real g l x0" + by (import hollight LIM_EQUAL) + +lemma LIM_TRANSFORM: "ALL (f::hollight.real => hollight.real) (g::hollight.real => hollight.real) + (x0::hollight.real) l::hollight.real. + tends_real_real (%x::hollight.real. real_sub (f x) (g x)) + (real_of_num (0::nat)) x0 & + tends_real_real g l x0 --> + tends_real_real f l x0" + by (import hollight LIM_TRANSFORM) + +constdefs + diffl :: "(hollight.real => hollight.real) => hollight.real => hollight.real => bool" + "diffl == +%(u::hollight.real => hollight.real) (ua::hollight.real) ub::hollight.real. + tends_real_real + (%h::hollight.real. real_div (real_sub (u (real_add ub h)) (u ub)) h) ua + (real_of_num (0::nat))" + +lemma DEF_diffl: "diffl = +(%(u::hollight.real => hollight.real) (ua::hollight.real) ub::hollight.real. + tends_real_real + (%h::hollight.real. real_div (real_sub (u (real_add ub h)) (u ub)) h) + ua (real_of_num (0::nat)))" + by (import hollight DEF_diffl) + +constdefs + contl :: "(hollight.real => hollight.real) => hollight.real => bool" + "contl == +%(u::hollight.real => hollight.real) ua::hollight.real. + tends_real_real (%h::hollight.real. u (real_add ua h)) (u ua) + (real_of_num (0::nat))" + +lemma DEF_contl: "contl = +(%(u::hollight.real => hollight.real) ua::hollight.real. + tends_real_real (%h::hollight.real. u (real_add ua h)) (u ua) + (real_of_num (0::nat)))" + by (import hollight DEF_contl) + +constdefs + differentiable :: "(hollight.real => hollight.real) => hollight.real => bool" + "differentiable == +%(u::hollight.real => hollight.real) ua::hollight.real. + EX l::hollight.real. diffl u l ua" + +lemma DEF_differentiable: "differentiable = +(%(u::hollight.real => hollight.real) ua::hollight.real. + EX l::hollight.real. diffl u l ua)" + by (import hollight DEF_differentiable) + +lemma DIFF_UNIQ: "ALL (f::hollight.real => hollight.real) (l::hollight.real) + (m::hollight.real) x::hollight.real. diffl f l x & diffl f m x --> l = m" + by (import hollight DIFF_UNIQ) + +lemma DIFF_CONT: "ALL (f::hollight.real => hollight.real) (l::hollight.real) x::hollight.real. + diffl f l x --> contl f x" + by (import hollight DIFF_CONT) + +lemma CONTL_LIM: "ALL (f::hollight.real => hollight.real) x::hollight.real. + contl f x = tends_real_real f (f x) x" + by (import hollight CONTL_LIM) + +lemma CONT_X: "All (contl (%x::hollight.real. x))" + by (import hollight CONT_X) + +lemma CONT_CONST: "All (contl (%x::hollight.real. k::hollight.real))" + by (import hollight CONT_CONST) + +lemma CONT_ADD: "ALL x::hollight.real. + contl (f::hollight.real => hollight.real) x & + contl (g::hollight.real => hollight.real) x --> + contl (%x::hollight.real. real_add (f x) (g x)) x" + by (import hollight CONT_ADD) + +lemma CONT_MUL: "ALL x::hollight.real. + contl (f::hollight.real => hollight.real) x & + contl (g::hollight.real => hollight.real) x --> + contl (%x::hollight.real. real_mul (f x) (g x)) x" + by (import hollight CONT_MUL) + +lemma CONT_NEG: "ALL x::hollight.real. + contl (f::hollight.real => hollight.real) x --> + contl (%x::hollight.real. real_neg (f x)) x" + by (import hollight CONT_NEG) + +lemma CONT_INV: "ALL x::hollight.real. + contl (f::hollight.real => hollight.real) x & + f x ~= real_of_num (0::nat) --> + contl (%x::hollight.real. real_inv (f x)) x" + by (import hollight CONT_INV) + +lemma CONT_SUB: "ALL x::hollight.real. + contl (f::hollight.real => hollight.real) x & + contl (g::hollight.real => hollight.real) x --> + contl (%x::hollight.real. real_sub (f x) (g x)) x" + by (import hollight CONT_SUB) + +lemma CONT_DIV: "ALL x::hollight.real. + contl (f::hollight.real => hollight.real) x & + contl (g::hollight.real => hollight.real) x & + g x ~= real_of_num (0::nat) --> + contl (%x::hollight.real. real_div (f x) (g x)) x" + by (import hollight CONT_DIV) + +lemma CONT_COMPOSE: "ALL (f::hollight.real => hollight.real) (g::hollight.real => hollight.real) + x::hollight.real. + contl f x & contl g (f x) --> contl (%x::hollight.real. g (f x)) x" + by (import hollight CONT_COMPOSE) + +lemma IVT: "ALL (f::hollight.real => hollight.real) (a::hollight.real) + (b::hollight.real) y::hollight.real. + real_le a b & + (real_le (f a) y & real_le y (f b)) & + (ALL x::hollight.real. real_le a x & real_le x b --> contl f x) --> + (EX x::hollight.real. real_le a x & real_le x b & f x = y)" + by (import hollight IVT) + +lemma IVT2: "ALL (f::hollight.real => hollight.real) (a::hollight.real) + (b::hollight.real) y::hollight.real. + real_le a b & + (real_le (f b) y & real_le y (f a)) & + (ALL x::hollight.real. real_le a x & real_le x b --> contl f x) --> + (EX x::hollight.real. real_le a x & real_le x b & f x = y)" + by (import hollight IVT2) + +lemma DIFF_CONST: "ALL k::hollight.real. + All (diffl (%x::hollight.real. k) (real_of_num (0::nat)))" + by (import hollight DIFF_CONST) + +lemma DIFF_ADD: "ALL (f::hollight.real => hollight.real) (g::hollight.real => hollight.real) + (l::hollight.real) (m::hollight.real) x::hollight.real. + diffl f l x & diffl g m x --> + diffl (%x::hollight.real. real_add (f x) (g x)) (real_add l m) x" + by (import hollight DIFF_ADD) + +lemma DIFF_MUL: "ALL (f::hollight.real => hollight.real) (g::hollight.real => hollight.real) + (l::hollight.real) (m::hollight.real) x::hollight.real. + diffl f l x & diffl g m x --> + diffl (%x::hollight.real. real_mul (f x) (g x)) + (real_add (real_mul l (g x)) (real_mul m (f x))) x" + by (import hollight DIFF_MUL) + +lemma DIFF_CMUL: "ALL (f::hollight.real => hollight.real) (c::hollight.real) + (l::hollight.real) x::hollight.real. + diffl f l x --> + diffl (%x::hollight.real. real_mul c (f x)) (real_mul c l) x" + by (import hollight DIFF_CMUL) + +lemma DIFF_NEG: "ALL (f::hollight.real => hollight.real) (l::hollight.real) x::hollight.real. + diffl f l x --> diffl (%x::hollight.real. real_neg (f x)) (real_neg l) x" + by (import hollight DIFF_NEG) + +lemma DIFF_SUB: "ALL (f::hollight.real => hollight.real) (g::hollight.real => hollight.real) + (l::hollight.real) (m::hollight.real) x::hollight.real. + diffl f l x & diffl g m x --> + diffl (%x::hollight.real. real_sub (f x) (g x)) (real_sub l m) x" + by (import hollight DIFF_SUB) + +lemma DIFF_CARAT: "ALL (f::hollight.real => hollight.real) (l::hollight.real) x::hollight.real. + diffl f l x = + (EX xa::hollight.real => hollight.real. + (ALL z::hollight.real. + real_sub (f z) (f x) = real_mul (xa z) (real_sub z x)) & + contl xa x & xa x = l)" + by (import hollight DIFF_CARAT) + +lemma DIFF_CHAIN: "ALL (f::hollight.real => hollight.real) (g::hollight.real => hollight.real) + (l::hollight.real) (m::hollight.real) x::hollight.real. + diffl f l (g x) & diffl g m x --> + diffl (%x::hollight.real. f (g x)) (real_mul l m) x" + by (import hollight DIFF_CHAIN) + +lemma DIFF_X: "All (diffl (%x::hollight.real. x) (real_of_num (NUMERAL_BIT1 (0::nat))))" + by (import hollight DIFF_X) + +lemma DIFF_POW: "ALL (n::nat) x::hollight.real. + diffl (%x::hollight.real. real_pow x n) + (real_mul (real_of_num n) (real_pow x (n - NUMERAL_BIT1 (0::nat)))) x" + by (import hollight DIFF_POW) + +lemma DIFF_XM1: "ALL x::hollight.real. + x ~= real_of_num (0::nat) --> + diffl real_inv + (real_neg + (real_pow (real_inv x) (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) + x" + by (import hollight DIFF_XM1) + +lemma DIFF_INV: "ALL (f::hollight.real => hollight.real) (l::hollight.real) x::hollight.real. + diffl f l x & f x ~= real_of_num (0::nat) --> + diffl (%x::hollight.real. real_inv (f x)) + (real_neg + (real_div l (real_pow (f x) (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))))) + x" + by (import hollight DIFF_INV) + +lemma DIFF_DIV: "ALL (f::hollight.real => hollight.real) (g::hollight.real => hollight.real) + (l::hollight.real) m::hollight.real. + diffl f l (x::hollight.real) & + diffl g m x & g x ~= real_of_num (0::nat) --> + diffl (%x::hollight.real. real_div (f x) (g x)) + (real_div (real_sub (real_mul l (g x)) (real_mul m (f x))) + (real_pow (g x) (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) + x" + by (import hollight DIFF_DIV) + +lemma DIFF_SUM: "ALL (f::nat => hollight.real => hollight.real) + (f'::nat => hollight.real => hollight.real) (m::nat) (n::nat) + x::hollight.real. + (ALL r::nat. <= m r & < r (m + n) --> diffl (f r) (f' r x) x) --> + diffl (%x::hollight.real. psum (m, n) (%n::nat. f n x)) + (psum (m, n) (%r::nat. f' r x)) x" + by (import hollight DIFF_SUM) + +lemma CONT_BOUNDED: "ALL (f::hollight.real => hollight.real) (a::hollight.real) b::hollight.real. + real_le a b & + (ALL x::hollight.real. real_le a x & real_le x b --> contl f x) --> + (EX M::hollight.real. + ALL x::hollight.real. real_le a x & real_le x b --> real_le (f x) M)" + by (import hollight CONT_BOUNDED) + +lemma CONT_BOUNDED_ABS: "ALL (f::hollight.real => hollight.real) (a::hollight.real) b::hollight.real. + (ALL x::hollight.real. real_le a x & real_le x b --> contl f x) --> + (EX M::hollight.real. + ALL x::hollight.real. + real_le a x & real_le x b --> real_le (real_abs (f x)) M)" + by (import hollight CONT_BOUNDED_ABS) + +lemma CONT_HASSUP: "ALL (f::hollight.real => hollight.real) (a::hollight.real) b::hollight.real. + real_le a b & + (ALL x::hollight.real. real_le a x & real_le x b --> contl f x) --> + (EX M::hollight.real. + (ALL x::hollight.real. + real_le a x & real_le x b --> real_le (f x) M) & + (ALL N::hollight.real. + real_lt N M --> + (EX x::hollight.real. + real_le a x & real_le x b & real_lt N (f x))))" + by (import hollight CONT_HASSUP) + +lemma CONT_ATTAINS: "ALL (f::hollight.real => hollight.real) (a::hollight.real) b::hollight.real. + real_le a b & + (ALL x::hollight.real. real_le a x & real_le x b --> contl f x) --> + (EX x::hollight.real. + (ALL xa::hollight.real. + real_le a xa & real_le xa b --> real_le (f xa) x) & + (EX xa::hollight.real. real_le a xa & real_le xa b & f xa = x))" + by (import hollight CONT_ATTAINS) + +lemma CONT_ATTAINS2: "ALL (f::hollight.real => hollight.real) (a::hollight.real) b::hollight.real. + real_le a b & + (ALL x::hollight.real. real_le a x & real_le x b --> contl f x) --> + (EX M::hollight.real. + (ALL x::hollight.real. + real_le a x & real_le x b --> real_le M (f x)) & + (EX x::hollight.real. real_le a x & real_le x b & f x = M))" + by (import hollight CONT_ATTAINS2) + +lemma CONT_ATTAINS_ALL: "ALL (f::hollight.real => hollight.real) (a::hollight.real) b::hollight.real. + real_le a b & + (ALL x::hollight.real. real_le a x & real_le x b --> contl f x) --> + (EX (L::hollight.real) M::hollight.real. + (ALL x::hollight.real. + real_le a x & real_le x b --> + real_le L (f x) & real_le (f x) M) & + (ALL y::hollight.real. + real_le L y & real_le y M --> + (EX x::hollight.real. real_le a x & real_le x b & f x = y)))" + by (import hollight CONT_ATTAINS_ALL) + +lemma DIFF_LINC: "ALL (f::hollight.real => hollight.real) (x::hollight.real) l::hollight.real. + diffl f l x & real_lt (real_of_num (0::nat)) l --> + (EX d::hollight.real. + real_lt (real_of_num (0::nat)) d & + (ALL h::hollight.real. + real_lt (real_of_num (0::nat)) h & real_lt h d --> + real_lt (f x) (f (real_add x h))))" + by (import hollight DIFF_LINC) + +lemma DIFF_LDEC: "ALL (f::hollight.real => hollight.real) (x::hollight.real) l::hollight.real. + diffl f l x & real_lt l (real_of_num (0::nat)) --> + (EX d::hollight.real. + real_lt (real_of_num (0::nat)) d & + (ALL h::hollight.real. + real_lt (real_of_num (0::nat)) h & real_lt h d --> + real_lt (f x) (f (real_sub x h))))" + by (import hollight DIFF_LDEC) + +lemma DIFF_LMAX: "ALL (f::hollight.real => hollight.real) (x::hollight.real) l::hollight.real. + diffl f l x & + (EX d::hollight.real. + real_lt (real_of_num (0::nat)) d & + (ALL y::hollight.real. + real_lt (real_abs (real_sub x y)) d --> real_le (f y) (f x))) --> + l = real_of_num (0::nat)" + by (import hollight DIFF_LMAX) + +lemma DIFF_LMIN: "ALL (f::hollight.real => hollight.real) (x::hollight.real) l::hollight.real. + diffl f l x & + (EX d::hollight.real. + real_lt (real_of_num (0::nat)) d & + (ALL y::hollight.real. + real_lt (real_abs (real_sub x y)) d --> real_le (f x) (f y))) --> + l = real_of_num (0::nat)" + by (import hollight DIFF_LMIN) + +lemma DIFF_LCONST: "ALL (f::hollight.real => hollight.real) (x::hollight.real) l::hollight.real. + diffl f l x & + (EX d::hollight.real. + real_lt (real_of_num (0::nat)) d & + (ALL y::hollight.real. + real_lt (real_abs (real_sub x y)) d --> f y = f x)) --> + l = real_of_num (0::nat)" + by (import hollight DIFF_LCONST) + +lemma INTERVAL_LEMMA_LT: "ALL (a::hollight.real) (b::hollight.real) x::hollight.real. + real_lt a x & real_lt x b --> + (EX xa::hollight.real. + real_lt (real_of_num (0::nat)) xa & + (ALL xb::hollight.real. + real_lt (real_abs (real_sub x xb)) xa --> + real_lt a xb & real_lt xb b))" + by (import hollight INTERVAL_LEMMA_LT) + +lemma INTERVAL_LEMMA: "ALL (a::hollight.real) (b::hollight.real) x::hollight.real. + real_lt a x & real_lt x b --> + (EX xa::hollight.real. + real_lt (real_of_num (0::nat)) xa & + (ALL y::hollight.real. + real_lt (real_abs (real_sub x y)) xa --> + real_le a y & real_le y b))" + by (import hollight INTERVAL_LEMMA) + +lemma ROLLE: "ALL (f::hollight.real => hollight.real) (a::hollight.real) b::hollight.real. + real_lt a b & + f a = f b & + (ALL x::hollight.real. real_le a x & real_le x b --> contl f x) & + (ALL x::hollight.real. + real_lt a x & real_lt x b --> differentiable f x) --> + (EX z::hollight.real. + real_lt a z & real_lt z b & diffl f (real_of_num (0::nat)) z)" + by (import hollight ROLLE) + +lemma MVT_LEMMA: "ALL (f::hollight.real => hollight.real) (a::hollight.real) b::hollight.real. + real_sub (f a) + (real_mul (real_div (real_sub (f b) (f a)) (real_sub b a)) a) = + real_sub (f b) + (real_mul (real_div (real_sub (f b) (f a)) (real_sub b a)) b)" + by (import hollight MVT_LEMMA) + +lemma MVT: "ALL (f::hollight.real => hollight.real) (a::hollight.real) b::hollight.real. + real_lt a b & + (ALL x::hollight.real. real_le a x & real_le x b --> contl f x) & + (ALL x::hollight.real. + real_lt a x & real_lt x b --> differentiable f x) --> + (EX (l::hollight.real) z::hollight.real. + real_lt a z & + real_lt z b & + diffl f l z & real_sub (f b) (f a) = real_mul (real_sub b a) l)" + by (import hollight MVT) + +lemma MVT_ALT: "ALL (f::hollight.real => hollight.real) (f'::hollight.real => hollight.real) + (a::hollight.real) b::hollight.real. + real_lt a b & + (ALL x::hollight.real. + real_le a x & real_le x b --> diffl f (f' x) x) --> + (EX z::hollight.real. + real_lt a z & + real_lt z b & real_sub (f b) (f a) = real_mul (real_sub b a) (f' z))" + by (import hollight MVT_ALT) + +lemma DIFF_ISCONST_END: "ALL (f::hollight.real => hollight.real) (a::hollight.real) b::hollight.real. + real_lt a b & + (ALL x::hollight.real. real_le a x & real_le x b --> contl f x) & + (ALL x::hollight.real. + real_lt a x & real_lt x b --> diffl f (real_of_num (0::nat)) x) --> + f b = f a" + by (import hollight DIFF_ISCONST_END) + +lemma DIFF_ISCONST: "ALL (f::hollight.real => hollight.real) (a::hollight.real) b::hollight.real. + real_lt a b & + (ALL x::hollight.real. real_le a x & real_le x b --> contl f x) & + (ALL x::hollight.real. + real_lt a x & real_lt x b --> diffl f (real_of_num (0::nat)) x) --> + (ALL x::hollight.real. real_le a x & real_le x b --> f x = f a)" + by (import hollight DIFF_ISCONST) + +lemma DIFF_ISCONST_END_SIMPLE: "ALL (f::hollight.real => hollight.real) (a::hollight.real) b::hollight.real. + real_lt a b & + (ALL x::hollight.real. + real_le a x & real_le x b --> diffl f (real_of_num (0::nat)) x) --> + f b = f a" + by (import hollight DIFF_ISCONST_END_SIMPLE) + +lemma DIFF_ISCONST_ALL: "ALL (x::hollight.real => hollight.real) (xa::hollight.real) + xb::hollight.real. All (diffl x (real_of_num (0::nat))) --> x xa = x xb" + by (import hollight DIFF_ISCONST_ALL) + +lemma INTERVAL_ABS: "ALL (x::hollight.real) (z::hollight.real) d::hollight.real. + (real_le (real_sub x d) z & real_le z (real_add x d)) = + real_le (real_abs (real_sub z x)) d" + by (import hollight INTERVAL_ABS) + +lemma CONT_INJ_LEMMA: "ALL (f::hollight.real => hollight.real) (g::hollight.real => hollight.real) + (x::hollight.real) d::hollight.real. + real_lt (real_of_num (0::nat)) d & + (ALL z::hollight.real. + real_le (real_abs (real_sub z x)) d --> g (f z) = z) & + (ALL z::hollight.real. + real_le (real_abs (real_sub z x)) d --> contl f z) --> + ~ (ALL z::hollight.real. + real_le (real_abs (real_sub z x)) d --> real_le (f z) (f x))" + by (import hollight CONT_INJ_LEMMA) + +lemma CONT_INJ_LEMMA2: "ALL (f::hollight.real => hollight.real) (g::hollight.real => hollight.real) + (x::hollight.real) d::hollight.real. + real_lt (real_of_num (0::nat)) d & + (ALL z::hollight.real. + real_le (real_abs (real_sub z x)) d --> g (f z) = z) & + (ALL z::hollight.real. + real_le (real_abs (real_sub z x)) d --> contl f z) --> + ~ (ALL z::hollight.real. + real_le (real_abs (real_sub z x)) d --> real_le (f x) (f z))" + by (import hollight CONT_INJ_LEMMA2) + +lemma CONT_INJ_RANGE: "ALL (f::hollight.real => hollight.real) (g::hollight.real => hollight.real) + (x::hollight.real) d::hollight.real. + real_lt (real_of_num (0::nat)) d & + (ALL z::hollight.real. + real_le (real_abs (real_sub z x)) d --> g (f z) = z) & + (ALL z::hollight.real. + real_le (real_abs (real_sub z x)) d --> contl f z) --> + (EX e::hollight.real. + real_lt (real_of_num (0::nat)) e & + (ALL y::hollight.real. + real_le (real_abs (real_sub y (f x))) e --> + (EX z::hollight.real. + real_le (real_abs (real_sub z x)) d & f z = y)))" + by (import hollight CONT_INJ_RANGE) + +lemma CONT_INVERSE: "ALL (f::hollight.real => hollight.real) (g::hollight.real => hollight.real) + (x::hollight.real) d::hollight.real. + real_lt (real_of_num (0::nat)) d & + (ALL z::hollight.real. + real_le (real_abs (real_sub z x)) d --> g (f z) = z) & + (ALL z::hollight.real. + real_le (real_abs (real_sub z x)) d --> contl f z) --> + contl g (f x)" + by (import hollight CONT_INVERSE) + +lemma DIFF_INVERSE: "ALL (f::hollight.real => hollight.real) (g::hollight.real => hollight.real) + (l::hollight.real) (x::hollight.real) d::hollight.real. + real_lt (real_of_num (0::nat)) d & + (ALL z::hollight.real. + real_le (real_abs (real_sub z x)) d --> g (f z) = z) & + (ALL z::hollight.real. + real_le (real_abs (real_sub z x)) d --> contl f z) & + diffl f l x & l ~= real_of_num (0::nat) --> + diffl g (real_inv l) (f x)" + by (import hollight DIFF_INVERSE) + +lemma DIFF_INVERSE_LT: "ALL (f::hollight.real => hollight.real) (g::hollight.real => hollight.real) + (l::hollight.real) (x::hollight.real) d::hollight.real. + real_lt (real_of_num (0::nat)) d & + (ALL z::hollight.real. + real_lt (real_abs (real_sub z x)) d --> g (f z) = z) & + (ALL z::hollight.real. + real_lt (real_abs (real_sub z x)) d --> contl f z) & + diffl f l x & l ~= real_of_num (0::nat) --> + diffl g (real_inv l) (f x)" + by (import hollight DIFF_INVERSE_LT) + +lemma IVT_DERIVATIVE_0: "ALL (f::hollight.real => hollight.real) (f'::hollight.real => hollight.real) + (a::hollight.real) b::hollight.real. + real_le a b & + (ALL x::hollight.real. real_le a x & real_le x b --> diffl f (f' x) x) & + hollight.real_gt (f' a) (real_of_num (0::nat)) & + real_lt (f' b) (real_of_num (0::nat)) --> + (EX z::hollight.real. + real_lt a z & real_lt z b & f' z = real_of_num (0::nat))" + by (import hollight IVT_DERIVATIVE_0) + +lemma IVT_DERIVATIVE_POS: "ALL (x::hollight.real => hollight.real) (xa::hollight.real => hollight.real) + (xb::hollight.real) (xc::hollight.real) xd::hollight.real. + real_le xb xc & + (ALL xd::hollight.real. + real_le xb xd & real_le xd xc --> diffl x (xa xd) xd) & + hollight.real_gt (xa xb) xd & real_lt (xa xc) xd --> + (EX z::hollight.real. real_lt xb z & real_lt z xc & xa z = xd)" + by (import hollight IVT_DERIVATIVE_POS) + +lemma IVT_DERIVATIVE_NEG: "ALL (x::hollight.real => hollight.real) (xa::hollight.real => hollight.real) + (xb::hollight.real) (xc::hollight.real) xd::hollight.real. + real_le xb xc & + (ALL xd::hollight.real. + real_le xb xd & real_le xd xc --> diffl x (xa xd) xd) & + real_lt (xa xb) xd & hollight.real_gt (xa xc) xd --> + (EX z::hollight.real. real_lt xb z & real_lt z xc & xa z = xd)" + by (import hollight IVT_DERIVATIVE_NEG) + +lemma SEQ_CONT_UNIFORM: "ALL (s::nat => hollight.real => hollight.real) + (f::hollight.real => hollight.real) x0::hollight.real. + (ALL e::hollight.real. + real_lt (real_of_num (0::nat)) e --> + (EX (N::nat) d::hollight.real. + real_lt (real_of_num (0::nat)) d & + (ALL (x::hollight.real) n::nat. + real_lt (real_abs (real_sub x x0)) d & >= n N --> + real_lt (real_abs (real_sub (s n x) (f x))) e))) & + (EX N::nat. ALL n::nat. >= n N --> contl (s n) x0) --> + contl f x0" + by (import hollight SEQ_CONT_UNIFORM) + +lemma SER_COMPARA_UNIFORM: "ALL (s::hollight.real => nat => hollight.real) (x0::hollight.real) + g::nat => hollight.real. + (EX (N::nat) d::hollight.real. + real_lt (real_of_num (0::nat)) d & + (ALL (n::nat) x::hollight.real. + real_lt (real_abs (real_sub x x0)) d & >= n N --> + real_le (real_abs (s x n)) (g n))) & + summable g --> + (EX (f::hollight.real => hollight.real) d::hollight.real. + real_lt (real_of_num (0::nat)) d & + (ALL e::hollight.real. + real_lt (real_of_num (0::nat)) e --> + (EX N::nat. + ALL (x::hollight.real) n::nat. + real_lt (real_abs (real_sub x x0)) d & >= n N --> + real_lt + (real_abs (real_sub (psum (0::nat, n) (s x)) (f x))) e)))" + by (import hollight SER_COMPARA_UNIFORM) + +lemma SER_COMPARA_UNIFORM_WEAK: "ALL (s::hollight.real => nat => hollight.real) (x0::hollight.real) + g::nat => hollight.real. + (EX (N::nat) d::hollight.real. + real_lt (real_of_num (0::nat)) d & + (ALL (n::nat) x::hollight.real. + real_lt (real_abs (real_sub x x0)) d & >= n N --> + real_le (real_abs (s x n)) (g n))) & + summable g --> + (EX f::hollight.real => hollight.real. + ALL e::hollight.real. + real_lt (real_of_num (0::nat)) e --> + (EX (N::nat) d::hollight.real. + real_lt (real_of_num (0::nat)) d & + (ALL (x::hollight.real) n::nat. + real_lt (real_abs (real_sub x x0)) d & >= n N --> + real_lt + (real_abs (real_sub (psum (0::nat, n) (s x)) (f x))) e)))" + by (import hollight SER_COMPARA_UNIFORM_WEAK) + +lemma POWDIFF_LEMMA: "ALL (n::nat) (x::hollight.real) y::hollight.real. + psum (0::nat, Suc n) + (%p::nat. real_mul (real_pow x p) (real_pow y (Suc n - p))) = + real_mul y + (psum (0::nat, Suc n) + (%p::nat. real_mul (real_pow x p) (real_pow y (n - p))))" + by (import hollight POWDIFF_LEMMA) + +lemma POWDIFF: "ALL (n::nat) (x::hollight.real) y::hollight.real. + real_sub (real_pow x (Suc n)) (real_pow y (Suc n)) = + real_mul (real_sub x y) + (psum (0::nat, Suc n) + (%p::nat. real_mul (real_pow x p) (real_pow y (n - p))))" + by (import hollight POWDIFF) + +lemma POWREV: "ALL (n::nat) (x::hollight.real) y::hollight.real. + psum (0::nat, Suc n) + (%xa::nat. real_mul (real_pow x xa) (real_pow y (n - xa))) = + psum (0::nat, Suc n) + (%xa::nat. real_mul (real_pow x (n - xa)) (real_pow y xa))" + by (import hollight POWREV) + +lemma POWSER_INSIDEA: "ALL (f::nat => hollight.real) (x::hollight.real) z::hollight.real. + summable (%n::nat. real_mul (f n) (real_pow x n)) & + real_lt (real_abs z) (real_abs x) --> + summable (%n::nat. real_mul (real_abs (f n)) (real_pow z n))" + by (import hollight POWSER_INSIDEA) + +lemma POWSER_INSIDE: "ALL (f::nat => hollight.real) (x::hollight.real) z::hollight.real. + summable (%n::nat. real_mul (f n) (real_pow x n)) & + real_lt (real_abs z) (real_abs x) --> + summable (%n::nat. real_mul (f n) (real_pow z n))" + by (import hollight POWSER_INSIDE) + +constdefs + diffs :: "(nat => hollight.real) => nat => hollight.real" + "diffs == +%(u::nat => hollight.real) n::nat. + real_mul (real_of_num (Suc n)) (u (Suc n))" + +lemma DEF_diffs: "diffs = +(%(u::nat => hollight.real) n::nat. + real_mul (real_of_num (Suc n)) (u (Suc n)))" + by (import hollight DEF_diffs) + +lemma DIFFS_NEG: "ALL c::nat => hollight.real. + diffs (%n::nat. real_neg (c n)) = (%x::nat. real_neg (diffs c x))" + by (import hollight DIFFS_NEG) + +lemma DIFFS_LEMMA: "ALL (n::nat) (c::nat => hollight.real) x::hollight.real. + psum (0::nat, n) (%n::nat. real_mul (diffs c n) (real_pow x n)) = + real_add + (psum (0::nat, n) + (%n::nat. + real_mul (real_of_num n) + (real_mul (c n) (real_pow x (n - NUMERAL_BIT1 (0::nat)))))) + (real_mul (real_of_num n) + (real_mul (c n) (real_pow x (n - NUMERAL_BIT1 (0::nat)))))" + by (import hollight DIFFS_LEMMA) + +lemma DIFFS_LEMMA2: "ALL (n::nat) (c::nat => hollight.real) x::hollight.real. + psum (0::nat, n) + (%n::nat. + real_mul (real_of_num n) + (real_mul (c n) (real_pow x (n - NUMERAL_BIT1 (0::nat))))) = + real_sub + (psum (0::nat, n) (%n::nat. real_mul (diffs c n) (real_pow x n))) + (real_mul (real_of_num n) + (real_mul (c n) (real_pow x (n - NUMERAL_BIT1 (0::nat)))))" + by (import hollight DIFFS_LEMMA2) + +lemma DIFFS_EQUIV: "ALL (c::nat => hollight.real) x::hollight.real. + summable (%n::nat. real_mul (diffs c n) (real_pow x n)) --> + sums + (%n::nat. + real_mul (real_of_num n) + (real_mul (c n) (real_pow x (n - NUMERAL_BIT1 (0::nat))))) + (suminf (%n::nat. real_mul (diffs c n) (real_pow x n)))" + by (import hollight DIFFS_EQUIV) + +lemma TERMDIFF_LEMMA1: "ALL (m::nat) (z::hollight.real) h::hollight.real. + psum (0::nat, m) + (%p::nat. + real_sub (real_mul (real_pow (real_add z h) (m - p)) (real_pow z p)) + (real_pow z m)) = + psum (0::nat, m) + (%p::nat. + real_mul (real_pow z p) + (real_sub (real_pow (real_add z h) (m - p)) (real_pow z (m - p))))" + by (import hollight TERMDIFF_LEMMA1) + +lemma TERMDIFF_LEMMA2: "ALL (z::hollight.real) h::hollight.real. + h ~= real_of_num (0::nat) --> + real_sub + (real_div (real_sub (real_pow (real_add z h) (n::nat)) (real_pow z n)) + h) + (real_mul (real_of_num n) (real_pow z (n - NUMERAL_BIT1 (0::nat)))) = + real_mul h + (psum (0::nat, n - NUMERAL_BIT1 (0::nat)) + (%p::nat. + real_mul (real_pow z p) + (psum (0::nat, n - NUMERAL_BIT1 (0::nat) - p) + (%q::nat. + real_mul (real_pow (real_add z h) q) + (real_pow z + (n - NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)) - p - q))))))" + by (import hollight TERMDIFF_LEMMA2) + +lemma TERMDIFF_LEMMA3: "ALL (z::hollight.real) (h::hollight.real) (n::nat) K::hollight.real. + h ~= real_of_num (0::nat) & + real_le (real_abs z) K & real_le (real_abs (real_add z h)) K --> + real_le + (real_abs + (real_sub + (real_div (real_sub (real_pow (real_add z h) n) (real_pow z n)) h) + (real_mul (real_of_num n) + (real_pow z (n - NUMERAL_BIT1 (0::nat)))))) + (real_mul (real_of_num n) + (real_mul (real_of_num (n - NUMERAL_BIT1 (0::nat))) + (real_mul (real_pow K (n - NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) + (real_abs h))))" + by (import hollight TERMDIFF_LEMMA3) + +lemma TERMDIFF_LEMMA4: "ALL (f::hollight.real => hollight.real) (K::hollight.real) k::hollight.real. + real_lt (real_of_num (0::nat)) k & + (ALL h::hollight.real. + real_lt (real_of_num (0::nat)) (real_abs h) & + real_lt (real_abs h) k --> + real_le (real_abs (f h)) (real_mul K (real_abs h))) --> + tends_real_real f (real_of_num (0::nat)) (real_of_num (0::nat))" + by (import hollight TERMDIFF_LEMMA4) + +lemma TERMDIFF_LEMMA5: "ALL (f::nat => hollight.real) (g::hollight.real => nat => hollight.real) + k::hollight.real. + real_lt (real_of_num (0::nat)) k & + summable f & + (ALL h::hollight.real. + real_lt (real_of_num (0::nat)) (real_abs h) & + real_lt (real_abs h) k --> + (ALL n::nat. + real_le (real_abs (g h n)) (real_mul (f n) (real_abs h)))) --> + tends_real_real (%h::hollight.real. suminf (g h)) (real_of_num (0::nat)) + (real_of_num (0::nat))" + by (import hollight TERMDIFF_LEMMA5) + +lemma TERMDIFF: "ALL (c::nat => hollight.real) K::hollight.real. + summable (%n::nat. real_mul (c n) (real_pow K n)) & + summable (%n::nat. real_mul (diffs c n) (real_pow K n)) & + summable (%n::nat. real_mul (diffs (diffs c) n) (real_pow K n)) & + real_lt (real_abs (x::hollight.real)) (real_abs K) --> + diffl + (%x::hollight.real. suminf (%n::nat. real_mul (c n) (real_pow x n))) + (suminf (%n::nat. real_mul (diffs c n) (real_pow x n))) x" + by (import hollight TERMDIFF) + +lemma SEQ_NPOW: "ALL x::hollight.real. + real_lt (real_abs x) (real_of_num (NUMERAL_BIT1 (0::nat))) --> + tends_num_real (%n::nat. real_mul (real_of_num n) (real_pow x n)) + (real_of_num (0::nat))" + by (import hollight SEQ_NPOW) + +lemma TERMDIFF_CONVERGES: "ALL K::hollight.real. + (ALL x::hollight.real. + real_lt (real_abs x) K --> + summable + (%n::nat. + real_mul ((c::nat => hollight.real) n) (real_pow x n))) --> + (ALL x::hollight.real. + real_lt (real_abs x) K --> + summable (%n::nat. real_mul (diffs c n) (real_pow x n)))" + by (import hollight TERMDIFF_CONVERGES) + +lemma TERMDIFF_STRONG: "ALL (c::nat => hollight.real) (K::hollight.real) x::hollight.real. + summable (%n::nat. real_mul (c n) (real_pow K n)) & + real_lt (real_abs x) (real_abs K) --> + diffl + (%x::hollight.real. suminf (%n::nat. real_mul (c n) (real_pow x n))) + (suminf (%n::nat. real_mul (diffs c n) (real_pow x n))) x" + by (import hollight TERMDIFF_STRONG) + +lemma POWSER_0: "ALL a::nat => hollight.real. + sums (%n::nat. real_mul (a n) (real_pow (real_of_num (0::nat)) n)) + (a (0::nat))" + by (import hollight POWSER_0) + +lemma POWSER_LIMIT_0: "ALL (f::hollight.real => hollight.real) (a::nat => hollight.real) + s::hollight.real. + real_lt (real_of_num (0::nat)) s & + (ALL x::hollight.real. + real_lt (real_abs x) s --> + sums (%n::nat. real_mul (a n) (real_pow x n)) (f x)) --> + tends_real_real f (a (0::nat)) (real_of_num (0::nat))" + by (import hollight POWSER_LIMIT_0) + +lemma POWSER_LIMIT_0_STRONG: "ALL (f::hollight.real => hollight.real) (a::nat => hollight.real) + s::hollight.real. + real_lt (real_of_num (0::nat)) s & + (ALL x::hollight.real. + real_lt (real_of_num (0::nat)) (real_abs x) & + real_lt (real_abs x) s --> + sums (%n::nat. real_mul (a n) (real_pow x n)) (f x)) --> + tends_real_real f (a (0::nat)) (real_of_num (0::nat))" + by (import hollight POWSER_LIMIT_0_STRONG) + +lemma POWSER_EQUAL_0: "ALL (f::hollight.real => hollight.real) (a::nat => hollight.real) + (b::nat => hollight.real) P::hollight.real => bool. + (ALL e::hollight.real. + real_lt (real_of_num (0::nat)) e --> + (EX x::hollight.real. + P x & + real_lt (real_of_num (0::nat)) (real_abs x) & + real_lt (real_abs x) e)) & + (ALL x::hollight.real. + real_lt (real_of_num (0::nat)) (real_abs x) & P x --> + sums (%n::nat. real_mul (a n) (real_pow x n)) (f x) & + sums (%n::nat. real_mul (b n) (real_pow x n)) (f x)) --> + a (0::nat) = b (0::nat)" + by (import hollight POWSER_EQUAL_0) + +lemma POWSER_EQUAL: "ALL (f::hollight.real => hollight.real) (a::nat => hollight.real) + (b::nat => hollight.real) P::hollight.real => bool. + (ALL e::hollight.real. + real_lt (real_of_num (0::nat)) e --> + (EX x::hollight.real. + P x & + real_lt (real_of_num (0::nat)) (real_abs x) & + real_lt (real_abs x) e)) & + (ALL x::hollight.real. + P x --> + sums (%n::nat. real_mul (a n) (real_pow x n)) (f x) & + sums (%n::nat. real_mul (b n) (real_pow x n)) (f x)) --> + a = b" + by (import hollight POWSER_EQUAL) + +lemma MULT_DIV_2: "ALL n::nat. + DIV (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)) * n) + (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))) = + n" + by (import hollight MULT_DIV_2) + +lemma EVEN_DIV2: "ALL n::nat. + ~ EVEN n --> + DIV (Suc n) (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))) = + Suc (DIV (n - NUMERAL_BIT1 (0::nat)) + (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))" + by (import hollight EVEN_DIV2) + +lemma POW_ZERO: "ALL (n::nat) x::hollight.real. + real_pow x n = real_of_num (0::nat) --> x = real_of_num (0::nat)" + by (import hollight POW_ZERO) + +lemma POW_ZERO_EQ: "ALL (n::nat) x::hollight.real. + (real_pow x (Suc n) = real_of_num (0::nat)) = (x = real_of_num (0::nat))" + by (import hollight POW_ZERO_EQ) + +lemma POW_LT: "ALL (n::nat) (x::hollight.real) y::hollight.real. + real_le (real_of_num (0::nat)) x & real_lt x y --> + real_lt (real_pow x (Suc n)) (real_pow y (Suc n))" + by (import hollight POW_LT) + +lemma POW_EQ: "ALL (n::nat) (x::hollight.real) y::hollight.real. + real_le (real_of_num (0::nat)) x & + real_le (real_of_num (0::nat)) y & + real_pow x (Suc n) = real_pow y (Suc n) --> + x = y" + by (import hollight POW_EQ) + +constdefs + exp :: "hollight.real => hollight.real" + "exp == +%u::hollight.real. + suminf + (%n::nat. real_mul (real_inv (real_of_num (FACT n))) (real_pow u n))" + +lemma DEF_exp: "exp = +(%u::hollight.real. + suminf + (%n::nat. real_mul (real_inv (real_of_num (FACT n))) (real_pow u n)))" + by (import hollight DEF_exp) + +constdefs + sin :: "hollight.real => hollight.real" + "sin == +%u::hollight.real. + suminf + (%n::nat. + real_mul + (COND (EVEN n) (real_of_num (0::nat)) + (real_div + (real_pow (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) + (DIV (n - NUMERAL_BIT1 (0::nat)) + (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) + (real_of_num (FACT n)))) + (real_pow u n))" + +lemma DEF_sin: "sin = +(%u::hollight.real. + suminf + (%n::nat. + real_mul + (COND (EVEN n) (real_of_num (0::nat)) + (real_div + (real_pow (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) + (DIV (n - NUMERAL_BIT1 (0::nat)) + (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) + (real_of_num (FACT n)))) + (real_pow u n)))" + by (import hollight DEF_sin) + +constdefs + cos :: "hollight.real => hollight.real" + "cos == +%u::hollight.real. + suminf + (%n::nat. + real_mul + (COND (EVEN n) + (real_div + (real_pow (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) + (DIV n (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) + (real_of_num (FACT n))) + (real_of_num (0::nat))) + (real_pow u n))" + +lemma DEF_cos: "cos = +(%u::hollight.real. + suminf + (%n::nat. + real_mul + (COND (EVEN n) + (real_div + (real_pow (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) + (DIV n (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) + (real_of_num (FACT n))) + (real_of_num (0::nat))) + (real_pow u n)))" + by (import hollight DEF_cos) + +lemma REAL_EXP_CONVERGES: "ALL x::hollight.real. + sums (%n::nat. real_mul (real_inv (real_of_num (FACT n))) (real_pow x n)) + (exp x)" + by (import hollight REAL_EXP_CONVERGES) + +lemma SIN_CONVERGES: "ALL x::hollight.real. + sums + (%n::nat. + real_mul + (COND (EVEN n) (real_of_num (0::nat)) + (real_div + (real_pow (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) + (DIV (n - NUMERAL_BIT1 (0::nat)) + (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) + (real_of_num (FACT n)))) + (real_pow x n)) + (sin x)" + by (import hollight SIN_CONVERGES) + +lemma COS_CONVERGES: "ALL x::hollight.real. + sums + (%n::nat. + real_mul + (COND (EVEN n) + (real_div + (real_pow (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) + (DIV n (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) + (real_of_num (FACT n))) + (real_of_num (0::nat))) + (real_pow x n)) + (cos x)" + by (import hollight COS_CONVERGES) + +lemma REAL_EXP_FDIFF: "diffs (%n::nat. real_inv (real_of_num (FACT n))) = +(%n::nat. real_inv (real_of_num (FACT n)))" + by (import hollight REAL_EXP_FDIFF) + +lemma SIN_FDIFF: "diffs + (%n::nat. + COND (EVEN n) (real_of_num (0::nat)) + (real_div + (real_pow (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) + (DIV (n - NUMERAL_BIT1 (0::nat)) + (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) + (real_of_num (FACT n)))) = +(%n::nat. + COND (EVEN n) + (real_div + (real_pow (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) + (DIV n (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) + (real_of_num (FACT n))) + (real_of_num (0::nat)))" + by (import hollight SIN_FDIFF) + +lemma COS_FDIFF: "diffs + (%n::nat. + COND (EVEN n) + (real_div + (real_pow (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) + (DIV n (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) + (real_of_num (FACT n))) + (real_of_num (0::nat))) = +(%n::nat. + real_neg + (COND (EVEN n) (real_of_num (0::nat)) + (real_div + (real_pow (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) + (DIV (n - NUMERAL_BIT1 (0::nat)) + (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) + (real_of_num (FACT n)))))" + by (import hollight COS_FDIFF) + +lemma SIN_NEGLEMMA: "ALL x::hollight.real. + real_neg (sin x) = + suminf + (%n::nat. + real_neg + (real_mul + (COND (EVEN n) (real_of_num (0::nat)) + (real_div + (real_pow (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) + (DIV (n - NUMERAL_BIT1 (0::nat)) + (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) + (real_of_num (FACT n)))) + (real_pow x n)))" + by (import hollight SIN_NEGLEMMA) + +lemma DIFF_EXP: "ALL x::hollight.real. diffl exp (exp x) x" + by (import hollight DIFF_EXP) + +lemma DIFF_SIN: "ALL x::hollight.real. diffl sin (cos x) x" + by (import hollight DIFF_SIN) + +lemma DIFF_COS: "ALL x::hollight.real. diffl cos (real_neg (sin x)) x" + by (import hollight DIFF_COS) + +lemma DIFF_COMPOSITE: "(diffl (f::hollight.real => hollight.real) (l::hollight.real) + (x::hollight.real) & + f x ~= real_of_num (0::nat) --> + diffl (%x::hollight.real. real_inv (f x)) + (real_neg + (real_div l (real_pow (f x) (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))))) + x) & +(diffl f l x & + diffl (g::hollight.real => hollight.real) (m::hollight.real) x & + g x ~= real_of_num (0::nat) --> + diffl (%x::hollight.real. real_div (f x) (g x)) + (real_div (real_sub (real_mul l (g x)) (real_mul m (f x))) + (real_pow (g x) (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) + x) & +(diffl f l x & diffl g m x --> + diffl (%x::hollight.real. real_add (f x) (g x)) (real_add l m) x) & +(diffl f l x & diffl g m x --> + diffl (%x::hollight.real. real_mul (f x) (g x)) + (real_add (real_mul l (g x)) (real_mul m (f x))) x) & +(diffl f l x & diffl g m x --> + diffl (%x::hollight.real. real_sub (f x) (g x)) (real_sub l m) x) & +(diffl f l x --> diffl (%x::hollight.real. real_neg (f x)) (real_neg l) x) & +(diffl g m x --> + diffl (%x::hollight.real. real_pow (g x) (n::nat)) + (real_mul + (real_mul (real_of_num n) (real_pow (g x) (n - NUMERAL_BIT1 (0::nat)))) + m) + x) & +(diffl g m x --> + diffl (%x::hollight.real. exp (g x)) (real_mul (exp (g x)) m) x) & +(diffl g m x --> + diffl (%x::hollight.real. sin (g x)) (real_mul (cos (g x)) m) x) & +(diffl g m x --> + diffl (%x::hollight.real. cos (g x)) (real_mul (real_neg (sin (g x))) m) x)" + by (import hollight DIFF_COMPOSITE) + +lemma REAL_EXP_0: "exp (real_of_num (0::nat)) = real_of_num (NUMERAL_BIT1 (0::nat))" + by (import hollight REAL_EXP_0) + +lemma REAL_EXP_LE_X: "ALL x::hollight.real. + real_le (real_of_num (0::nat)) x --> + real_le (real_add (real_of_num (NUMERAL_BIT1 (0::nat))) x) (exp x)" + by (import hollight REAL_EXP_LE_X) + +lemma REAL_EXP_LT_1: "ALL x::hollight.real. + real_lt (real_of_num (0::nat)) x --> + real_lt (real_of_num (NUMERAL_BIT1 (0::nat))) (exp x)" + by (import hollight REAL_EXP_LT_1) + +lemma REAL_EXP_ADD_MUL: "ALL (x::hollight.real) y::hollight.real. + real_mul (exp (real_add x y)) (exp (real_neg x)) = exp y" + by (import hollight REAL_EXP_ADD_MUL) + +lemma REAL_EXP_NEG_MUL: "ALL x::hollight.real. + real_mul (exp x) (exp (real_neg x)) = real_of_num (NUMERAL_BIT1 (0::nat))" + by (import hollight REAL_EXP_NEG_MUL) + +lemma REAL_EXP_NEG_MUL2: "ALL x::hollight.real. + real_mul (exp (real_neg x)) (exp x) = real_of_num (NUMERAL_BIT1 (0::nat))" + by (import hollight REAL_EXP_NEG_MUL2) + +lemma REAL_EXP_NEG: "ALL x::hollight.real. exp (real_neg x) = real_inv (exp x)" + by (import hollight REAL_EXP_NEG) + +lemma REAL_EXP_ADD: "ALL (x::hollight.real) y::hollight.real. + exp (real_add x y) = real_mul (exp x) (exp y)" + by (import hollight REAL_EXP_ADD) + +lemma REAL_EXP_POS_LE: "ALL x::hollight.real. real_le (real_of_num (0::nat)) (exp x)" + by (import hollight REAL_EXP_POS_LE) + +lemma REAL_EXP_NZ: "ALL x::hollight.real. exp x ~= real_of_num (0::nat)" + by (import hollight REAL_EXP_NZ) + +lemma REAL_EXP_POS_LT: "ALL x::hollight.real. real_lt (real_of_num (0::nat)) (exp x)" + by (import hollight REAL_EXP_POS_LT) + +lemma REAL_EXP_N: "ALL (n::nat) x::hollight.real. + exp (real_mul (real_of_num n) x) = real_pow (exp x) n" + by (import hollight REAL_EXP_N) + +lemma REAL_EXP_SUB: "ALL (x::hollight.real) y::hollight.real. + exp (real_sub x y) = real_div (exp x) (exp y)" + by (import hollight REAL_EXP_SUB) + +lemma REAL_EXP_MONO_IMP: "ALL (x::hollight.real) y::hollight.real. + real_lt x y --> real_lt (exp x) (exp y)" + by (import hollight REAL_EXP_MONO_IMP) + +lemma REAL_EXP_MONO_LT: "ALL (x::hollight.real) y::hollight.real. + real_lt (exp x) (exp y) = real_lt x y" + by (import hollight REAL_EXP_MONO_LT) + +lemma REAL_EXP_MONO_LE: "ALL (x::hollight.real) y::hollight.real. + real_le (exp x) (exp y) = real_le x y" + by (import hollight REAL_EXP_MONO_LE) + +lemma REAL_EXP_INJ: "ALL (x::hollight.real) y::hollight.real. (exp x = exp y) = (x = y)" + by (import hollight REAL_EXP_INJ) + +lemma REAL_EXP_TOTAL_LEMMA: "ALL y::hollight.real. + real_le (real_of_num (NUMERAL_BIT1 (0::nat))) y --> + (EX x::hollight.real. + real_le (real_of_num (0::nat)) x & + real_le x (real_sub y (real_of_num (NUMERAL_BIT1 (0::nat)))) & + exp x = y)" + by (import hollight REAL_EXP_TOTAL_LEMMA) + +lemma REAL_EXP_TOTAL: "ALL y::hollight.real. + real_lt (real_of_num (0::nat)) y --> (EX x::hollight.real. exp x = y)" + by (import hollight REAL_EXP_TOTAL) + +lemma REAL_EXP_BOUND_LEMMA: "ALL x::hollight.real. + real_le (real_of_num (0::nat)) x & + real_le x + (real_inv (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) --> + real_le (exp x) + (real_add (real_of_num (NUMERAL_BIT1 (0::nat))) + (real_mul (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) x))" + by (import hollight REAL_EXP_BOUND_LEMMA) + +constdefs + ln :: "hollight.real => hollight.real" + "ln == %u::hollight.real. SOME ua::hollight.real. exp ua = u" + +lemma DEF_ln: "ln = (%u::hollight.real. SOME ua::hollight.real. exp ua = u)" + by (import hollight DEF_ln) + +lemma LN_EXP: "ALL x::hollight.real. ln (exp x) = x" + by (import hollight LN_EXP) + +lemma REAL_EXP_LN: "ALL x::hollight.real. (exp (ln x) = x) = real_lt (real_of_num (0::nat)) x" + by (import hollight REAL_EXP_LN) + +lemma LN_MUL: "ALL (x::hollight.real) y::hollight.real. + real_lt (real_of_num (0::nat)) x & real_lt (real_of_num (0::nat)) y --> + ln (real_mul x y) = real_add (ln x) (ln y)" + by (import hollight LN_MUL) + +lemma LN_INJ: "ALL (x::hollight.real) y::hollight.real. + real_lt (real_of_num (0::nat)) x & real_lt (real_of_num (0::nat)) y --> + (ln x = ln y) = (x = y)" + by (import hollight LN_INJ) + +lemma LN_1: "ln (real_of_num (NUMERAL_BIT1 (0::nat))) = real_of_num (0::nat)" + by (import hollight LN_1) + +lemma LN_INV: "ALL x::hollight.real. + real_lt (real_of_num (0::nat)) x --> ln (real_inv x) = real_neg (ln x)" + by (import hollight LN_INV) + +lemma LN_DIV: "ALL x::hollight.real. + real_lt (real_of_num (0::nat)) x & + real_lt (real_of_num (0::nat)) (y::hollight.real) --> + ln (real_div x y) = real_sub (ln x) (ln y)" + by (import hollight LN_DIV) + +lemma LN_MONO_LT: "ALL (x::hollight.real) y::hollight.real. + real_lt (real_of_num (0::nat)) x & real_lt (real_of_num (0::nat)) y --> + real_lt (ln x) (ln y) = real_lt x y" + by (import hollight LN_MONO_LT) + +lemma LN_MONO_LE: "ALL (x::hollight.real) y::hollight.real. + real_lt (real_of_num (0::nat)) x & real_lt (real_of_num (0::nat)) y --> + real_le (ln x) (ln y) = real_le x y" + by (import hollight LN_MONO_LE) + +lemma LN_POW: "ALL (n::nat) x::hollight.real. + real_lt (real_of_num (0::nat)) x --> + ln (real_pow x n) = real_mul (real_of_num n) (ln x)" + by (import hollight LN_POW) + +lemma LN_LE: "ALL x::hollight.real. + real_le (real_of_num (0::nat)) x --> + real_le (ln (real_add (real_of_num (NUMERAL_BIT1 (0::nat))) x)) x" + by (import hollight LN_LE) + +lemma LN_LT_X: "ALL x::hollight.real. real_lt (real_of_num (0::nat)) x --> real_lt (ln x) x" + by (import hollight LN_LT_X) + +lemma LN_POS: "ALL x::hollight.real. + real_le (real_of_num (NUMERAL_BIT1 (0::nat))) x --> + real_le (real_of_num (0::nat)) (ln x)" + by (import hollight LN_POS) + +lemma LN_POS_LT: "ALL x::hollight.real. + real_lt (real_of_num (NUMERAL_BIT1 (0::nat))) x --> + real_lt (real_of_num (0::nat)) (ln x)" + by (import hollight LN_POS_LT) + +lemma DIFF_LN: "ALL x::hollight.real. + real_lt (real_of_num (0::nat)) x --> diffl ln (real_inv x) x" + by (import hollight DIFF_LN) + +constdefs + root :: "nat => hollight.real => hollight.real" + "root == +%(u::nat) ua::hollight.real. + SOME ub::hollight.real. + (real_lt (real_of_num (0::nat)) ua --> + real_lt (real_of_num (0::nat)) ub) & + real_pow ub u = ua" + +lemma DEF_root: "root = +(%(u::nat) ua::hollight.real. + SOME ub::hollight.real. + (real_lt (real_of_num (0::nat)) ua --> + real_lt (real_of_num (0::nat)) ub) & + real_pow ub u = ua)" + by (import hollight DEF_root) + +constdefs + sqrt :: "hollight.real => hollight.real" + "sqrt == +%u::hollight.real. + SOME y::hollight.real. + real_le (real_of_num (0::nat)) y & + real_pow y (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))) = u" + +lemma DEF_sqrt: "sqrt = +(%u::hollight.real. + SOME y::hollight.real. + real_le (real_of_num (0::nat)) y & + real_pow y (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))) = u)" + by (import hollight DEF_sqrt) + +lemma sqrt: "sqrt (x::hollight.real) = root (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))) x" + by (import hollight sqrt) + +lemma ROOT_LT_LEMMA: "ALL (n::nat) x::hollight.real. + real_lt (real_of_num (0::nat)) x --> + real_pow (exp (real_div (ln x) (real_of_num (Suc n)))) (Suc n) = x" + by (import hollight ROOT_LT_LEMMA) + +lemma ROOT_LN: "ALL x::hollight.real. + real_lt (real_of_num (0::nat)) x --> + (ALL n::nat. + root (Suc n) x = exp (real_div (ln x) (real_of_num (Suc n))))" + by (import hollight ROOT_LN) + +lemma ROOT_0: "ALL n::nat. root (Suc n) (real_of_num (0::nat)) = real_of_num (0::nat)" + by (import hollight ROOT_0) + +lemma ROOT_1: "ALL n::nat. + root (Suc n) (real_of_num (NUMERAL_BIT1 (0::nat))) = + real_of_num (NUMERAL_BIT1 (0::nat))" + by (import hollight ROOT_1) + +lemma ROOT_POW_POS: "ALL (n::nat) x::hollight.real. + real_le (real_of_num (0::nat)) x --> + real_pow (root (Suc n) x) (Suc n) = x" + by (import hollight ROOT_POW_POS) + +lemma POW_ROOT_POS: "ALL (n::nat) x::hollight.real. + real_le (real_of_num (0::nat)) x --> + root (Suc n) (real_pow x (Suc n)) = x" + by (import hollight POW_ROOT_POS) + +lemma ROOT_POS_POSITIVE: "ALL (x::hollight.real) n::nat. + real_le (real_of_num (0::nat)) x --> + real_le (real_of_num (0::nat)) (root (Suc n) x)" + by (import hollight ROOT_POS_POSITIVE) + +lemma ROOT_POS_UNIQ: "ALL (n::nat) (x::hollight.real) y::hollight.real. + real_le (real_of_num (0::nat)) x & + real_le (real_of_num (0::nat)) y & real_pow y (Suc n) = x --> + root (Suc n) x = y" + by (import hollight ROOT_POS_UNIQ) + +lemma ROOT_MUL: "ALL (n::nat) (x::hollight.real) y::hollight.real. + real_le (real_of_num (0::nat)) x & real_le (real_of_num (0::nat)) y --> + root (Suc n) (real_mul x y) = real_mul (root (Suc n) x) (root (Suc n) y)" + by (import hollight ROOT_MUL) + +lemma ROOT_INV: "ALL (n::nat) x::hollight.real. + real_le (real_of_num (0::nat)) x --> + root (Suc n) (real_inv x) = real_inv (root (Suc n) x)" + by (import hollight ROOT_INV) + +lemma ROOT_DIV: "ALL (x::nat) (xa::hollight.real) xb::hollight.real. + real_le (real_of_num (0::nat)) xa & real_le (real_of_num (0::nat)) xb --> + root (Suc x) (real_div xa xb) = + real_div (root (Suc x) xa) (root (Suc x) xb)" + by (import hollight ROOT_DIV) + +lemma ROOT_MONO_LT: "ALL (x::hollight.real) y::hollight.real. + real_le (real_of_num (0::nat)) x & real_lt x y --> + real_lt (root (Suc (n::nat)) x) (root (Suc n) y)" + by (import hollight ROOT_MONO_LT) + +lemma ROOT_MONO_LE: "ALL (x::hollight.real) y::hollight.real. + real_le (real_of_num (0::nat)) x & real_le x y --> + real_le (root (Suc (n::nat)) x) (root (Suc n) y)" + by (import hollight ROOT_MONO_LE) + +lemma ROOT_MONO_LT_EQ: "ALL (x::hollight.real) y::hollight.real. + real_le (real_of_num (0::nat)) x & real_le (real_of_num (0::nat)) y --> + real_lt (root (Suc (n::nat)) x) (root (Suc n) y) = real_lt x y" + by (import hollight ROOT_MONO_LT_EQ) + +lemma ROOT_MONO_LE_EQ: "ALL (x::hollight.real) y::hollight.real. + real_le (real_of_num (0::nat)) x & real_le (real_of_num (0::nat)) y --> + real_le (root (Suc (n::nat)) x) (root (Suc n) y) = real_le x y" + by (import hollight ROOT_MONO_LE_EQ) + +lemma ROOT_INJ: "ALL (x::hollight.real) xa::hollight.real. + real_le (real_of_num (0::nat)) x & real_le (real_of_num (0::nat)) xa --> + (root (Suc (n::nat)) x = root (Suc n) xa) = (x = xa)" + by (import hollight ROOT_INJ) + +lemma SQRT_0: "sqrt (real_of_num (0::nat)) = real_of_num (0::nat)" + by (import hollight SQRT_0) + +lemma SQRT_1: "sqrt (real_of_num (NUMERAL_BIT1 (0::nat))) = +real_of_num (NUMERAL_BIT1 (0::nat))" + by (import hollight SQRT_1) + +lemma SQRT_POS_LT: "ALL x::hollight.real. + real_lt (real_of_num (0::nat)) x --> + real_lt (real_of_num (0::nat)) (sqrt x)" + by (import hollight SQRT_POS_LT) + +lemma SQRT_POS_LE: "ALL x::hollight.real. + real_le (real_of_num (0::nat)) x --> + real_le (real_of_num (0::nat)) (sqrt x)" + by (import hollight SQRT_POS_LE) + +lemma SQRT_POW2: "ALL x::hollight.real. + (real_pow (sqrt x) (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))) = x) = + real_le (real_of_num (0::nat)) x" + by (import hollight SQRT_POW2) + +lemma SQRT_POW_2: "ALL x::hollight.real. + real_le (real_of_num (0::nat)) x --> + real_pow (sqrt x) (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))) = x" + by (import hollight SQRT_POW_2) + +lemma POW_2_SQRT: "real_le (real_of_num (0::nat)) (x::hollight.real) --> +sqrt (real_pow x (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) = x" + by (import hollight POW_2_SQRT) + +lemma SQRT_POS_UNIQ: "ALL (x::hollight.real) xa::hollight.real. + real_le (real_of_num (0::nat)) x & + real_le (real_of_num (0::nat)) xa & + real_pow xa (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))) = x --> + sqrt x = xa" + by (import hollight SQRT_POS_UNIQ) + +lemma SQRT_MUL: "ALL (x::hollight.real) xa::hollight.real. + real_le (real_of_num (0::nat)) x & real_le (real_of_num (0::nat)) xa --> + sqrt (real_mul x xa) = real_mul (sqrt x) (sqrt xa)" + by (import hollight SQRT_MUL) + +lemma SQRT_INV: "ALL x::hollight.real. + real_le (real_of_num (0::nat)) x --> + sqrt (real_inv x) = real_inv (sqrt x)" + by (import hollight SQRT_INV) + +lemma SQRT_DIV: "ALL (x::hollight.real) xa::hollight.real. + real_le (real_of_num (0::nat)) x & real_le (real_of_num (0::nat)) xa --> + sqrt (real_div x xa) = real_div (sqrt x) (sqrt xa)" + by (import hollight SQRT_DIV) + +lemma SQRT_MONO_LT: "ALL (x::hollight.real) xa::hollight.real. + real_le (real_of_num (0::nat)) x & real_lt x xa --> + real_lt (sqrt x) (sqrt xa)" + by (import hollight SQRT_MONO_LT) + +lemma SQRT_MONO_LE: "ALL (x::hollight.real) xa::hollight.real. + real_le (real_of_num (0::nat)) x & real_le x xa --> + real_le (sqrt x) (sqrt xa)" + by (import hollight SQRT_MONO_LE) + +lemma SQRT_MONO_LT_EQ: "ALL (x::hollight.real) xa::hollight.real. + real_le (real_of_num (0::nat)) x & real_le (real_of_num (0::nat)) xa --> + real_lt (sqrt x) (sqrt xa) = real_lt x xa" + by (import hollight SQRT_MONO_LT_EQ) + +lemma SQRT_MONO_LE_EQ: "ALL (x::hollight.real) xa::hollight.real. + real_le (real_of_num (0::nat)) x & real_le (real_of_num (0::nat)) xa --> + real_le (sqrt x) (sqrt xa) = real_le x xa" + by (import hollight SQRT_MONO_LE_EQ) + +lemma SQRT_INJ: "ALL (x::hollight.real) xa::hollight.real. + real_le (real_of_num (0::nat)) x & real_le (real_of_num (0::nat)) xa --> + (sqrt x = sqrt xa) = (x = xa)" + by (import hollight SQRT_INJ) + +lemma SQRT_EVEN_POW2: "ALL n::nat. + EVEN n --> + sqrt (real_pow (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) n) = + real_pow (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) + (DIV n (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))" + by (import hollight SQRT_EVEN_POW2) + +lemma REAL_DIV_SQRT: "ALL x::hollight.real. + real_le (real_of_num (0::nat)) x --> real_div x (sqrt x) = sqrt x" + by (import hollight REAL_DIV_SQRT) + +lemma POW_2_SQRT_ABS: "ALL x::hollight.real. + sqrt (real_pow x (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) = real_abs x" + by (import hollight POW_2_SQRT_ABS) + +lemma SQRT_EQ_0: "ALL x::hollight.real. + real_le (real_of_num (0::nat)) x --> + (sqrt x = real_of_num (0::nat)) = (x = real_of_num (0::nat))" + by (import hollight SQRT_EQ_0) + +lemma REAL_LE_LSQRT: "ALL (x::hollight.real) y::hollight.real. + real_le (real_of_num (0::nat)) x & + real_le (real_of_num (0::nat)) y & + real_le x (real_pow y (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) --> + real_le (sqrt x) y" + by (import hollight REAL_LE_LSQRT) + +lemma REAL_LE_POW_2: "ALL x::hollight.real. + real_le (real_of_num (0::nat)) + (real_pow x (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))" + by (import hollight REAL_LE_POW_2) + +lemma REAL_LE_RSQRT: "ALL (x::hollight.real) y::hollight.real. + real_le (real_pow x (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) y --> + real_le x (sqrt y)" + by (import hollight REAL_LE_RSQRT) + +lemma SIN_0: "sin (real_of_num (0::nat)) = real_of_num (0::nat)" + by (import hollight SIN_0) + +lemma COS_0: "cos (real_of_num (0::nat)) = real_of_num (NUMERAL_BIT1 (0::nat))" + by (import hollight COS_0) + +lemma SIN_CIRCLE: "ALL x::hollight.real. + real_add (real_pow (sin x) (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) + (real_pow (cos x) (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) = + real_of_num (NUMERAL_BIT1 (0::nat))" + by (import hollight SIN_CIRCLE) + +lemma SIN_BOUND: "ALL x::hollight.real. + real_le (real_abs (sin x)) (real_of_num (NUMERAL_BIT1 (0::nat)))" + by (import hollight SIN_BOUND) + +lemma SIN_BOUNDS: "ALL x::hollight.real. + real_le (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) (sin x) & + real_le (sin x) (real_of_num (NUMERAL_BIT1 (0::nat)))" + by (import hollight SIN_BOUNDS) + +lemma COS_BOUND: "ALL x::hollight.real. + real_le (real_abs (cos x)) (real_of_num (NUMERAL_BIT1 (0::nat)))" + by (import hollight COS_BOUND) + +lemma COS_BOUNDS: "ALL x::hollight.real. + real_le (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) (cos x) & + real_le (cos x) (real_of_num (NUMERAL_BIT1 (0::nat)))" + by (import hollight COS_BOUNDS) + +lemma SIN_COS_ADD: "ALL (x::hollight.real) y::hollight.real. + real_add + (real_pow + (real_sub (sin (real_add x y)) + (real_add (real_mul (sin x) (cos y)) (real_mul (cos x) (sin y)))) + (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) + (real_pow + (real_sub (cos (real_add x y)) + (real_sub (real_mul (cos x) (cos y)) (real_mul (sin x) (sin y)))) + (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) = + real_of_num (0::nat)" + by (import hollight SIN_COS_ADD) + +lemma SIN_COS_NEG: "ALL x::hollight.real. + real_add + (real_pow (real_add (sin (real_neg x)) (sin x)) + (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) + (real_pow (real_sub (cos (real_neg x)) (cos x)) + (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) = + real_of_num (0::nat)" + by (import hollight SIN_COS_NEG) + +lemma SIN_ADD: "ALL (x::hollight.real) y::hollight.real. + sin (real_add x y) = + real_add (real_mul (sin x) (cos y)) (real_mul (cos x) (sin y))" + by (import hollight SIN_ADD) + +lemma COS_ADD: "ALL (x::hollight.real) y::hollight.real. + cos (real_add x y) = + real_sub (real_mul (cos x) (cos y)) (real_mul (sin x) (sin y))" + by (import hollight COS_ADD) + +lemma SIN_NEG: "ALL x::hollight.real. sin (real_neg x) = real_neg (sin x)" + by (import hollight SIN_NEG) + +lemma COS_NEG: "ALL x::hollight.real. cos (real_neg x) = cos x" + by (import hollight COS_NEG) + +lemma SIN_DOUBLE: "ALL x::hollight.real. + sin (real_mul (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) x) = + real_mul (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) + (real_mul (sin x) (cos x))" + by (import hollight SIN_DOUBLE) + +lemma COS_DOUBLE: "ALL x::hollight.real. + cos (real_mul (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) x) = + real_sub (real_pow (cos x) (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) + (real_pow (sin x) (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))" + by (import hollight COS_DOUBLE) + +lemma COS_ABS: "ALL x::hollight.real. cos (real_abs x) = cos x" + by (import hollight COS_ABS) + +lemma SIN_PAIRED: "ALL x::hollight.real. + sums + (%n::nat. + real_mul + (real_div + (real_pow (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) n) + (real_of_num + (FACT + (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)) * n + + NUMERAL_BIT1 (0::nat))))) + (real_pow x + (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)) * n + + NUMERAL_BIT1 (0::nat)))) + (sin x)" + by (import hollight SIN_PAIRED) + +lemma SIN_POS: "ALL x::hollight.real. + real_lt (real_of_num (0::nat)) x & + real_lt x (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) --> + real_lt (real_of_num (0::nat)) (sin x)" + by (import hollight SIN_POS) + +lemma COS_PAIRED: "ALL x::hollight.real. + sums + (%n::nat. + real_mul + (real_div + (real_pow (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) n) + (real_of_num (FACT (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)) * n)))) + (real_pow x (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)) * n))) + (cos x)" + by (import hollight COS_PAIRED) + +lemma COS_2: "real_lt (cos (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) + (real_of_num (0::nat))" + by (import hollight COS_2) + +lemma COS_ISZERO: "EX! x::hollight.real. + real_le (real_of_num (0::nat)) x & + real_le x (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) & + cos x = real_of_num (0::nat)" + by (import hollight COS_ISZERO) + +constdefs + pi :: "hollight.real" + "pi == +real_mul (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) + (SOME x::hollight.real. + real_le (real_of_num (0::nat)) x & + real_le x (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) & + cos x = real_of_num (0::nat))" + +lemma DEF_pi: "pi = +real_mul (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) + (SOME x::hollight.real. + real_le (real_of_num (0::nat)) x & + real_le x (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) & + cos x = real_of_num (0::nat))" + by (import hollight DEF_pi) + +lemma PI2: "real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) = +(SOME x::hollight.real. + real_le (real_of_num (0::nat)) x & + real_le x (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) & + cos x = real_of_num (0::nat))" + by (import hollight PI2) + +lemma COS_PI2: "cos (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) = +real_of_num (0::nat)" + by (import hollight COS_PI2) + +lemma PI2_BOUNDS: "real_lt (real_of_num (0::nat)) + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) & +real_lt (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) + (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))" + by (import hollight PI2_BOUNDS) + +lemma PI_POS: "real_lt (real_of_num (0::nat)) pi" + by (import hollight PI_POS) + +lemma SIN_PI2: "sin (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) = +real_of_num (NUMERAL_BIT1 (0::nat))" + by (import hollight SIN_PI2) + +lemma COS_PI: "cos pi = real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))" + by (import hollight COS_PI) + +lemma SIN_PI: "sin pi = real_of_num (0::nat)" + by (import hollight SIN_PI) + +lemma SIN_COS: "ALL x::hollight.real. + sin x = + cos (real_sub + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) + x)" + by (import hollight SIN_COS) + +lemma COS_SIN: "ALL x::hollight.real. + cos x = + sin (real_sub + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) + x)" + by (import hollight COS_SIN) + +lemma SIN_PERIODIC_PI: "ALL x::hollight.real. sin (real_add x pi) = real_neg (sin x)" + by (import hollight SIN_PERIODIC_PI) + +lemma COS_PERIODIC_PI: "ALL x::hollight.real. cos (real_add x pi) = real_neg (cos x)" + by (import hollight COS_PERIODIC_PI) + +lemma SIN_PERIODIC: "ALL x::hollight.real. + sin (real_add x + (real_mul (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) + pi)) = + sin x" + by (import hollight SIN_PERIODIC) + +lemma COS_PERIODIC: "ALL x::hollight.real. + cos (real_add x + (real_mul (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) + pi)) = + cos x" + by (import hollight COS_PERIODIC) + +lemma COS_NPI: "ALL n::nat. + cos (real_mul (real_of_num n) pi) = + real_pow (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) n" + by (import hollight COS_NPI) + +lemma SIN_NPI: "ALL n::nat. sin (real_mul (real_of_num n) pi) = real_of_num (0::nat)" + by (import hollight SIN_NPI) + +lemma SIN_POS_PI2: "ALL x::hollight.real. + real_lt (real_of_num (0::nat)) x & + real_lt x + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) --> + real_lt (real_of_num (0::nat)) (sin x)" + by (import hollight SIN_POS_PI2) + +lemma COS_POS_PI2: "ALL x::hollight.real. + real_lt (real_of_num (0::nat)) x & + real_lt x + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) --> + real_lt (real_of_num (0::nat)) (cos x)" + by (import hollight COS_POS_PI2) + +lemma COS_POS_PI: "ALL x::hollight.real. + real_lt + (real_neg + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))))) + x & + real_lt x + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) --> + real_lt (real_of_num (0::nat)) (cos x)" + by (import hollight COS_POS_PI) + +lemma SIN_POS_PI: "ALL x::hollight.real. + real_lt (real_of_num (0::nat)) x & real_lt x pi --> + real_lt (real_of_num (0::nat)) (sin x)" + by (import hollight SIN_POS_PI) + +lemma SIN_POS_PI_LE: "ALL x::hollight.real. + real_le (real_of_num (0::nat)) x & real_le x pi --> + real_le (real_of_num (0::nat)) (sin x)" + by (import hollight SIN_POS_PI_LE) + +lemma COS_TOTAL: "ALL y::hollight.real. + real_le (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) y & + real_le y (real_of_num (NUMERAL_BIT1 (0::nat))) --> + (EX! x::hollight.real. + real_le (real_of_num (0::nat)) x & real_le x pi & cos x = y)" + by (import hollight COS_TOTAL) + +lemma SIN_TOTAL: "ALL y::hollight.real. + real_le (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) y & + real_le y (real_of_num (NUMERAL_BIT1 (0::nat))) --> + (EX! x::hollight.real. + real_le + (real_neg + (real_div pi + (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))))) + x & + real_le x + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) & + sin x = y)" + by (import hollight SIN_TOTAL) + +lemma COS_ZERO_LEMMA: "ALL x::hollight.real. + real_le (real_of_num (0::nat)) x & cos x = real_of_num (0::nat) --> + (EX n::nat. + ~ EVEN n & + x = + real_mul (real_of_num n) + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))))" + by (import hollight COS_ZERO_LEMMA) + +lemma SIN_ZERO_LEMMA: "ALL x::hollight.real. + real_le (real_of_num (0::nat)) x & sin x = real_of_num (0::nat) --> + (EX n::nat. + EVEN n & + x = + real_mul (real_of_num n) + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))))" + by (import hollight SIN_ZERO_LEMMA) + +lemma COS_ZERO: "ALL x::hollight.real. + (cos x = real_of_num (0::nat)) = + ((EX n::nat. + ~ EVEN n & + x = + real_mul (real_of_num n) + (real_div pi + (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))))) | + (EX n::nat. + ~ EVEN n & + x = + real_neg + (real_mul (real_of_num n) + (real_div pi + (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))))))" + by (import hollight COS_ZERO) + +lemma SIN_ZERO: "ALL x::hollight.real. + (sin x = real_of_num (0::nat)) = + ((EX n::nat. + EVEN n & + x = + real_mul (real_of_num n) + (real_div pi + (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))))) | + (EX n::nat. + EVEN n & + x = + real_neg + (real_mul (real_of_num n) + (real_div pi + (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))))))" + by (import hollight SIN_ZERO) + +lemma SIN_ZERO_PI: "ALL x::hollight.real. + (sin x = real_of_num (0::nat)) = + ((EX n::nat. x = real_mul (real_of_num n) pi) | + (EX n::nat. x = real_neg (real_mul (real_of_num n) pi)))" + by (import hollight SIN_ZERO_PI) + +lemma COS_ONE_2PI: "ALL x::hollight.real. + (cos x = real_of_num (NUMERAL_BIT1 (0::nat))) = + ((EX n::nat. + x = + real_mul (real_of_num n) + (real_mul (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) + pi)) | + (EX n::nat. + x = + real_neg + (real_mul (real_of_num n) + (real_mul (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) + pi))))" + by (import hollight COS_ONE_2PI) + +constdefs + tan :: "hollight.real => hollight.real" + "tan == %u::hollight.real. real_div (sin u) (cos u)" + +lemma DEF_tan: "tan = (%u::hollight.real. real_div (sin u) (cos u))" + by (import hollight DEF_tan) + +lemma TAN_0: "tan (real_of_num (0::nat)) = real_of_num (0::nat)" + by (import hollight TAN_0) + +lemma TAN_PI: "tan pi = real_of_num (0::nat)" + by (import hollight TAN_PI) + +lemma TAN_NPI: "ALL n::nat. tan (real_mul (real_of_num n) pi) = real_of_num (0::nat)" + by (import hollight TAN_NPI) + +lemma TAN_NEG: "ALL x::hollight.real. tan (real_neg x) = real_neg (tan x)" + by (import hollight TAN_NEG) + +lemma TAN_PERIODIC: "ALL x::hollight.real. + tan (real_add x + (real_mul (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) + pi)) = + tan x" + by (import hollight TAN_PERIODIC) + +lemma TAN_PERIODIC_PI: "ALL x::hollight.real. tan (real_add x pi) = tan x" + by (import hollight TAN_PERIODIC_PI) + +lemma TAN_PERIODIC_NPI: "ALL (x::hollight.real) n::nat. + tan (real_add x (real_mul (real_of_num n) pi)) = tan x" + by (import hollight TAN_PERIODIC_NPI) + +lemma TAN_ADD: "ALL (x::hollight.real) y::hollight.real. + cos x ~= real_of_num (0::nat) & + cos y ~= real_of_num (0::nat) & + cos (real_add x y) ~= real_of_num (0::nat) --> + tan (real_add x y) = + real_div (real_add (tan x) (tan y)) + (real_sub (real_of_num (NUMERAL_BIT1 (0::nat))) + (real_mul (tan x) (tan y)))" + by (import hollight TAN_ADD) + +lemma TAN_DOUBLE: "ALL x::hollight.real. + cos x ~= real_of_num (0::nat) & + cos (real_mul (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) x) ~= + real_of_num (0::nat) --> + tan (real_mul (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) x) = + real_div + (real_mul (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) (tan x)) + (real_sub (real_of_num (NUMERAL_BIT1 (0::nat))) + (real_pow (tan x) (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))))" + by (import hollight TAN_DOUBLE) + +lemma TAN_POS_PI2: "ALL x::hollight.real. + real_lt (real_of_num (0::nat)) x & + real_lt x + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) --> + real_lt (real_of_num (0::nat)) (tan x)" + by (import hollight TAN_POS_PI2) + +lemma DIFF_TAN: "ALL x::hollight.real. + cos x ~= real_of_num (0::nat) --> + diffl tan + (real_inv (real_pow (cos x) (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) x" + by (import hollight DIFF_TAN) + +lemma DIFF_TAN_COMPOSITE: "diffl (g::hollight.real => hollight.real) (m::hollight.real) + (x::hollight.real) & +cos (g x) ~= real_of_num (0::nat) --> +diffl (%x::hollight.real. tan (g x)) + (real_mul + (real_inv (real_pow (cos (g x)) (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) + m) + x" + by (import hollight DIFF_TAN_COMPOSITE) + +lemma TAN_TOTAL_LEMMA: "ALL y::hollight.real. + real_lt (real_of_num (0::nat)) y --> + (EX x::hollight.real. + real_lt (real_of_num (0::nat)) x & + real_lt x + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) & + real_lt y (tan x))" + by (import hollight TAN_TOTAL_LEMMA) + +lemma TAN_TOTAL_POS: "ALL y::hollight.real. + real_le (real_of_num (0::nat)) y --> + (EX x::hollight.real. + real_le (real_of_num (0::nat)) x & + real_lt x + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) & + tan x = y)" + by (import hollight TAN_TOTAL_POS) + +lemma TAN_TOTAL: "ALL y::hollight.real. + EX! x::hollight.real. + real_lt + (real_neg + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))))) + x & + real_lt x + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) & + tan x = y" + by (import hollight TAN_TOTAL) + +lemma PI2_PI4: "real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) = +real_mul (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) + (real_div pi + (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))))" + by (import hollight PI2_PI4) + +lemma TAN_PI4: "tan (real_div pi + (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))))) = +real_of_num (NUMERAL_BIT1 (0::nat))" + by (import hollight TAN_PI4) + +lemma TAN_COT: "ALL x::hollight.real. + tan (real_sub + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) + x) = + real_inv (tan x)" + by (import hollight TAN_COT) + +lemma TAN_BOUND_PI2: "ALL x::hollight.real. + real_lt (real_abs x) + (real_div pi + (real_of_num + (NUMERAL_BIT0 (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))))) --> + real_lt (real_abs (tan x)) (real_of_num (NUMERAL_BIT1 (0::nat)))" + by (import hollight TAN_BOUND_PI2) + +lemma TAN_ABS_GE_X: "ALL x::hollight.real. + real_lt (real_abs x) + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) --> + real_le (real_abs x) (real_abs (tan x))" + by (import hollight TAN_ABS_GE_X) + +constdefs + asn :: "hollight.real => hollight.real" + "asn == +%u::hollight.real. + SOME x::hollight.real. + real_le + (real_neg + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))))) + x & + real_le x + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) & + sin x = u" + +lemma DEF_asn: "asn = +(%u::hollight.real. + SOME x::hollight.real. + real_le + (real_neg + (real_div pi + (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))))) + x & + real_le x + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) & + sin x = u)" + by (import hollight DEF_asn) + +constdefs + acs :: "hollight.real => hollight.real" + "acs == +%u::hollight.real. + SOME x::hollight.real. + real_le (real_of_num (0::nat)) x & real_le x pi & cos x = u" + +lemma DEF_acs: "acs = +(%u::hollight.real. + SOME x::hollight.real. + real_le (real_of_num (0::nat)) x & real_le x pi & cos x = u)" + by (import hollight DEF_acs) + +constdefs + atn :: "hollight.real => hollight.real" + "atn == +%u::hollight.real. + SOME x::hollight.real. + real_lt + (real_neg + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))))) + x & + real_lt x + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) & + tan x = u" + +lemma DEF_atn: "atn = +(%u::hollight.real. + SOME x::hollight.real. + real_lt + (real_neg + (real_div pi + (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))))) + x & + real_lt x + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) & + tan x = u)" + by (import hollight DEF_atn) + +lemma ASN: "ALL y::hollight.real. + real_le (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) y & + real_le y (real_of_num (NUMERAL_BIT1 (0::nat))) --> + real_le + (real_neg + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))))) + (asn y) & + real_le (asn y) + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) & + sin (asn y) = y" + by (import hollight ASN) + +lemma ASN_SIN: "ALL y::hollight.real. + real_le (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) y & + real_le y (real_of_num (NUMERAL_BIT1 (0::nat))) --> + sin (asn y) = y" + by (import hollight ASN_SIN) + +lemma ASN_BOUNDS: "ALL y::hollight.real. + real_le (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) y & + real_le y (real_of_num (NUMERAL_BIT1 (0::nat))) --> + real_le + (real_neg + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))))) + (asn y) & + real_le (asn y) + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))))" + by (import hollight ASN_BOUNDS) + +lemma ASN_BOUNDS_LT: "ALL y::hollight.real. + real_lt (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) y & + real_lt y (real_of_num (NUMERAL_BIT1 (0::nat))) --> + real_lt + (real_neg + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))))) + (asn y) & + real_lt (asn y) + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))))" + by (import hollight ASN_BOUNDS_LT) + +lemma SIN_ASN: "ALL x::hollight.real. + real_le + (real_neg + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))))) + x & + real_le x + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) --> + asn (sin x) = x" + by (import hollight SIN_ASN) + +lemma ACS: "ALL y::hollight.real. + real_le (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) y & + real_le y (real_of_num (NUMERAL_BIT1 (0::nat))) --> + real_le (real_of_num (0::nat)) (acs y) & + real_le (acs y) pi & cos (acs y) = y" + by (import hollight ACS) + +lemma ACS_COS: "ALL y::hollight.real. + real_le (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) y & + real_le y (real_of_num (NUMERAL_BIT1 (0::nat))) --> + cos (acs y) = y" + by (import hollight ACS_COS) + +lemma ACS_BOUNDS: "ALL y::hollight.real. + real_le (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) y & + real_le y (real_of_num (NUMERAL_BIT1 (0::nat))) --> + real_le (real_of_num (0::nat)) (acs y) & real_le (acs y) pi" + by (import hollight ACS_BOUNDS) + +lemma ACS_BOUNDS_LT: "ALL y::hollight.real. + real_lt (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) y & + real_lt y (real_of_num (NUMERAL_BIT1 (0::nat))) --> + real_lt (real_of_num (0::nat)) (acs y) & real_lt (acs y) pi" + by (import hollight ACS_BOUNDS_LT) + +lemma COS_ACS: "ALL x::hollight.real. + real_le (real_of_num (0::nat)) x & real_le x pi --> acs (cos x) = x" + by (import hollight COS_ACS) + +lemma ATN: "ALL y::hollight.real. + real_lt + (real_neg + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))))) + (atn y) & + real_lt (atn y) + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) & + tan (atn y) = y" + by (import hollight ATN) + +lemma ATN_TAN: "ALL x::hollight.real. tan (atn x) = x" + by (import hollight ATN_TAN) + +lemma ATN_BOUNDS: "ALL x::hollight.real. + real_lt + (real_neg + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))))) + (atn x) & + real_lt (atn x) + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))))" + by (import hollight ATN_BOUNDS) + +lemma TAN_ATN: "ALL x::hollight.real. + real_lt + (real_neg + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))))) + x & + real_lt x + (real_div pi (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) --> + atn (tan x) = x" + by (import hollight TAN_ATN) + +lemma ATN_0: "atn (real_of_num (0::nat)) = real_of_num (0::nat)" + by (import hollight ATN_0) + +lemma ATN_1: "atn (real_of_num (NUMERAL_BIT1 (0::nat))) = +real_div pi + (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))))" + by (import hollight ATN_1) + +lemma ATN_NEG: "ALL x::hollight.real. atn (real_neg x) = real_neg (atn x)" + by (import hollight ATN_NEG) + +lemma COS_ATN_NZ: "ALL x::hollight.real. cos (atn x) ~= real_of_num (0::nat)" + by (import hollight COS_ATN_NZ) + +lemma TAN_SEC: "ALL x::hollight.real. + cos x ~= real_of_num (0::nat) --> + real_add (real_of_num (NUMERAL_BIT1 (0::nat))) + (real_pow (tan x) (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) = + real_pow (real_inv (cos x)) (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))" + by (import hollight TAN_SEC) + +lemma DIFF_ATN: "ALL x::hollight.real. + diffl atn + (real_inv + (real_add (real_of_num (NUMERAL_BIT1 (0::nat))) + (real_pow x (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))))) + x" + by (import hollight DIFF_ATN) + +lemma DIFF_ATN_COMPOSITE: "diffl (g::hollight.real => hollight.real) (m::hollight.real) + (x::hollight.real) --> +diffl (%x::hollight.real. atn (g x)) + (real_mul + (real_inv + (real_add (real_of_num (NUMERAL_BIT1 (0::nat))) + (real_pow (g x) (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))))) + m) + x" + by (import hollight DIFF_ATN_COMPOSITE) + +lemma ATN_MONO_LT: "ALL (x::hollight.real) y::hollight.real. + real_lt x y --> real_lt (atn x) (atn y)" + by (import hollight ATN_MONO_LT) + +lemma ATN_MONO_LT_EQ: "ALL (x::hollight.real) y::hollight.real. + real_lt (atn x) (atn y) = real_lt x y" + by (import hollight ATN_MONO_LT_EQ) + +lemma ATN_MONO_LE_EQ: "ALL (x::hollight.real) xa::hollight.real. + real_le (atn x) (atn xa) = real_le x xa" + by (import hollight ATN_MONO_LE_EQ) + +lemma ATN_INJ: "ALL (x::hollight.real) xa::hollight.real. (atn x = atn xa) = (x = xa)" + by (import hollight ATN_INJ) + +lemma ATN_POS_LT: "real_lt (real_of_num (0::nat)) (atn (x::hollight.real)) = +real_lt (real_of_num (0::nat)) x" + by (import hollight ATN_POS_LT) + +lemma ATN_POS_LE: "real_le (real_of_num (0::nat)) (atn (x::hollight.real)) = +real_le (real_of_num (0::nat)) x" + by (import hollight ATN_POS_LE) + +lemma ATN_LT_PI4_POS: "ALL x::hollight.real. + real_lt x (real_of_num (NUMERAL_BIT1 (0::nat))) --> + real_lt (atn x) + (real_div pi + (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))))" + by (import hollight ATN_LT_PI4_POS) + +lemma ATN_LT_PI4_NEG: "ALL x::hollight.real. + real_lt (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) x --> + real_lt + (real_neg + (real_div pi + (real_of_num + (NUMERAL_BIT0 (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))))) + (atn x)" + by (import hollight ATN_LT_PI4_NEG) + +lemma ATN_LT_PI4: "ALL x::hollight.real. + real_lt (real_abs x) (real_of_num (NUMERAL_BIT1 (0::nat))) --> + real_lt (real_abs (atn x)) + (real_div pi + (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))))" + by (import hollight ATN_LT_PI4) + +lemma ATN_LE_PI4: "ALL x::hollight.real. + real_le (real_abs x) (real_of_num (NUMERAL_BIT1 (0::nat))) --> + real_le (real_abs (atn x)) + (real_div pi + (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))))" + by (import hollight ATN_LE_PI4) + +lemma COS_SIN_SQRT: "ALL x::hollight.real. + real_le (real_of_num (0::nat)) (cos x) --> + cos x = + sqrt + (real_sub (real_of_num (NUMERAL_BIT1 (0::nat))) + (real_pow (sin x) (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))))" + by (import hollight COS_SIN_SQRT) + +lemma COS_ASN_NZ: "ALL x::hollight.real. + real_lt (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) x & + real_lt x (real_of_num (NUMERAL_BIT1 (0::nat))) --> + cos (asn x) ~= real_of_num (0::nat)" + by (import hollight COS_ASN_NZ) + +lemma DIFF_ASN_COS: "ALL x::hollight.real. + real_lt (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) x & + real_lt x (real_of_num (NUMERAL_BIT1 (0::nat))) --> + diffl asn (real_inv (cos (asn x))) x" + by (import hollight DIFF_ASN_COS) + +lemma DIFF_ASN: "ALL x::hollight.real. + real_lt (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) x & + real_lt x (real_of_num (NUMERAL_BIT1 (0::nat))) --> + diffl asn + (real_inv + (sqrt + (real_sub (real_of_num (NUMERAL_BIT1 (0::nat))) + (real_pow x (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))))) + x" + by (import hollight DIFF_ASN) + +lemma SIN_COS_SQRT: "ALL x::hollight.real. + real_le (real_of_num (0::nat)) (sin x) --> + sin x = + sqrt + (real_sub (real_of_num (NUMERAL_BIT1 (0::nat))) + (real_pow (cos x) (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))))" + by (import hollight SIN_COS_SQRT) + +lemma SIN_ACS_NZ: "ALL x::hollight.real. + real_lt (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) x & + real_lt x (real_of_num (NUMERAL_BIT1 (0::nat))) --> + sin (acs x) ~= real_of_num (0::nat)" + by (import hollight SIN_ACS_NZ) + +lemma DIFF_ACS_SIN: "ALL x::hollight.real. + real_lt (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) x & + real_lt x (real_of_num (NUMERAL_BIT1 (0::nat))) --> + diffl acs (real_inv (real_neg (sin (acs x)))) x" + by (import hollight DIFF_ACS_SIN) + +lemma DIFF_ACS: "ALL x::hollight.real. + real_lt (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) x & + real_lt x (real_of_num (NUMERAL_BIT1 (0::nat))) --> + diffl acs + (real_neg + (real_inv + (sqrt + (real_sub (real_of_num (NUMERAL_BIT1 (0::nat))) + (real_pow x (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))))))) + x" + by (import hollight DIFF_ACS) + +lemma CIRCLE_SINCOS: "ALL (x::hollight.real) y::hollight.real. + real_add (real_pow x (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) + (real_pow y (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))) = + real_of_num (NUMERAL_BIT1 (0::nat)) --> + (EX t::hollight.real. x = cos t & y = sin t)" + by (import hollight CIRCLE_SINCOS) + +lemma ACS_MONO_LT: "ALL (x::hollight.real) y::hollight.real. + real_lt (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) x & + real_lt x y & real_lt y (real_of_num (NUMERAL_BIT1 (0::nat))) --> + real_lt (acs y) (acs x)" + by (import hollight ACS_MONO_LT) + +lemma LESS_SUC_EQ: "ALL (m::nat) n::nat. < m (Suc n) = <= m n" + by (import hollight LESS_SUC_EQ) + +lemma LESS_1: "ALL x::nat. < x (NUMERAL_BIT1 (0::nat)) = (x = (0::nat))" + by (import hollight LESS_1) + +constdefs + division :: "hollight.real * hollight.real => (nat => hollight.real) => bool" + "division == +%(u::hollight.real * hollight.real) ua::nat => hollight.real. + ua (0::nat) = fst u & + (EX N::nat. + (ALL n::nat. < n N --> real_lt (ua n) (ua (Suc n))) & + (ALL n::nat. >= n N --> ua n = snd u))" + +lemma DEF_division: "division = +(%(u::hollight.real * hollight.real) ua::nat => hollight.real. + ua (0::nat) = fst u & + (EX N::nat. + (ALL n::nat. < n N --> real_lt (ua n) (ua (Suc n))) & + (ALL n::nat. >= n N --> ua n = snd u)))" + by (import hollight DEF_division) + +constdefs + dsize :: "(nat => hollight.real) => nat" + "dsize == +%u::nat => hollight.real. + SOME N::nat. + (ALL n::nat. < n N --> real_lt (u n) (u (Suc n))) & + (ALL n::nat. >= n N --> u n = u N)" + +lemma DEF_dsize: "dsize = +(%u::nat => hollight.real. + SOME N::nat. + (ALL n::nat. < n N --> real_lt (u n) (u (Suc n))) & + (ALL n::nat. >= n N --> u n = u N))" + by (import hollight DEF_dsize) + +constdefs + tdiv :: "hollight.real * hollight.real +=> (nat => hollight.real) * (nat => hollight.real) => bool" + "tdiv == +%(u::hollight.real * hollight.real) + ua::(nat => hollight.real) * (nat => hollight.real). + division (fst u, snd u) (fst ua) & + (ALL n::nat. + real_le (fst ua n) (snd ua n) & real_le (snd ua n) (fst ua (Suc n)))" + +lemma DEF_tdiv: "tdiv = +(%(u::hollight.real * hollight.real) + ua::(nat => hollight.real) * (nat => hollight.real). + division (fst u, snd u) (fst ua) & + (ALL n::nat. + real_le (fst ua n) (snd ua n) & + real_le (snd ua n) (fst ua (Suc n))))" + by (import hollight DEF_tdiv) + +constdefs + gauge :: "(hollight.real => bool) => (hollight.real => hollight.real) => bool" + "gauge == +%(u::hollight.real => bool) ua::hollight.real => hollight.real. + ALL x::hollight.real. u x --> real_lt (real_of_num (0::nat)) (ua x)" + +lemma DEF_gauge: "gauge = +(%(u::hollight.real => bool) ua::hollight.real => hollight.real. + ALL x::hollight.real. u x --> real_lt (real_of_num (0::nat)) (ua x))" + by (import hollight DEF_gauge) + +constdefs + fine :: "(hollight.real => hollight.real) +=> (nat => hollight.real) * (nat => hollight.real) => bool" + "fine == +%(u::hollight.real => hollight.real) + ua::(nat => hollight.real) * (nat => hollight.real). + ALL n::nat. + < n (dsize (fst ua)) --> + real_lt (real_sub (fst ua (Suc n)) (fst ua n)) (u (snd ua n))" + +lemma DEF_fine: "fine = +(%(u::hollight.real => hollight.real) + ua::(nat => hollight.real) * (nat => hollight.real). + ALL n::nat. + < n (dsize (fst ua)) --> + real_lt (real_sub (fst ua (Suc n)) (fst ua n)) (u (snd ua n)))" + by (import hollight DEF_fine) + +constdefs + rsum :: "(nat => hollight.real) * (nat => hollight.real) +=> (hollight.real => hollight.real) => hollight.real" + "rsum == +%(u::(nat => hollight.real) * (nat => hollight.real)) + ua::hollight.real => hollight.real. + psum (0::nat, dsize (fst u)) + (%n::nat. real_mul (ua (snd u n)) (real_sub (fst u (Suc n)) (fst u n)))" + +lemma DEF_rsum: "rsum = +(%(u::(nat => hollight.real) * (nat => hollight.real)) + ua::hollight.real => hollight.real. + psum (0::nat, dsize (fst u)) + (%n::nat. + real_mul (ua (snd u n)) (real_sub (fst u (Suc n)) (fst u n))))" + by (import hollight DEF_rsum) + +constdefs + defint :: "hollight.real * hollight.real +=> (hollight.real => hollight.real) => hollight.real => bool" + "defint == +%(u::hollight.real * hollight.real) (ua::hollight.real => hollight.real) + ub::hollight.real. + ALL e::hollight.real. + real_lt (real_of_num (0::nat)) e --> + (EX g::hollight.real => hollight.real. + gauge (%x::hollight.real. real_le (fst u) x & real_le x (snd u)) + g & + (ALL (D::nat => hollight.real) p::nat => hollight.real. + tdiv (fst u, snd u) (D, p) & fine g (D, p) --> + real_lt (real_abs (real_sub (rsum (D, p) ua) ub)) e))" + +lemma DEF_defint: "defint = +(%(u::hollight.real * hollight.real) (ua::hollight.real => hollight.real) + ub::hollight.real. + ALL e::hollight.real. + real_lt (real_of_num (0::nat)) e --> + (EX g::hollight.real => hollight.real. + gauge (%x::hollight.real. real_le (fst u) x & real_le x (snd u)) + g & + (ALL (D::nat => hollight.real) p::nat => hollight.real. + tdiv (fst u, snd u) (D, p) & fine g (D, p) --> + real_lt (real_abs (real_sub (rsum (D, p) ua) ub)) e)))" + by (import hollight DEF_defint) + +lemma DIVISION_0: "ALL (a::hollight.real) b::hollight.real. + a = b --> dsize (%n::nat. COND (n = (0::nat)) a b) = (0::nat)" + by (import hollight DIVISION_0) + +lemma DIVISION_1: "ALL (a::hollight.real) b::hollight.real. + real_lt a b --> + dsize (%n::nat. COND (n = (0::nat)) a b) = NUMERAL_BIT1 (0::nat)" + by (import hollight DIVISION_1) + +lemma DIVISION_SINGLE: "ALL (a::hollight.real) b::hollight.real. + real_le a b --> division (a, b) (%n::nat. COND (n = (0::nat)) a b)" + by (import hollight DIVISION_SINGLE) + +lemma DIVISION_LHS: "ALL (D::nat => hollight.real) (a::hollight.real) b::hollight.real. + division (a, b) D --> D (0::nat) = a" + by (import hollight DIVISION_LHS) + +lemma DIVISION_THM: "ALL (D::nat => hollight.real) (a::hollight.real) b::hollight.real. + division (a, b) D = + (D (0::nat) = a & + (ALL n::nat. < n (dsize D) --> real_lt (D n) (D (Suc n))) & + (ALL n::nat. >= n (dsize D) --> D n = b))" + by (import hollight DIVISION_THM) + +lemma DIVISION_RHS: "ALL (D::nat => hollight.real) (a::hollight.real) b::hollight.real. + division (a, b) D --> D (dsize D) = b" + by (import hollight DIVISION_RHS) + +lemma DIVISION_LT_GEN: "ALL (D::nat => hollight.real) (a::hollight.real) (b::hollight.real) (m::nat) + n::nat. + division (a, b) D & < m n & <= n (dsize D) --> real_lt (D m) (D n)" + by (import hollight DIVISION_LT_GEN) + +lemma DIVISION_LT: "ALL (D::nat => hollight.real) (a::hollight.real) b::hollight.real. + division (a, b) D --> + (ALL n::nat. < n (dsize D) --> real_lt (D (0::nat)) (D (Suc n)))" + by (import hollight DIVISION_LT) + +lemma DIVISION_LE: "ALL (D::nat => hollight.real) (a::hollight.real) b::hollight.real. + division (a, b) D --> real_le a b" + by (import hollight DIVISION_LE) + +lemma DIVISION_GT: "ALL (D::nat => hollight.real) (a::hollight.real) b::hollight.real. + division (a, b) D --> + (ALL n::nat. < n (dsize D) --> real_lt (D n) (D (dsize D)))" + by (import hollight DIVISION_GT) + +lemma DIVISION_EQ: "ALL (D::nat => hollight.real) (a::hollight.real) b::hollight.real. + division (a, b) D --> (a = b) = (dsize D = (0::nat))" + by (import hollight DIVISION_EQ) + +lemma DIVISION_LBOUND: "ALL (x::nat => hollight.real) (xa::hollight.real) (xb::hollight.real) + xc::nat. division (xa, xb) x --> real_le xa (x xc)" + by (import hollight DIVISION_LBOUND) + +lemma DIVISION_LBOUND_LT: "ALL (x::nat => hollight.real) (xa::hollight.real) (xb::hollight.real) + xc::nat. + division (xa, xb) x & dsize x ~= (0::nat) --> real_lt xa (x (Suc xc))" + by (import hollight DIVISION_LBOUND_LT) + +lemma DIVISION_UBOUND: "ALL (x::nat => hollight.real) (xa::hollight.real) (xb::hollight.real) + xc::nat. division (xa, xb) x --> real_le (x xc) xb" + by (import hollight DIVISION_UBOUND) + +lemma DIVISION_UBOUND_LT: "ALL (D::nat => hollight.real) (a::hollight.real) (b::hollight.real) n::nat. + division (a, b) D & < n (dsize D) --> real_lt (D n) b" + by (import hollight DIVISION_UBOUND_LT) + +lemma DIVISION_APPEND_LEMMA1: "ALL (a::hollight.real) (b::hollight.real) (c::hollight.real) + (D1::nat => hollight.real) D2::nat => hollight.real. + division (a, b) D1 & division (b, c) D2 --> + (ALL n::nat. + < n (dsize D1 + dsize D2) --> + real_lt (COND (< n (dsize D1)) (D1 n) (D2 (n - dsize D1))) + (COND (< (Suc n) (dsize D1)) (D1 (Suc n)) + (D2 (Suc n - dsize D1)))) & + (ALL n::nat. + >= n (dsize D1 + dsize D2) --> + COND (< n (dsize D1)) (D1 n) (D2 (n - dsize D1)) = + COND (< (dsize D1 + dsize D2) (dsize D1)) (D1 (dsize D1 + dsize D2)) + (D2 (dsize D1 + dsize D2 - dsize D1)))" + by (import hollight DIVISION_APPEND_LEMMA1) + +lemma DIVISION_APPEND_LEMMA2: "ALL (a::hollight.real) (b::hollight.real) (c::hollight.real) + (D1::nat => hollight.real) D2::nat => hollight.real. + division (a, b) D1 & division (b, c) D2 --> + dsize (%n::nat. COND (< n (dsize D1)) (D1 n) (D2 (n - dsize D1))) = + dsize D1 + dsize D2" + by (import hollight DIVISION_APPEND_LEMMA2) + +lemma DIVISION_APPEND: "ALL (a::hollight.real) (b::hollight.real) c::hollight.real. + (EX (D1::nat => hollight.real) p1::nat => hollight.real. + tdiv (a, b) (D1, p1) & + fine (g::hollight.real => hollight.real) (D1, p1)) & + (EX (D2::nat => hollight.real) p2::nat => hollight.real. + tdiv (b, c) (D2, p2) & fine g (D2, p2)) --> + (EX (x::nat => hollight.real) p::nat => hollight.real. + tdiv (a, c) (x, p) & fine g (x, p))" + by (import hollight DIVISION_APPEND) + +lemma DIVISION_EXISTS: "ALL (a::hollight.real) (b::hollight.real) g::hollight.real => hollight.real. + real_le a b & gauge (%x::hollight.real. real_le a x & real_le x b) g --> + (EX (D::nat => hollight.real) p::nat => hollight.real. + tdiv (a, b) (D, p) & fine g (D, p))" + by (import hollight DIVISION_EXISTS) + +lemma GAUGE_MIN: "ALL (E::hollight.real => bool) (g1::hollight.real => hollight.real) + g2::hollight.real => hollight.real. + gauge E g1 & gauge E g2 --> + gauge E (%x::hollight.real. COND (real_lt (g1 x) (g2 x)) (g1 x) (g2 x))" + by (import hollight GAUGE_MIN) + +lemma FINE_MIN: "ALL (g1::hollight.real => hollight.real) + (g2::hollight.real => hollight.real) (D::nat => hollight.real) + p::nat => hollight.real. + fine (%x::hollight.real. COND (real_lt (g1 x) (g2 x)) (g1 x) (g2 x)) + (D, p) --> + fine g1 (D, p) & fine g2 (D, p)" + by (import hollight FINE_MIN) + +lemma DINT_UNIQ: "ALL (a::hollight.real) (b::hollight.real) + (f::hollight.real => hollight.real) (k1::hollight.real) + k2::hollight.real. + real_le a b & defint (a, b) f k1 & defint (a, b) f k2 --> k1 = k2" + by (import hollight DINT_UNIQ) + +lemma INTEGRAL_NULL: "ALL (f::hollight.real => hollight.real) a::hollight.real. + defint (a, a) f (real_of_num (0::nat))" + by (import hollight INTEGRAL_NULL) + +lemma STRADDLE_LEMMA: "ALL (f::hollight.real => hollight.real) (f'::hollight.real => hollight.real) + (a::hollight.real) (b::hollight.real) e::hollight.real. + (ALL x::hollight.real. real_le a x & real_le x b --> diffl f (f' x) x) & + real_lt (real_of_num (0::nat)) e --> + (EX x::hollight.real => hollight.real. + gauge (%x::hollight.real. real_le a x & real_le x b) x & + (ALL (xa::hollight.real) (u::hollight.real) v::hollight.real. + real_le a u & + real_le u xa & + real_le xa v & real_le v b & real_lt (real_sub v u) (x xa) --> + real_le + (real_abs + (real_sub (real_sub (f v) (f u)) + (real_mul (f' xa) (real_sub v u)))) + (real_mul e (real_sub v u))))" + by (import hollight STRADDLE_LEMMA) + +lemma FTC1: "ALL (f::hollight.real => hollight.real) (f'::hollight.real => hollight.real) + (a::hollight.real) b::hollight.real. + real_le a b & + (ALL x::hollight.real. + real_le a x & real_le x b --> diffl f (f' x) x) --> + defint (a, b) f' (real_sub (f b) (f a))" + by (import hollight FTC1) + +lemma MCLAURIN: "ALL (f::hollight.real => hollight.real) + (diff::nat => hollight.real => hollight.real) (h::hollight.real) n::nat. + real_lt (real_of_num (0::nat)) h & + < (0::nat) n & + diff (0::nat) = f & + (ALL (m::nat) t::hollight.real. + < m n & real_le (real_of_num (0::nat)) t & real_le t h --> + diffl (diff m) (diff (Suc m) t) t) --> + (EX t::hollight.real. + real_lt (real_of_num (0::nat)) t & + real_lt t h & + f h = + real_add + (psum (0::nat, n) + (%m::nat. + real_mul + (real_div (diff m (real_of_num (0::nat))) + (real_of_num (FACT m))) + (real_pow h m))) + (real_mul (real_div (diff n t) (real_of_num (FACT n))) + (real_pow h n)))" + by (import hollight MCLAURIN) + +lemma MCLAURIN_NEG: "ALL (f::hollight.real => hollight.real) + (diff::nat => hollight.real => hollight.real) (h::hollight.real) n::nat. + real_lt h (real_of_num (0::nat)) & + < (0::nat) n & + diff (0::nat) = f & + (ALL (m::nat) t::hollight.real. + < m n & real_le h t & real_le t (real_of_num (0::nat)) --> + diffl (diff m) (diff (Suc m) t) t) --> + (EX t::hollight.real. + real_lt h t & + real_lt t (real_of_num (0::nat)) & + f h = + real_add + (psum (0::nat, n) + (%m::nat. + real_mul + (real_div (diff m (real_of_num (0::nat))) + (real_of_num (FACT m))) + (real_pow h m))) + (real_mul (real_div (diff n t) (real_of_num (FACT n))) + (real_pow h n)))" + by (import hollight MCLAURIN_NEG) + +lemma MCLAURIN_BI_LE: "ALL (f::hollight.real => hollight.real) + (diff::nat => hollight.real => hollight.real) (x::hollight.real) n::nat. + diff (0::nat) = f & + (ALL (m::nat) t::hollight.real. + < m n & real_le (real_abs t) (real_abs x) --> + diffl (diff m) (diff (Suc m) t) t) --> + (EX xa::hollight.real. + real_le (real_abs xa) (real_abs x) & + f x = + real_add + (psum (0::nat, n) + (%m::nat. + real_mul + (real_div (diff m (real_of_num (0::nat))) + (real_of_num (FACT m))) + (real_pow x m))) + (real_mul (real_div (diff n xa) (real_of_num (FACT n))) + (real_pow x n)))" + by (import hollight MCLAURIN_BI_LE) + +lemma MCLAURIN_ALL_LT: "ALL (f::hollight.real => hollight.real) + diff::nat => hollight.real => hollight.real. + diff (0::nat) = f & + (ALL (m::nat) x::hollight.real. diffl (diff m) (diff (Suc m) x) x) --> + (ALL (x::hollight.real) n::nat. + x ~= real_of_num (0::nat) & < (0::nat) n --> + (EX t::hollight.real. + real_lt (real_of_num (0::nat)) (real_abs t) & + real_lt (real_abs t) (real_abs x) & + f x = + real_add + (psum (0::nat, n) + (%m::nat. + real_mul + (real_div (diff m (real_of_num (0::nat))) + (real_of_num (FACT m))) + (real_pow x m))) + (real_mul (real_div (diff n t) (real_of_num (FACT n))) + (real_pow x n))))" + by (import hollight MCLAURIN_ALL_LT) + +lemma MCLAURIN_ZERO: "ALL (diff::nat => hollight.real => hollight.real) (n::nat) x::hollight.real. + x = real_of_num (0::nat) & < (0::nat) n --> + psum (0::nat, n) + (%m::nat. + real_mul + (real_div (diff m (real_of_num (0::nat))) (real_of_num (FACT m))) + (real_pow x m)) = + diff (0::nat) (real_of_num (0::nat))" + by (import hollight MCLAURIN_ZERO) + +lemma MCLAURIN_ALL_LE: "ALL (f::hollight.real => hollight.real) + diff::nat => hollight.real => hollight.real. + diff (0::nat) = f & + (ALL (m::nat) x::hollight.real. diffl (diff m) (diff (Suc m) x) x) --> + (ALL (x::hollight.real) n::nat. + EX t::hollight.real. + real_le (real_abs t) (real_abs x) & + f x = + real_add + (psum (0::nat, n) + (%m::nat. + real_mul + (real_div (diff m (real_of_num (0::nat))) + (real_of_num (FACT m))) + (real_pow x m))) + (real_mul (real_div (diff n t) (real_of_num (FACT n))) + (real_pow x n)))" + by (import hollight MCLAURIN_ALL_LE) + +lemma MCLAURIN_EXP_LEMMA: "exp = exp & (ALL (x::nat) xa::hollight.real. diffl exp (exp xa) xa)" + by (import hollight MCLAURIN_EXP_LEMMA) + +lemma MCLAURIN_EXP_LT: "ALL (x::hollight.real) n::nat. + x ~= real_of_num (0::nat) & < (0::nat) n --> + (EX t::hollight.real. + real_lt (real_of_num (0::nat)) (real_abs t) & + real_lt (real_abs t) (real_abs x) & + exp x = + real_add + (psum (0::nat, n) + (%m::nat. real_div (real_pow x m) (real_of_num (FACT m)))) + (real_mul (real_div (exp t) (real_of_num (FACT n))) (real_pow x n)))" + by (import hollight MCLAURIN_EXP_LT) + +lemma MCLAURIN_EXP_LE: "ALL (x::hollight.real) n::nat. + EX t::hollight.real. + real_le (real_abs t) (real_abs x) & + exp x = + real_add + (psum (0::nat, n) + (%m::nat. real_div (real_pow x m) (real_of_num (FACT m)))) + (real_mul (real_div (exp t) (real_of_num (FACT n))) (real_pow x n))" + by (import hollight MCLAURIN_EXP_LE) + +lemma DIFF_LN_COMPOSITE: "ALL (g::hollight.real => hollight.real) (m::hollight.real) x::hollight.real. + diffl g m x & real_lt (real_of_num (0::nat)) (g x) --> + diffl (%x::hollight.real. ln (g x)) (real_mul (real_inv (g x)) m) x" + by (import hollight DIFF_LN_COMPOSITE) + +lemma MCLAURIN_LN_POS: "ALL (x::hollight.real) n::nat. + real_lt (real_of_num (0::nat)) x & < (0::nat) n --> + (EX t::hollight.real. + real_lt (real_of_num (0::nat)) t & + real_lt t x & + ln (real_add (real_of_num (NUMERAL_BIT1 (0::nat))) x) = + real_add + (psum (0::nat, n) + (%m::nat. + real_mul + (real_pow (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) + (Suc m)) + (real_div (real_pow x m) (real_of_num m)))) + (real_mul + (real_pow (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) + (Suc n)) + (real_div (real_pow x n) + (real_mul (real_of_num n) + (real_pow (real_add (real_of_num (NUMERAL_BIT1 (0::nat))) t) + n)))))" + by (import hollight MCLAURIN_LN_POS) + +lemma MCLAURIN_LN_NEG: "ALL (x::hollight.real) n::nat. + real_lt (real_of_num (0::nat)) x & + real_lt x (real_of_num (NUMERAL_BIT1 (0::nat))) & < (0::nat) n --> + (EX t::hollight.real. + real_lt (real_of_num (0::nat)) t & + real_lt t x & + real_neg (ln (real_sub (real_of_num (NUMERAL_BIT1 (0::nat))) x)) = + real_add + (psum (0::nat, n) + (%m::nat. real_div (real_pow x m) (real_of_num m))) + (real_div (real_pow x n) + (real_mul (real_of_num n) + (real_pow (real_sub (real_of_num (NUMERAL_BIT1 (0::nat))) t) + n))))" + by (import hollight MCLAURIN_LN_NEG) + +lemma MCLAURIN_SIN: "ALL (x::hollight.real) n::nat. + real_le + (real_abs + (real_sub (sin x) + (psum (0::nat, n) + (%m::nat. + real_mul + (COND (EVEN m) (real_of_num (0::nat)) + (real_div + (real_pow + (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) + (DIV (m - NUMERAL_BIT1 (0::nat)) + (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) + (real_of_num (FACT m)))) + (real_pow x m))))) + (real_mul (real_inv (real_of_num (FACT n))) (real_pow (real_abs x) n))" + by (import hollight MCLAURIN_SIN) + +lemma MCLAURIN_COS: "ALL (x::hollight.real) n::nat. + real_le + (real_abs + (real_sub (cos x) + (psum (0::nat, n) + (%m::nat. + real_mul + (COND (EVEN m) + (real_div + (real_pow + (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) + (DIV m (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) + (real_of_num (FACT m))) + (real_of_num (0::nat))) + (real_pow x m))))) + (real_mul (real_inv (real_of_num (FACT n))) (real_pow (real_abs x) n))" + by (import hollight MCLAURIN_COS) + +lemma REAL_ATN_POWSER_SUMMABLE: "ALL x::hollight.real. + real_lt (real_abs x) (real_of_num (NUMERAL_BIT1 (0::nat))) --> + summable + (%n::nat. + real_mul + (COND (EVEN n) (real_of_num (0::nat)) + (real_div + (real_pow (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) + (DIV (n - NUMERAL_BIT1 (0::nat)) + (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) + (real_of_num n))) + (real_pow x n))" + by (import hollight REAL_ATN_POWSER_SUMMABLE) + +lemma REAL_ATN_POWSER_DIFFS_SUMMABLE: "ALL x::hollight.real. + real_lt (real_abs x) (real_of_num (NUMERAL_BIT1 (0::nat))) --> + summable + (%xa::nat. + real_mul + (diffs + (%n::nat. + COND (EVEN n) (real_of_num (0::nat)) + (real_div + (real_pow (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) + (DIV (n - NUMERAL_BIT1 (0::nat)) + (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) + (real_of_num n))) + xa) + (real_pow x xa))" + by (import hollight REAL_ATN_POWSER_DIFFS_SUMMABLE) + +lemma REAL_ATN_POWSER_DIFFS_SUM: "ALL x::hollight.real. + real_lt (real_abs x) (real_of_num (NUMERAL_BIT1 (0::nat))) --> + sums + (%n::nat. + real_mul + (diffs + (%n::nat. + COND (EVEN n) (real_of_num (0::nat)) + (real_div + (real_pow (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) + (DIV (n - NUMERAL_BIT1 (0::nat)) + (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) + (real_of_num n))) + n) + (real_pow x n)) + (real_inv + (real_add (real_of_num (NUMERAL_BIT1 (0::nat))) + (real_pow x (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))))" + by (import hollight REAL_ATN_POWSER_DIFFS_SUM) + +lemma REAL_ATN_POWSER_DIFFS_DIFFS_SUMMABLE: "ALL x::hollight.real. + real_lt (real_abs x) (real_of_num (NUMERAL_BIT1 (0::nat))) --> + summable + (%xa::nat. + real_mul + (diffs + (diffs + (%n::nat. + COND (EVEN n) (real_of_num (0::nat)) + (real_div + (real_pow + (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) + (DIV (n - NUMERAL_BIT1 (0::nat)) + (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) + (real_of_num n)))) + xa) + (real_pow x xa))" + by (import hollight REAL_ATN_POWSER_DIFFS_DIFFS_SUMMABLE) + +lemma REAL_ATN_POWSER_DIFFL: "ALL x::hollight.real. + real_lt (real_abs x) (real_of_num (NUMERAL_BIT1 (0::nat))) --> + diffl + (%x::hollight.real. + suminf + (%n::nat. + real_mul + (COND (EVEN n) (real_of_num (0::nat)) + (real_div + (real_pow (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) + (DIV (n - NUMERAL_BIT1 (0::nat)) + (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) + (real_of_num n))) + (real_pow x n))) + (real_inv + (real_add (real_of_num (NUMERAL_BIT1 (0::nat))) + (real_pow x (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat)))))) + x" + by (import hollight REAL_ATN_POWSER_DIFFL) + +lemma REAL_ATN_POWSER: "ALL x::hollight.real. + real_lt (real_abs x) (real_of_num (NUMERAL_BIT1 (0::nat))) --> + sums + (%n::nat. + real_mul + (COND (EVEN n) (real_of_num (0::nat)) + (real_div + (real_pow (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) + (DIV (n - NUMERAL_BIT1 (0::nat)) + (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) + (real_of_num n))) + (real_pow x n)) + (atn x)" + by (import hollight REAL_ATN_POWSER) + +lemma MCLAURIN_ATN: "ALL (x::hollight.real) n::nat. + real_lt (real_abs x) (real_of_num (NUMERAL_BIT1 (0::nat))) --> + real_le + (real_abs + (real_sub (atn x) + (psum (0::nat, n) + (%m::nat. + real_mul + (COND (EVEN m) (real_of_num (0::nat)) + (real_div + (real_pow + (real_neg (real_of_num (NUMERAL_BIT1 (0::nat)))) + (DIV (m - NUMERAL_BIT1 (0::nat)) + (NUMERAL_BIT0 (NUMERAL_BIT1 (0::nat))))) + (real_of_num m))) + (real_pow x m))))) + (real_div (real_pow (real_abs x) n) + (real_sub (real_of_num (NUMERAL_BIT1 (0::nat))) (real_abs x)))" + by (import hollight MCLAURIN_ATN) + +;end_setup + +end + diff -r 7e417a7cbf9f -r bd59bfd4bf37 src/HOL/Import/HOLLight/hollight.imp --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Import/HOLLight/hollight.imp Mon Sep 26 02:27:14 2005 +0200 @@ -0,0 +1,2980 @@ +import + +import_segment "hollight" + +def_maps + "two_2" > "two_2_def" + "two_1" > "two_1_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" + "three_3" > "three_3_def" + "three_2" > "three_2_def" + "three_1" > "three_1_def" + "tendsto" > "tendsto_def" + "tends_real_real" > "tends_real_real_def" + "tends_num_real" > "tends_num_real_def" + "tends" > "tends_def" + "tdiv" > "tdiv_def" + "tan" > "tan_def" + "tailadmissible" > "tailadmissible_def" + "support" > "support_def" + "superadmissible" > "superadmissible_def" + "sup" > "sup_def" + "sums" > "sums_def" + "summable" > "summable_def" + "suminf" > "suminf_def" + "sum" > "sum_def" + "subseq" > "subseq_def" + "sqrt" > "sqrt_def" + "sndcart" > "sndcart_def" + "sin" > "sin_def" + "set_of_list" > "set_of_list_def" + "rsum" > "rsum_def" + "root" > "root_def" + "real_sub" > "real_sub_def" + "real_pow" > "real_pow_def" + "real_of_num" > "real_of_num_def" + "real_neg" > "real_neg_def" + "real_mul" > "real_mul_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" + "re_universe" > "re_universe_def" + "re_union" > "re_union_def" + "re_subset" > "re_subset_def" + "re_null" > "re_null_def" + "re_intersect" > "re_intersect_def" + "re_compl" > "re_compl_def" + "re_Union" > "re_Union_def" + "psum" > "psum_def" + "pi" > "pi_def" + "pastecart" > "pastecart_def" + "pairwise" > "pairwise_def" + "nsum" > "nsum_def" + "neutral" > "neutral_def" + "neigh" > "neigh_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" + "mtop" > "mtop_def" + "mr1" > "mr1_def" + "monoidal" > "monoidal_def" + "mono" > "mono_def" + "mod_real" > "mod_real_def" + "mod_nat" > "mod_nat_def" + "mod_int" > "mod_int_def" + "minimal" > "minimal_def" + "measure" > "measure_def" + "ln" > "ln_def" + "list_of_set" > "list_of_set_def" + "limpt" > "limpt_def" + "lim" > "lim_def" + "lambda" > "lambda_def" + "iterate" > "iterate_def" + "istopology" > "istopology_def" + "ismet" > "ismet_def" + "is_nadd" > "is_nadd_def" + "is_int" > "is_int_def" + "int_sub" > "int_sub_def" + "int_pow" > "int_pow_def" + "int_of_num" > "int_of_num_def" + "int_neg" > "int_neg_def" + "int_mul" > "int_mul_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_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" + "gauge" > "gauge_def" + "fstcart" > "fstcart_def" + "finite_index" > "finite_index_def" + "fine" > "fine_def" + "exp" > "exp_def" + "eqeq" > "eqeq_def" + "dsize" > "dsize_def" + "dotdot" > "dotdot_def" + "dorder" > "dorder_def" + "division" > "division_def" + "dist" > "dist_def" + "dimindex" > "dimindex_def" + "diffs" > "diffs_def" + "diffl" > "diffl_def" + "differentiable" > "differentiable_def" + "defint" > "defint_def" + "cos" > "cos_def" + "convergent" > "convergent_def" + "contl" > "contl_def" + "closed" > "closed_def" + "cauchy" > "cauchy_def" + "bounded" > "bounded_def" + "ball" > "ball_def" + "atn" > "atn_def" + "asn" > "asn_def" + "admissible" > "admissible_def" + "acs" > "acs_def" + "_FALSITY_" > "_FALSITY__def" + "_10314" > "_10314_def" + "_10313" > "_10313_def" + "_10312" > "_10312_def" + "_10289" > "_10289_def" + "_10288" > "_10288_def" + "ZRECSPACE" > "ZRECSPACE_def" + "ZIP" > "ZIP_def" + "ZCONSTR" > "ZCONSTR_def" + "ZBOT" > "ZBOT_def" + "WF" > "WF_def" + "UNIV" > "UNIV_def" + "UNIONS" > "UNIONS_def" + "UNION" > "UNION_def" + "UNCURRY" > "UNCURRY_def" + "TL" > "TL_def" + "SURJ" > "SURJ_def" + "SUBSET" > "SUBSET_def" + "SOME" > "SOME_def" + "SING" > "SING_def" + "SETSPEC" > "SETSPEC_def" + "REVERSE" > "REVERSE_def" + "REST" > "REST_def" + "REPLICATE" > "REPLICATE_def" + "PSUBSET" > "PSUBSET_def" + "PASSOC" > "PASSOC_def" + "PAIRWISE" > "PAIRWISE_def" + "OUTR" > "OUTR_def" + "OUTL" > "OUTL_def" + "ODD" > "ODD_def" + "NUMSUM" > "NUMSUM_def" + "NUMSND" > "NUMSND_def" + "NUMRIGHT" > "NUMRIGHT_def" + "NUMPAIR" > "NUMPAIR_def" + "NUMLEFT" > "NUMLEFT_def" + "NUMFST" > "NUMFST_def" + "NULL" > "NULL_def" + "NONE" > "NONE_def" + "NIL" > "NIL_def" + "MOD" > "MOD_def" + "MEM" > "MEM_def" + "MAP2" > "MAP2_def" + "MAP" > "MAP_def" + "LET_END" > "LET_END_def" + "LENGTH" > "LENGTH_def" + "LAST" > "LAST_def" + "ITSET" > "ITSET_def" + "ITLIST2" > "ITLIST2_def" + "ITLIST" > "ITLIST_def" + "ISO" > "ISO_def" + "INTERS" > "INTERS_def" + "INTER" > "INTER_def" + "INSERT" > "INSERT_def" + "INR" > "INR_def" + "INL" > "INL_def" + "INJP" > "INJP_def" + "INJN" > "INJN_def" + "INJF" > "INJF_def" + "INJA" > "INJA_def" + "INJ" > "INJ_def" + "INFINITE" > "INFINITE_def" + "IN" > "IN_def" + "IMAGE" > "IMAGE_def" + "HD" > "HD_def" + "HAS_SIZE" > "HAS_SIZE_def" + "GSPEC" > "GSPEC_def" + "GEQ" > "GEQ_def" + "GABS" > "GABS_def" + "FNIL" > "FNIL_def" + "FINREC" > "FINREC_def" + "FINITE" > "FINITE_def" + "FILTER" > "FILTER_def" + "FCONS" > "FCONS_def" + "FACT" > "FACT_def" + "EXP" > "EXP_def" + "EX" > "EX_def" + "EVEN" > "EVEN_def" + "EMPTY" > "EMPTY_def" + "EL" > "EL_def" + "DIV" > "DIV_def" + "DISJOINT" > "DISJOINT_def" + "DIFF" > "DIFF_def" + "DELETE" > "DELETE_def" + "DECIMAL" > "DECIMAL_def" + "CURRY" > "CURRY_def" + "COUNTABLE" > "COUNTABLE_def" + "CONSTR" > "CONSTR_def" + "CONS" > "CONS_def" + "COND" > "COND_def" + "CHOICE" > "CHOICE_def" + "CASEWISE" > "CASEWISE_def" + "CARD_LT" > "CARD_LT_def" + "CARD_LE" > "CARD_LE_def" + "CARD_GT" > "CARD_GT_def" + "CARD_GE" > "CARD_GE_def" + "CARD_EQ" > "CARD_EQ_def" + "CARD" > "CARD_def" + "BOTTOM" > "BOTTOM_def" + "BIJ" > "BIJ_def" + "ASSOC" > "ASSOC_def" + "APPEND" > "APPEND_def" + "ALL_list" > "ALL_list_def" + "ALL2" > "ALL2_def" + ">=" > ">=_def" + ">" > ">_def" + "<=" > "<=_def" + "<" > "<_def" + "$" > "$_def" + +type_maps + "topology" > "HOLLight.hollight.topology" + "sum" > "HOLLight.hollight.sum" + "recspace" > "HOLLight.hollight.recspace" + "real" > "HOLLight.hollight.real" + "prod" > "*" + "option" > "HOLLight.hollight.option" + "num" > "nat" + "nadd" > "HOLLight.hollight.nadd" + "metric" > "HOLLight.hollight.metric" + "list" > "HOLLight.hollight.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" + "cart" > "HOLLight.hollight.cart" + "bool" > "bool" + "N_3" > "HOLLight.hollight.N_3" + "N_2" > "HOLLight.hollight.N_2" + "N_1" > "Product_Type.unit" + +const_maps + "~" > "Not" + "two_2" > "HOLLight.hollight.two_2" + "two_1" > "HOLLight.hollight.two_1" + "treal_of_num" > "HOLLight.hollight.treal_of_num" + "treal_neg" > "HOLLight.hollight.treal_neg" + "treal_mul" > "HOLLight.hollight.treal_mul" + "treal_le" > "HOLLight.hollight.treal_le" + "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" + "tendsto" > "HOLLight.hollight.tendsto" + "tends_real_real" > "HOLLight.hollight.tends_real_real" + "tends_num_real" > "HOLLight.hollight.tends_num_real" + "tends" > "HOLLight.hollight.tends" + "tdiv" > "HOLLight.hollight.tdiv" + "tan" > "HOLLight.hollight.tan" + "tailadmissible" > "HOLLight.hollight.tailadmissible" + "support" > "HOLLight.hollight.support" + "superadmissible" > "HOLLight.hollight.superadmissible" + "sup" > "HOLLight.hollight.sup" + "sums" > "HOLLight.hollight.sums" + "summable" > "HOLLight.hollight.summable" + "suminf" > "HOLLight.hollight.suminf" + "sum" > "HOLLight.hollight.sum" + "subseq" > "HOLLight.hollight.subseq" + "sqrt" > "HOLLight.hollight.sqrt" + "sndcart" > "HOLLight.hollight.sndcart" + "sin" > "HOLLight.hollight.sin" + "set_of_list" > "HOLLight.hollight.set_of_list" + "rsum" > "HOLLight.hollight.rsum" + "root" > "HOLLight.hollight.root" + "real_sub" > "HOLLight.hollight.real_sub" + "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_min" > "HOLLight.hollight.real_min" + "real_max" > "HOLLight.hollight.real_max" + "real_lt" > "HOLLight.hollight.real_lt" + "real_le" > "HOLLight.hollight.real_le" + "real_inv" > "HOLLight.hollight.real_inv" + "real_gt" > "HOLLight.hollight.real_gt" + "real_ge" > "HOLLight.hollight.real_ge" + "real_div" > "HOLLight.hollight.real_div" + "real_add" > "HOLLight.hollight.real_add" + "real_abs" > "HOLLight.hollight.real_abs" + "re_universe" > "HOLLight.hollight.re_universe" + "re_union" > "HOLLight.hollight.re_union" + "re_subset" > "HOLLight.hollight.re_subset" + "re_null" > "HOLLight.hollight.re_null" + "re_intersect" > "HOLLight.hollight.re_intersect" + "re_compl" > "HOLLight.hollight.re_compl" + "re_Union" > "HOLLight.hollight.re_Union" + "psum" > "HOLLight.hollight.psum" + "pi" > "HOLLight.hollight.pi" + "pastecart" > "HOLLight.hollight.pastecart" + "pairwise" > "HOLLight.hollight.pairwise" + "one" > "Product_Type.Unity" + "o" > "Fun.comp" + "nsum" > "HOLLight.hollight.nsum" + "neutral" > "HOLLight.hollight.neutral" + "neigh" > "HOLLight.hollight.neigh" + "nadd_rinv" > "HOLLight.hollight.nadd_rinv" + "nadd_of_num" > "HOLLight.hollight.nadd_of_num" + "nadd_mul" > "HOLLight.hollight.nadd_mul" + "nadd_le" > "HOLLight.hollight.nadd_le" + "nadd_inv" > "HOLLight.hollight.nadd_inv" + "nadd_eq" > "HOLLight.hollight.nadd_eq" + "nadd_add" > "HOLLight.hollight.nadd_add" + "mtop" > "HOLLight.hollight.mtop" + "mr1" > "HOLLight.hollight.mr1" + "monoidal" > "HOLLight.hollight.monoidal" + "mono" > "HOLLight.hollight.mono" + "mod_real" > "HOLLight.hollight.mod_real" + "mod_nat" > "HOLLight.hollight.mod_nat" + "mod_int" > "HOLLight.hollight.mod_int" + "mk_pair" > "Product_Type.Pair_Rep" + "minimal" > "HOLLight.hollight.minimal" + "measure" > "HOLLight.hollight.measure" + "ln" > "HOLLight.hollight.ln" + "list_of_set" > "HOLLight.hollight.list_of_set" + "limpt" > "HOLLight.hollight.limpt" + "lim" > "HOLLight.hollight.lim" + "lambda" > "HOLLight.hollight.lambda" + "iterate" > "HOLLight.hollight.iterate" + "istopology" > "HOLLight.hollight.istopology" + "ismet" > "HOLLight.hollight.ismet" + "is_nadd" > "HOLLight.hollight.is_nadd" + "is_int" > "HOLLight.hollight.is_int" + "int_sub" > "HOLLight.hollight.int_sub" + "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_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_add" > "HOLLight.hollight.int_add" + "int_abs" > "HOLLight.hollight.int_abs" + "hreal_of_num" > "HOLLight.hollight.hreal_of_num" + "hreal_mul" > "HOLLight.hollight.hreal_mul" + "hreal_le" > "HOLLight.hollight.hreal_le" + "hreal_inv" > "HOLLight.hollight.hreal_inv" + "hreal_add" > "HOLLight.hollight.hreal_add" + "gauge" > "HOLLight.hollight.gauge" + "fstcart" > "HOLLight.hollight.fstcart" + "finite_index" > "HOLLight.hollight.finite_index" + "fine" > "HOLLight.hollight.fine" + "exp" > "HOLLight.hollight.exp" + "eqeq" > "HOLLight.hollight.eqeq" + "dsize" > "HOLLight.hollight.dsize" + "dotdot" > "HOLLight.hollight.dotdot" + "dorder" > "HOLLight.hollight.dorder" + "division" > "HOLLight.hollight.division" + "dist" > "HOLLight.hollight.dist" + "dimindex" > "HOLLight.hollight.dimindex" + "diffs" > "HOLLight.hollight.diffs" + "diffl" > "HOLLight.hollight.diffl" + "differentiable" > "HOLLight.hollight.differentiable" + "defint" > "HOLLight.hollight.defint" + "cos" > "HOLLight.hollight.cos" + "convergent" > "HOLLight.hollight.convergent" + "contl" > "HOLLight.hollight.contl" + "closed" > "HOLLight.hollight.closed" + "cauchy" > "HOLLight.hollight.cauchy" + "bounded" > "HOLLight.hollight.bounded" + "ball" > "HOLLight.hollight.ball" + "atn" > "HOLLight.hollight.atn" + "asn" > "HOLLight.hollight.asn" + "admissible" > "HOLLight.hollight.admissible" + "acs" > "HOLLight.hollight.acs" + "_FALSITY_" > "HOLLight.hollight._FALSITY_" + "_10314" > "HOLLight.hollight._10314" + "_10313" > "HOLLight.hollight._10313" + "_10312" > "HOLLight.hollight._10312" + "_10289" > "HOLLight.hollight._10289" + "_10288" > "HOLLight.hollight._10288" + "_0" > "0" :: "nat" + "\\/" > "op |" + "ZRECSPACE" > "HOLLight.hollight.ZRECSPACE" + "ZIP" > "HOLLight.hollight.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" + "UNCURRY" > "HOLLight.hollight.UNCURRY" + "TL" > "HOLLight.hollight.TL" + "T" > "True" + "SURJ" > "HOLLight.hollight.SURJ" + "SUC" > "Suc" + "SUBSET" > "HOLLight.hollight.SUBSET" + "SOME" > "HOLLight.hollight.SOME" + "SND" > "snd" + "SING" > "HOLLight.hollight.SING" + "SETSPEC" > "HOLLight.hollight.SETSPEC" + "REVERSE" > "HOLLight.hollight.REVERSE" + "REST" > "HOLLight.hollight.REST" + "REP_prod" > "Rep_Prod" + "REPLICATE" > "HOLLight.hollight.REPLICATE" + "PSUBSET" > "HOLLight.hollight.PSUBSET" + "PRE" > "HOLLightCompat.Pred" + "PASSOC" > "HOLLight.hollight.PASSOC" + "PAIRWISE" > "HOLLight.hollight.PAIRWISE" + "OUTR" > "HOLLight.hollight.OUTR" + "OUTL" > "HOLLight.hollight.OUTL" + "ONTO" > "Fun.surj" + "ONE_ONE" > "HOL4Setup.ONE_ONE" + "ODD" > "HOLLight.hollight.ODD" + "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" + "LET_END" > "HOLLight.hollight.LET_END" + "LET" > "HOL4Compat.LET" + "LENGTH" > "HOLLight.hollight.LENGTH" + "LAST" > "HOLLight.hollight.LAST" + "ITSET" > "HOLLight.hollight.ITSET" + "ITLIST2" > "HOLLight.hollight.ITLIST2" + "ITLIST" > "HOLLight.hollight.ITLIST" + "ISO" > "HOLLight.hollight.ISO" + "INTERS" > "HOLLight.hollight.INTERS" + "INTER" > "HOLLight.hollight.INTER" + "INSERT" > "HOLLight.hollight.INSERT" + "INR" > "HOLLight.hollight.INR" + "INL" > "HOLLight.hollight.INL" + "INJP" > "HOLLight.hollight.INJP" + "INJN" > "HOLLight.hollight.INJN" + "INJF" > "HOLLight.hollight.INJF" + "INJA" > "HOLLight.hollight.INJA" + "INJ" > "HOLLight.hollight.INJ" + "INFINITE" > "HOLLight.hollight.INFINITE" + "IN" > "HOLLight.hollight.IN" + "IMAGE" > "HOLLight.hollight.IMAGE" + "I" > "Fun.id" + "HD" > "HOLLight.hollight.HD" + "HAS_SIZE" > "HOLLight.hollight.HAS_SIZE" + "GSPEC" > "HOLLight.hollight.GSPEC" + "GEQ" > "HOLLight.hollight.GEQ" + "GABS" > "HOLLight.hollight.GABS" + "FST" > "fst" + "FNIL" > "HOLLight.hollight.FNIL" + "FINREC" > "HOLLight.hollight.FINREC" + "FINITE" > "HOLLight.hollight.FINITE" + "FILTER" > "HOLLight.hollight.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" + "DECIMAL" > "HOLLight.hollight.DECIMAL" + "CURRY" > "HOLLight.hollight.CURRY" + "COUNTABLE" > "HOLLight.hollight.COUNTABLE" + "CONSTR" > "HOLLight.hollight.CONSTR" + "CONS" > "HOLLight.hollight.CONS" + "COND" > "HOLLight.hollight.COND" + "CHOICE" > "HOLLight.hollight.CHOICE" + "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" + "BOTTOM" > "HOLLight.hollight.BOTTOM" + "BIT1" > "HOL4Compat.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" + "@" > "Hilbert_Choice.Eps" + "?!" > "Ex1" + "?" > "Ex" + ">=" > "HOLLight.hollight.>=" + ">" > "HOLLight.hollight.>" + "==>" > "op -->" + "=" > "op =" + "<=" > "HOLLight.hollight.<=" + "<" > "HOLLight.hollight.<" + "/\\" > "op &" + "-" > "op -" :: "nat => nat => nat" + "," > "Pair" + "+" > "op +" :: "nat => nat => nat" + "*" > "op *" :: "nat => nat => nat" + "$" > "HOLLight.hollight.$" + "!" > "All" + +const_renames + "ALL" > "ALL_list" + "==" > "eqeq" + ".." > "dotdot" + +thm_maps + "two_2_def" > "HOLLight.hollight.two_2_def" + "two_1_def" > "HOLLight.hollight.two_1_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" + "treal_le_def" > "HOLLight.hollight.treal_le_def" + "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" > "HOLLight.hollight.th" + "tendsto_def" > "HOLLight.hollight.tendsto_def" + "tends_real_real_def" > "HOLLight.hollight.tends_real_real_def" + "tends_num_real_def" > "HOLLight.hollight.tends_num_real_def" + "tends_def" > "HOLLight.hollight.tends_def" + "tdiv_def" > "HOLLight.hollight.tdiv_def" + "tan_def" > "HOLLight.hollight.tan_def" + "tailadmissible_def" > "HOLLight.hollight.tailadmissible_def" + "support_def" > "HOLLight.hollight.support_def" + "superadmissible_def" > "HOLLight.hollight.superadmissible_def" + "sup_def" > "HOLLight.hollight.sup_def" + "sup" > "HOLLight.hollight.sup" + "sums_def" > "HOLLight.hollight.sums_def" + "summable_def" > "HOLLight.hollight.summable_def" + "suminf_def" > "HOLLight.hollight.suminf_def" + "sum_def" > "HOLLight.hollight.sum_def" + "sum_EXISTS" > "HOLLight.hollight.sum_EXISTS" + "sum" > "HOLLight.hollight.sum" + "subseq_def" > "HOLLight.hollight.subseq_def" + "sth" > "HOLLight.hollight.sth" + "sqrt_def" > "HOLLight.hollight.sqrt_def" + "sqrt" > "HOLLight.hollight.sqrt" + "sndcart_def" > "HOLLight.hollight.sndcart_def" + "sin_def" > "HOLLight.hollight.sin_def" + "set_of_list_def" > "HOLLight.hollight.set_of_list_def" + "rsum_def" > "HOLLight.hollight.rsum_def" + "root_def" > "HOLLight.hollight.root_def" + "right_th" > "HOLLight.hollight.right_th" + "real_sub_def" > "HOLLight.hollight.real_sub_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_min_def" > "HOLLight.hollight.real_min_def" + "real_max_def" > "HOLLight.hollight.real_max_def" + "real_lt_def" > "HOLLight.hollight.real_lt_def" + "real_le_def" > "HOLLight.hollight.real_le_def" + "real_le" > "HOLLight.hollight.real_le" + "real_inv_def" > "HOLLight.hollight.real_inv_def" + "real_gt_def" > "HOLLight.hollight.real_gt_def" + "real_ge_def" > "HOLLight.hollight.real_ge_def" + "real_div_def" > "HOLLight.hollight.real_div_def" + "real_add_def" > "HOLLight.hollight.real_add_def" + "real_abs_def" > "HOLLight.hollight.real_abs_def" + "re_universe_def" > "HOLLight.hollight.re_universe_def" + "re_union_def" > "HOLLight.hollight.re_union_def" + "re_subset_def" > "HOLLight.hollight.re_subset_def" + "re_null_def" > "HOLLight.hollight.re_null_def" + "re_intersect_def" > "HOLLight.hollight.re_intersect_def" + "re_compl_def" > "HOLLight.hollight.re_compl_def" + "re_Union_def" > "HOLLight.hollight.re_Union_def" + "psum_def" > "HOLLight.hollight.psum_def" + "pi_def" > "HOLLight.hollight.pi_def" + "pastecart_def" > "HOLLight.hollight.pastecart_def" + "pairwise_def" > "HOLLight.hollight.pairwise_def" + "pair_RECURSION" > "HOLLight.hollight.pair_RECURSION" + "pair_INDUCT" > "Datatype.prod.induct" + "one_axiom" > "HOLLight.hollight.one_axiom" + "one_RECURSION" > "HOLLight.hollight.one_RECURSION" + "one_INDUCT" > "Datatype.unit.induct" + "one_Axiom" > "HOLLight.hollight.one_Axiom" + "one" > "HOL4Compat.one" + "o_THM" > "Fun.o_apply" + "o_ASSOC" > "HOLLight.hollight.o_ASSOC" + "num_WOP" > "HOLLight.hollight.num_WOP" + "num_WF" > "HOLLight.hollight.num_WF" + "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_FINITE_AVOID" > "HOLLight.hollight.num_FINITE_AVOID" + "num_FINITE" > "HOLLight.hollight.num_FINITE" + "num_CASES" > "Nat.nat.nchotomy" + "num_Axiom" > "HOLLight.hollight.num_Axiom" + "nsum_def" > "HOLLight.hollight.nsum_def" + "neutral_def" > "HOLLight.hollight.neutral_def" + "neigh_def" > "HOLLight.hollight.neigh_def" + "nadd_rinv_def" > "HOLLight.hollight.nadd_rinv_def" + "nadd_of_num_def" > "HOLLight.hollight.nadd_of_num_def" + "nadd_mul_def" > "HOLLight.hollight.nadd_mul_def" + "nadd_le_def" > "HOLLight.hollight.nadd_le_def" + "nadd_inv_def" > "HOLLight.hollight.nadd_inv_def" + "nadd_eq_def" > "HOLLight.hollight.nadd_eq_def" + "nadd_add_def" > "HOLLight.hollight.nadd_add_def" + "mtop_istopology" > "HOLLight.hollight.mtop_istopology" + "mtop_def" > "HOLLight.hollight.mtop_def" + "mr1_def" > "HOLLight.hollight.mr1_def" + "monoidal_def" > "HOLLight.hollight.monoidal_def" + "mono_def" > "HOLLight.hollight.mono_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" + "ln_def" > "HOLLight.hollight.ln_def" + "list_of_set_def" > "HOLLight.hollight.list_of_set_def" + "list_INDUCT" > "HOLLight.hollight.list_INDUCT" + "list_CASES" > "HOLLight.hollight.list_CASES" + "limpt_def" > "HOLLight.hollight.limpt_def" + "lim_def" > "HOLLight.hollight.lim_def" + "lambda_def" > "HOLLight.hollight.lambda_def" + "iterate_def" > "HOLLight.hollight.iterate_def" + "istopology_def" > "HOLLight.hollight.istopology_def" + "ismet_def" > "HOLLight.hollight.ismet_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" + "int_sub_th" > "HOLLight.hollight.int_sub_th" + "int_sub_def" > "HOLLight.hollight.int_sub_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" + "int_of_num_def" > "HOLLight.hollight.int_of_num_def" + "int_neg_th" > "HOLLight.hollight.int_neg_th" + "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_min_th" > "HOLLight.hollight.int_min_th" + "int_min_def" > "HOLLight.hollight.int_min_def" + "int_max_th" > "HOLLight.hollight.int_max_th" + "int_max_def" > "HOLLight.hollight.int_max_def" + "int_lt_def" > "HOLLight.hollight.int_lt_def" + "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_add_th" > "HOLLight.hollight.int_add_th" + "int_add_def" > "HOLLight.hollight.int_add_def" + "int_abs_th" > "HOLLight.hollight.int_abs_th" + "int_abs_def" > "HOLLight.hollight.int_abs_def" + "hreal_of_num_def" > "HOLLight.hollight.hreal_of_num_def" + "hreal_mul_def" > "HOLLight.hollight.hreal_mul_def" + "hreal_le_def" > "HOLLight.hollight.hreal_le_def" + "hreal_inv_def" > "HOLLight.hollight.hreal_inv_def" + "hreal_add_def" > "HOLLight.hollight.hreal_add_def" + "gauge_def" > "HOLLight.hollight.gauge_def" + "fstcart_def" > "HOLLight.hollight.fstcart_def" + "finite_index_def" > "HOLLight.hollight.finite_index_def" + "fine_def" > "HOLLight.hollight.fine_def" + "exp_def" > "HOLLight.hollight.exp_def" + "eqeq_def" > "HOLLight.hollight.eqeq_def" + "dsize_def" > "HOLLight.hollight.dsize_def" + "dotdot_def" > "HOLLight.hollight.dotdot_def" + "dorder_def" > "HOLLight.hollight.dorder_def" + "division_def" > "HOLLight.hollight.division_def" + "dist_def" > "HOLLight.hollight.dist_def" + "dimindex_def" > "HOLLight.hollight.dimindex_def" + "diffs_def" > "HOLLight.hollight.diffs_def" + "diffl_def" > "HOLLight.hollight.diffl_def" + "differentiable_def" > "HOLLight.hollight.differentiable_def" + "dest_int_rep" > "HOLLight.hollight.dest_int_rep" + "defint_def" > "HOLLight.hollight.defint_def" + "cth" > "HOLLight.hollight.cth" + "cos_def" > "HOLLight.hollight.cos_def" + "convergent_def" > "HOLLight.hollight.convergent_def" + "contl_def" > "HOLLight.hollight.contl_def" + "closed_def" > "HOLLight.hollight.closed_def" + "cauchy_def" > "HOLLight.hollight.cauchy_def" + "bounded_def" > "HOLLight.hollight.bounded_def" + "ball_def" > "HOLLight.hollight.ball_def" + "ax__3" > "HOL4Setup.INFINITY_AX" + "ax__2" > "Hilbert_Choice.tfl_some" + "ax__1" > "Presburger.fm_modd_pinf" + "atn_def" > "HOLLight.hollight.atn_def" + "asn_def" > "HOLLight.hollight.asn_def" + "admissible_def" > "HOLLight.hollight.admissible_def" + "acs_def" > "HOLLight.hollight.acs_def" + "_FALSITY__def" > "HOLLight.hollight._FALSITY__def" + "_10314_def" > "HOLLight.hollight._10314_def" + "_10313_def" > "HOLLight.hollight._10313_def" + "_10312_def" > "HOLLight.hollight._10312_def" + "_10289_def" > "HOLLight.hollight._10289_def" + "_10288_def" > "HOLLight.hollight._10288_def" + "ZRECSPACE_def" > "HOLLight.hollight.ZRECSPACE_def" + "ZIP_def" > "HOLLight.hollight.ZIP_def" + "ZIP" > "HOLLight.hollight.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" + "WF_REFL" > "HOLLight.hollight.WF_REFL" + "WF_REC_num" > "HOLLight.hollight.WF_REC_num" + "WF_REC_WF" > "HOLLight.hollight.WF_REC_WF" + "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_IND" > "HOLLight.hollight.WF_IND" + "WF_FALSE" > "HOLLight.hollight.WF_FALSE" + "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" + "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_EMPTY" > "HOLLight.hollight.UNION_EMPTY" + "UNION_COMM" > "HOLLight.hollight.UNION_COMM" + "UNION_ASSOC" > "HOLLight.hollight.UNION_ASSOC" + "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" + "UNCURRY_def" > "HOLLight.hollight.UNCURRY_def" + "TYDEF_topology" > "HOLLight.hollight.TYDEF_topology" + "TYDEF_sum" > "HOLLight.hollight.TYDEF_sum" + "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_metric" > "HOLLight.hollight.TYDEF_metric" + "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_cart" > "HOLLight.hollight.TYDEF_cart" + "TYDEF_3" > "HOLLight.hollight.TYDEF_3" + "TYDEF_2" > "HOLLight.hollight.TYDEF_2" + "TWO" > "HOLLight.hollight.TWO" + "TRIV_OR_FORALL_THM" > "HOLLight.hollight.TRIV_OR_FORALL_THM" + "TRIV_FORALL_OR_THM" > "HOLLight.hollight.TRIV_FORALL_OR_THM" + "TRIV_FORALL_IMP_THM" > "HOLLight.hollight.TRIV_FORALL_IMP_THM" + "TRIV_EXISTS_IMP_THM" > "HOLLight.hollight.TRIV_EXISTS_IMP_THM" + "TRIV_EXISTS_AND_THM" > "HOLLight.hollight.TRIV_EXISTS_AND_THM" + "TRIV_AND_EXISTS_THM" > "HOLLight.hollight.TRIV_AND_EXISTS_THM" + "TREAL_OF_NUM_WELLDEF" > "HOLLight.hollight.TREAL_OF_NUM_WELLDEF" + "TREAL_OF_NUM_MUL" > "HOLLight.hollight.TREAL_OF_NUM_MUL" + "TREAL_OF_NUM_LE" > "HOLLight.hollight.TREAL_OF_NUM_LE" + "TREAL_OF_NUM_EQ" > "HOLLight.hollight.TREAL_OF_NUM_EQ" + "TREAL_OF_NUM_ADD" > "HOLLight.hollight.TREAL_OF_NUM_ADD" + "TREAL_NEG_WELLDEF" > "HOLLight.hollight.TREAL_NEG_WELLDEF" + "TREAL_MUL_WELLDEFR" > "HOLLight.hollight.TREAL_MUL_WELLDEFR" + "TREAL_MUL_WELLDEF" > "HOLLight.hollight.TREAL_MUL_WELLDEF" + "TREAL_MUL_SYM_EQ" > "HOLLight.hollight.TREAL_MUL_SYM_EQ" + "TREAL_MUL_SYM" > "HOLLight.hollight.TREAL_MUL_SYM" + "TREAL_MUL_LINV" > "HOLLight.hollight.TREAL_MUL_LINV" + "TREAL_MUL_LID" > "HOLLight.hollight.TREAL_MUL_LID" + "TREAL_MUL_ASSOC" > "HOLLight.hollight.TREAL_MUL_ASSOC" + "TREAL_LE_WELLDEF" > "HOLLight.hollight.TREAL_LE_WELLDEF" + "TREAL_LE_TRANS" > "HOLLight.hollight.TREAL_LE_TRANS" + "TREAL_LE_TOTAL" > "HOLLight.hollight.TREAL_LE_TOTAL" + "TREAL_LE_REFL" > "HOLLight.hollight.TREAL_LE_REFL" + "TREAL_LE_MUL" > "HOLLight.hollight.TREAL_LE_MUL" + "TREAL_LE_LADD_IMP" > "HOLLight.hollight.TREAL_LE_LADD_IMP" + "TREAL_LE_ANTISYM" > "HOLLight.hollight.TREAL_LE_ANTISYM" + "TREAL_INV_WELLDEF" > "HOLLight.hollight.TREAL_INV_WELLDEF" + "TREAL_INV_0" > "HOLLight.hollight.TREAL_INV_0" + "TREAL_EQ_TRANS" > "HOLLight.hollight.TREAL_EQ_TRANS" + "TREAL_EQ_SYM" > "HOLLight.hollight.TREAL_EQ_SYM" + "TREAL_EQ_REFL" > "HOLLight.hollight.TREAL_EQ_REFL" + "TREAL_EQ_IMP_LE" > "HOLLight.hollight.TREAL_EQ_IMP_LE" + "TREAL_EQ_AP" > "HOLLight.hollight.TREAL_EQ_AP" + "TREAL_ADD_WELLDEFR" > "HOLLight.hollight.TREAL_ADD_WELLDEFR" + "TREAL_ADD_WELLDEF" > "HOLLight.hollight.TREAL_ADD_WELLDEF" + "TREAL_ADD_SYM_EQ" > "HOLLight.hollight.TREAL_ADD_SYM_EQ" + "TREAL_ADD_SYM" > "HOLLight.hollight.TREAL_ADD_SYM" + "TREAL_ADD_LINV" > "HOLLight.hollight.TREAL_ADD_LINV" + "TREAL_ADD_LID" > "HOLLight.hollight.TREAL_ADD_LID" + "TREAL_ADD_LDISTRIB" > "HOLLight.hollight.TREAL_ADD_LDISTRIB" + "TREAL_ADD_ASSOC" > "HOLLight.hollight.TREAL_ADD_ASSOC" + "TOPOLOGY_UNION" > "HOLLight.hollight.TOPOLOGY_UNION" + "TOPOLOGY" > "HOLLight.hollight.TOPOLOGY" + "TL_def" > "HOLLight.hollight.TL_def" + "TERMDIFF_STRONG" > "HOLLight.hollight.TERMDIFF_STRONG" + "TERMDIFF_LEMMA5" > "HOLLight.hollight.TERMDIFF_LEMMA5" + "TERMDIFF_LEMMA4" > "HOLLight.hollight.TERMDIFF_LEMMA4" + "TERMDIFF_LEMMA3" > "HOLLight.hollight.TERMDIFF_LEMMA3" + "TERMDIFF_LEMMA2" > "HOLLight.hollight.TERMDIFF_LEMMA2" + "TERMDIFF_LEMMA1" > "HOLLight.hollight.TERMDIFF_LEMMA1" + "TERMDIFF_CONVERGES" > "HOLLight.hollight.TERMDIFF_CONVERGES" + "TERMDIFF" > "HOLLight.hollight.TERMDIFF" + "TAN_TOTAL_POS" > "HOLLight.hollight.TAN_TOTAL_POS" + "TAN_TOTAL_LEMMA" > "HOLLight.hollight.TAN_TOTAL_LEMMA" + "TAN_TOTAL" > "HOLLight.hollight.TAN_TOTAL" + "TAN_SEC" > "HOLLight.hollight.TAN_SEC" + "TAN_POS_PI2" > "HOLLight.hollight.TAN_POS_PI2" + "TAN_PI4" > "HOLLight.hollight.TAN_PI4" + "TAN_PI" > "HOLLight.hollight.TAN_PI" + "TAN_PERIODIC_PI" > "HOLLight.hollight.TAN_PERIODIC_PI" + "TAN_PERIODIC_NPI" > "HOLLight.hollight.TAN_PERIODIC_NPI" + "TAN_PERIODIC" > "HOLLight.hollight.TAN_PERIODIC" + "TAN_NPI" > "HOLLight.hollight.TAN_NPI" + "TAN_NEG" > "HOLLight.hollight.TAN_NEG" + "TAN_DOUBLE" > "HOLLight.hollight.TAN_DOUBLE" + "TAN_COT" > "HOLLight.hollight.TAN_COT" + "TAN_BOUND_PI2" > "HOLLight.hollight.TAN_BOUND_PI2" + "TAN_ATN" > "HOLLight.hollight.TAN_ATN" + "TAN_ADD" > "HOLLight.hollight.TAN_ADD" + "TAN_ABS_GE_X" > "HOLLight.hollight.TAN_ABS_GE_X" + "TAN_0" > "HOLLight.hollight.TAN_0" + "TAIL_IMP_SUPERADMISSIBLE" > "HOLLight.hollight.TAIL_IMP_SUPERADMISSIBLE" + "SWAP_FORALL_THM" > "HOLLight.hollight.SWAP_FORALL_THM" + "SWAP_EXISTS_THM" > "HOLLight.hollight.SWAP_EXISTS_THM" + "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_IFF_INJECTIVE_GEN" > "HOLLight.hollight.SURJECTIVE_IFF_INJECTIVE_GEN" + "SURJECTIVE_IFF_INJECTIVE" > "HOLLight.hollight.SURJECTIVE_IFF_INJECTIVE" + "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_T" > "HOLLight.hollight.SUPERADMISSIBLE_T" + "SUPERADMISSIBLE_COND" > "HOLLight.hollight.SUPERADMISSIBLE_COND" + "SUM_ZERO" > "HOLLight.hollight.SUM_ZERO" + "SUM_UNIQ" > "HOLLight.hollight.SUM_UNIQ" + "SUM_UNION_RZERO" > "HOLLight.hollight.SUM_UNION_RZERO" + "SUM_UNION_LZERO" > "HOLLight.hollight.SUM_UNION_LZERO" + "SUM_UNION_EQ" > "HOLLight.hollight.SUM_UNION_EQ" + "SUM_UNION" > "HOLLight.hollight.SUM_UNION" + "SUM_TWO" > "HOLLight.hollight.SUM_TWO" + "SUM_TRIV_NUMSEG" > "HOLLight.hollight.SUM_TRIV_NUMSEG" + "SUM_SWAP_NUMSEG" > "HOLLight.hollight.SUM_SWAP_NUMSEG" + "SUM_SWAP" > "HOLLight.hollight.SUM_SWAP" + "SUM_SUPPORT" > "HOLLight.hollight.SUM_SUPPORT" + "SUM_SUPERSET" > "HOLLight.hollight.SUM_SUPERSET" + "SUM_SUM_RESTRICT" > "HOLLight.hollight.SUM_SUM_RESTRICT" + "SUM_SUMMABLE" > "HOLLight.hollight.SUM_SUMMABLE" + "SUM_SUB_NUMSEG" > "HOLLight.hollight.SUM_SUB_NUMSEG" + "SUM_SUBST" > "HOLLight.hollight.SUM_SUBST" + "SUM_SUBSET_SIMPLE" > "HOLLight.hollight.SUM_SUBSET_SIMPLE" + "SUM_SUBSET" > "HOLLight.hollight.SUM_SUBSET" + "SUM_SUB" > "HOLLight.hollight.SUM_SUB" + "SUM_SPLIT" > "HOLLight.hollight.SUM_SPLIT" + "SUM_SING_NUMSEG" > "HOLLight.hollight.SUM_SING_NUMSEG" + "SUM_SING" > "HOLLight.hollight.SUM_SING" + "SUM_RESTRICT_SET" > "HOLLight.hollight.SUM_RESTRICT_SET" + "SUM_RESTRICT" > "HOLLight.hollight.SUM_RESTRICT" + "SUM_REINDEX" > "HOLLight.hollight.SUM_REINDEX" + "SUM_POS_LE_NUMSEG" > "HOLLight.hollight.SUM_POS_LE_NUMSEG" + "SUM_POS_LE" > "HOLLight.hollight.SUM_POS_LE" + "SUM_POS_GEN" > "HOLLight.hollight.SUM_POS_GEN" + "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_POS" > "HOLLight.hollight.SUM_POS" + "SUM_OFFSET_0" > "HOLLight.hollight.SUM_OFFSET_0" + "SUM_OFFSET" > "HOLLight.hollight.SUM_OFFSET" + "SUM_NSUB" > "HOLLight.hollight.SUM_NSUB" + "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_MORETERMS_EQ" > "HOLLight.hollight.SUM_MORETERMS_EQ" + "SUM_LT_ALL" > "HOLLight.hollight.SUM_LT_ALL" + "SUM_LT" > "HOLLight.hollight.SUM_LT" + "SUM_LE_NUMSEG" > "HOLLight.hollight.SUM_LE_NUMSEG" + "SUM_LE" > "HOLLight.hollight.SUM_LE" + "SUM_IMAGE_GEN" > "HOLLight.hollight.SUM_IMAGE_GEN" + "SUM_IMAGE" > "HOLLight.hollight.SUM_IMAGE" + "SUM_HORNER" > "HOLLight.hollight.SUM_HORNER" + "SUM_GROUP" > "HOLLight.hollight.SUM_GROUP" + "SUM_EQ_SUPERSET" > "HOLLight.hollight.SUM_EQ_SUPERSET" + "SUM_EQ_NUMSEG" > "HOLLight.hollight.SUM_EQ_NUMSEG" + "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_DIFFERENCES_EQ" > "HOLLight.hollight.SUM_DIFFERENCES_EQ" + "SUM_DIFF" > "HOLLight.hollight.SUM_DIFF" + "SUM_DELTA" > "HOLLight.hollight.SUM_DELTA" + "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_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_CANCEL" > "HOLLight.hollight.SUM_CANCEL" + "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" + "SUM_BOUND_GEN" > "HOLLight.hollight.SUM_BOUND_GEN" + "SUM_BOUND" > "HOLLight.hollight.SUM_BOUND" + "SUM_ADD_SPLIT" > "HOLLight.hollight.SUM_ADD_SPLIT" + "SUM_ADD_NUMSEG" > "HOLLight.hollight.SUM_ADD_NUMSEG" + "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_2" > "HOLLight.hollight.SUM_2" + "SUM_1" > "HOLLight.hollight.SUM_1" + "SUM_0" > "HOLLight.hollight.SUM_0" + "SUMMABLE_SUM" > "HOLLight.hollight.SUMMABLE_SUM" + "SUC_SUB1" > "HOLLight.hollight.SUC_SUB1" + "SUC_INJ" > "Nat.nat.simps_3" + "SUB_SUC" > "Nat.diff_Suc_Suc" + "SUB_SUB" > "HOLLight.hollight.SUB_SUB" + "SUB_REFL" > "Nat.diff_self_eq_0" + "SUB_PRESUC" > "HOLLight.hollight.SUB_PRESUC" + "SUB_OLD" > "HOLLight.hollight.SUB_OLD" + "SUB_EQ_0" > "HOLLight.hollight.SUB_EQ_0" + "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_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_UNION" > "HOLLight.hollight.SUBSET_UNION" + "SUBSET_TRANS" > "HOLLight.hollight.SUBSET_TRANS" + "SUBSET_RESTRICT" > "HOLLight.hollight.SUBSET_RESTRICT" + "SUBSET_REFL" > "HOLLight.hollight.SUBSET_REFL" + "SUBSET_PSUBSET_TRANS" > "HOLLight.hollight.SUBSET_PSUBSET_TRANS" + "SUBSET_NUMSEG" > "HOLLight.hollight.SUBSET_NUMSEG" + "SUBSET_INTER_ABSORPTION" > "HOLLight.hollight.SUBSET_INTER_ABSORPTION" + "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_DELETE" > "HOLLight.hollight.SUBSET_DELETE" + "SUBSET_ANTISYM" > "HOLLight.hollight.SUBSET_ANTISYM" + "SUBSETA_TRANS" > "HOLLight.hollight.SUBSETA_TRANS" + "SUBSETA_REFL" > "HOLLight.hollight.SUBSETA_REFL" + "SUBSETA_ANTISYM" > "HOLLight.hollight.SUBSETA_ANTISYM" + "SUBSEQ_SUC" > "HOLLight.hollight.SUBSEQ_SUC" + "STRADDLE_LEMMA" > "HOLLight.hollight.STRADDLE_LEMMA" + "SQRT_POW_2" > "HOLLight.hollight.SQRT_POW_2" + "SQRT_POW2" > "HOLLight.hollight.SQRT_POW2" + "SQRT_POS_UNIQ" > "HOLLight.hollight.SQRT_POS_UNIQ" + "SQRT_POS_LT" > "HOLLight.hollight.SQRT_POS_LT" + "SQRT_POS_LE" > "HOLLight.hollight.SQRT_POS_LE" + "SQRT_MUL" > "HOLLight.hollight.SQRT_MUL" + "SQRT_MONO_LT_EQ" > "HOLLight.hollight.SQRT_MONO_LT_EQ" + "SQRT_MONO_LT" > "HOLLight.hollight.SQRT_MONO_LT" + "SQRT_MONO_LE_EQ" > "HOLLight.hollight.SQRT_MONO_LE_EQ" + "SQRT_MONO_LE" > "HOLLight.hollight.SQRT_MONO_LE" + "SQRT_INV" > "HOLLight.hollight.SQRT_INV" + "SQRT_INJ" > "HOLLight.hollight.SQRT_INJ" + "SQRT_EVEN_POW2" > "HOLLight.hollight.SQRT_EVEN_POW2" + "SQRT_EQ_0" > "HOLLight.hollight.SQRT_EQ_0" + "SQRT_DIV" > "HOLLight.hollight.SQRT_DIV" + "SQRT_1" > "HOLLight.hollight.SQRT_1" + "SQRT_0" > "HOLLight.hollight.SQRT_0" + "SOME_def" > "HOLLight.hollight.SOME_def" + "SNDCART_PASTECART" > "HOLLight.hollight.SNDCART_PASTECART" + "SND" > "Product_Type.snd_conv" + "SKOLEM_THM" > "HOLLight.hollight.SKOLEM_THM" + "SIN_ZERO_PI" > "HOLLight.hollight.SIN_ZERO_PI" + "SIN_ZERO_LEMMA" > "HOLLight.hollight.SIN_ZERO_LEMMA" + "SIN_ZERO" > "HOLLight.hollight.SIN_ZERO" + "SIN_TOTAL" > "HOLLight.hollight.SIN_TOTAL" + "SIN_POS_PI_LE" > "HOLLight.hollight.SIN_POS_PI_LE" + "SIN_POS_PI2" > "HOLLight.hollight.SIN_POS_PI2" + "SIN_POS_PI" > "HOLLight.hollight.SIN_POS_PI" + "SIN_POS" > "HOLLight.hollight.SIN_POS" + "SIN_PI2" > "HOLLight.hollight.SIN_PI2" + "SIN_PI" > "HOLLight.hollight.SIN_PI" + "SIN_PERIODIC_PI" > "HOLLight.hollight.SIN_PERIODIC_PI" + "SIN_PERIODIC" > "HOLLight.hollight.SIN_PERIODIC" + "SIN_PAIRED" > "HOLLight.hollight.SIN_PAIRED" + "SIN_NPI" > "HOLLight.hollight.SIN_NPI" + "SIN_NEGLEMMA" > "HOLLight.hollight.SIN_NEGLEMMA" + "SIN_NEG" > "HOLLight.hollight.SIN_NEG" + "SIN_FDIFF" > "HOLLight.hollight.SIN_FDIFF" + "SIN_DOUBLE" > "HOLLight.hollight.SIN_DOUBLE" + "SIN_COS_SQRT" > "HOLLight.hollight.SIN_COS_SQRT" + "SIN_COS_NEG" > "HOLLight.hollight.SIN_COS_NEG" + "SIN_COS_ADD" > "HOLLight.hollight.SIN_COS_ADD" + "SIN_COS" > "HOLLight.hollight.SIN_COS" + "SIN_CONVERGES" > "HOLLight.hollight.SIN_CONVERGES" + "SIN_CIRCLE" > "HOLLight.hollight.SIN_CIRCLE" + "SIN_BOUNDS" > "HOLLight.hollight.SIN_BOUNDS" + "SIN_BOUND" > "HOLLight.hollight.SIN_BOUND" + "SIN_ASN" > "HOLLight.hollight.SIN_ASN" + "SIN_ADD" > "HOLLight.hollight.SIN_ADD" + "SIN_ACS_NZ" > "HOLLight.hollight.SIN_ACS_NZ" + "SIN_0" > "HOLLight.hollight.SIN_0" + "SING_def" > "HOLLight.hollight.SING_def" + "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_CASES" > "HOLLight.hollight.SET_CASES" + "SETSPEC_def" > "HOLLight.hollight.SETSPEC_def" + "SETOK_LE_LT" > "HOLLight.hollight.SETOK_LE_LT" + "SER_ZERO" > "HOLLight.hollight.SER_ZERO" + "SER_UNIQ" > "HOLLight.hollight.SER_UNIQ" + "SER_SUB" > "HOLLight.hollight.SER_SUB" + "SER_RATIO" > "HOLLight.hollight.SER_RATIO" + "SER_POS_LT_PAIR" > "HOLLight.hollight.SER_POS_LT_PAIR" + "SER_POS_LT" > "HOLLight.hollight.SER_POS_LT" + "SER_POS_LE" > "HOLLight.hollight.SER_POS_LE" + "SER_PAIR" > "HOLLight.hollight.SER_PAIR" + "SER_OFFSET_REV" > "HOLLight.hollight.SER_OFFSET_REV" + "SER_OFFSET" > "HOLLight.hollight.SER_OFFSET" + "SER_NEG" > "HOLLight.hollight.SER_NEG" + "SER_LE2" > "HOLLight.hollight.SER_LE2" + "SER_LE" > "HOLLight.hollight.SER_LE" + "SER_GROUP" > "HOLLight.hollight.SER_GROUP" + "SER_COMPARA_UNIFORM_WEAK" > "HOLLight.hollight.SER_COMPARA_UNIFORM_WEAK" + "SER_COMPARA_UNIFORM" > "HOLLight.hollight.SER_COMPARA_UNIFORM" + "SER_COMPARA" > "HOLLight.hollight.SER_COMPARA" + "SER_COMPAR" > "HOLLight.hollight.SER_COMPAR" + "SER_CMUL" > "HOLLight.hollight.SER_CMUL" + "SER_CDIV" > "HOLLight.hollight.SER_CDIV" + "SER_CAUCHY" > "HOLLight.hollight.SER_CAUCHY" + "SER_ADD" > "HOLLight.hollight.SER_ADD" + "SER_ACONV" > "HOLLight.hollight.SER_ACONV" + "SER_ABS" > "HOLLight.hollight.SER_ABS" + "SER_0" > "HOLLight.hollight.SER_0" + "SEQ_UNIQ" > "HOLLight.hollight.SEQ_UNIQ" + "SEQ_TRUNCATION" > "HOLLight.hollight.SEQ_TRUNCATION" + "SEQ_TRANSFORM" > "HOLLight.hollight.SEQ_TRANSFORM" + "SEQ_TENDS" > "HOLLight.hollight.SEQ_TENDS" + "SEQ_SUM" > "HOLLight.hollight.SEQ_SUM" + "SEQ_SUC" > "HOLLight.hollight.SEQ_SUC" + "SEQ_SUBLE" > "HOLLight.hollight.SEQ_SUBLE" + "SEQ_SUB" > "HOLLight.hollight.SEQ_SUB" + "SEQ_SBOUNDED" > "HOLLight.hollight.SEQ_SBOUNDED" + "SEQ_POWER_ABS" > "HOLLight.hollight.SEQ_POWER_ABS" + "SEQ_POWER" > "HOLLight.hollight.SEQ_POWER" + "SEQ_NULL" > "HOLLight.hollight.SEQ_NULL" + "SEQ_NPOW" > "HOLLight.hollight.SEQ_NPOW" + "SEQ_NEG_CONV" > "HOLLight.hollight.SEQ_NEG_CONV" + "SEQ_NEG_BOUNDED" > "HOLLight.hollight.SEQ_NEG_BOUNDED" + "SEQ_NEG" > "HOLLight.hollight.SEQ_NEG" + "SEQ_MUL" > "HOLLight.hollight.SEQ_MUL" + "SEQ_MONOSUB" > "HOLLight.hollight.SEQ_MONOSUB" + "SEQ_LIM" > "HOLLight.hollight.SEQ_LIM" + "SEQ_LE_0" > "HOLLight.hollight.SEQ_LE_0" + "SEQ_LE" > "HOLLight.hollight.SEQ_LE" + "SEQ_INV0" > "HOLLight.hollight.SEQ_INV0" + "SEQ_INV" > "HOLLight.hollight.SEQ_INV" + "SEQ_ICONV" > "HOLLight.hollight.SEQ_ICONV" + "SEQ_DIV" > "HOLLight.hollight.SEQ_DIV" + "SEQ_DIRECT" > "HOLLight.hollight.SEQ_DIRECT" + "SEQ_CONT_UNIFORM" > "HOLLight.hollight.SEQ_CONT_UNIFORM" + "SEQ_CONST" > "HOLLight.hollight.SEQ_CONST" + "SEQ_CBOUNDED" > "HOLLight.hollight.SEQ_CBOUNDED" + "SEQ_CAUCHY" > "HOLLight.hollight.SEQ_CAUCHY" + "SEQ_BOUNDED_2" > "HOLLight.hollight.SEQ_BOUNDED_2" + "SEQ_BOUNDED" > "HOLLight.hollight.SEQ_BOUNDED" + "SEQ_BCONV" > "HOLLight.hollight.SEQ_BCONV" + "SEQ_ADD" > "HOLLight.hollight.SEQ_ADD" + "SEQ_ABS_IMP" > "HOLLight.hollight.SEQ_ABS_IMP" + "SEQ_ABS" > "HOLLight.hollight.SEQ_ABS" + "SEQ" > "HOLLight.hollight.SEQ" + "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" + "ROOT_POW_POS" > "HOLLight.hollight.ROOT_POW_POS" + "ROOT_POS_UNIQ" > "HOLLight.hollight.ROOT_POS_UNIQ" + "ROOT_POS_POSITIVE" > "HOLLight.hollight.ROOT_POS_POSITIVE" + "ROOT_MUL" > "HOLLight.hollight.ROOT_MUL" + "ROOT_MONO_LT_EQ" > "HOLLight.hollight.ROOT_MONO_LT_EQ" + "ROOT_MONO_LT" > "HOLLight.hollight.ROOT_MONO_LT" + "ROOT_MONO_LE_EQ" > "HOLLight.hollight.ROOT_MONO_LE_EQ" + "ROOT_MONO_LE" > "HOLLight.hollight.ROOT_MONO_LE" + "ROOT_LT_LEMMA" > "HOLLight.hollight.ROOT_LT_LEMMA" + "ROOT_LN" > "HOLLight.hollight.ROOT_LN" + "ROOT_INV" > "HOLLight.hollight.ROOT_INV" + "ROOT_INJ" > "HOLLight.hollight.ROOT_INJ" + "ROOT_DIV" > "HOLLight.hollight.ROOT_DIV" + "ROOT_1" > "HOLLight.hollight.ROOT_1" + "ROOT_0" > "HOLLight.hollight.ROOT_0" + "ROLLE" > "HOLLight.hollight.ROLLE" + "RIGHT_SUB_DISTRIB" > "Nat.nat_distrib_3" + "RIGHT_OR_FORALL_THM" > "HOL.all_simps_4" + "RIGHT_OR_EXISTS_THM" > "HOL.ex_simps_4" + "RIGHT_OR_DISTRIB" > "HOL.conj_disj_distribR" + "RIGHT_IMP_FORALL_THM" > "HOL.all_simps_6" + "RIGHT_IMP_EXISTS_THM" > "HOL.ex_simps_6" + "RIGHT_FORALL_OR_THM" > "HOL.all_simps_4" + "RIGHT_FORALL_IMP_THM" > "HOL.all_simps_6" + "RIGHT_EXISTS_IMP_THM" > "HOL.ex_simps_6" + "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" + "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" + "REAL_WLOG_LT" > "HOLLight.hollight.REAL_WLOG_LT" + "REAL_WLOG_LE" > "HOLLight.hollight.REAL_WLOG_LE" + "REAL_SUP_UBOUND_LE" > "HOLLight.hollight.REAL_SUP_UBOUND_LE" + "REAL_SUP_UBOUND" > "HOLLight.hollight.REAL_SUP_UBOUND" + "REAL_SUP_LE" > "HOLLight.hollight.REAL_SUP_LE" + "REAL_SUP_EXISTS" > "HOLLight.hollight.REAL_SUP_EXISTS" + "REAL_SUP" > "HOLLight.hollight.REAL_SUP" + "REAL_SUMSQ" > "HOLLight.hollight.REAL_SUMSQ" + "REAL_SUB_TRIANGLE" > "HOLLight.hollight.REAL_SUB_TRIANGLE" + "REAL_SUB_SUB2" > "HOLLight.hollight.REAL_SUB_SUB2" + "REAL_SUB_SUB" > "HOLLight.hollight.REAL_SUB_SUB" + "REAL_SUB_RZERO" > "HOLLight.hollight.REAL_SUB_RZERO" + "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_NEG2" > "HOLLight.hollight.REAL_SUB_NEG2" + "REAL_SUB_LZERO" > "HOLLight.hollight.REAL_SUB_LZERO" + "REAL_SUB_LT" > "HOLLight.hollight.REAL_SUB_LT" + "REAL_SUB_LNEG" > "HOLLight.hollight.REAL_SUB_LNEG" + "REAL_SUB_LE" > "HOLLight.hollight.REAL_SUB_LE" + "REAL_SUB_LDISTRIB" > "HOLLight.hollight.REAL_SUB_LDISTRIB" + "REAL_SUB_INV2" > "HOLLight.hollight.REAL_SUB_INV2" + "REAL_SUB_INV" > "HOLLight.hollight.REAL_SUB_INV" + "REAL_SUB_ADD2" > "HOLLight.hollight.REAL_SUB_ADD2" + "REAL_SUB_ADD" > "HOLLight.hollight.REAL_SUB_ADD" + "REAL_SUB_ABS" > "HOLLight.hollight.REAL_SUB_ABS" + "REAL_SUB_0" > "HOLLight.hollight.REAL_SUB_0" + "REAL_RNEG_UNIQ" > "HOLLight.hollight.REAL_RNEG_UNIQ" + "REAL_RINV_UNIQ" > "HOLLight.hollight.REAL_RINV_UNIQ" + "REAL_RDISTRIB" > "HOLLight.hollight.REAL_RDISTRIB" + "REAL_POW_SUB" > "HOLLight.hollight.REAL_POW_SUB" + "REAL_POW_POW" > "HOLLight.hollight.REAL_POW_POW" + "REAL_POW_ONE" > "HOLLight.hollight.REAL_POW_ONE" + "REAL_POW_NZ" > "HOLLight.hollight.REAL_POW_NZ" + "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" > "HOLLight.hollight.REAL_POW_MONO" + "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" > "HOLLight.hollight.REAL_POW_LE2" + "REAL_POW_LE" > "HOLLight.hollight.REAL_POW_LE" + "REAL_POW_INV" > "HOLLight.hollight.REAL_POW_INV" + "REAL_POW_EQ_0" > "HOLLight.hollight.REAL_POW_EQ_0" + "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_LE" > "HOLLight.hollight.REAL_POW_1_LE" + "REAL_POW_1" > "HOLLight.hollight.REAL_POW_1" + "REAL_POW2_ABS" > "HOLLight.hollight.REAL_POW2_ABS" + "REAL_POS_NZ" > "HOLLight.hollight.REAL_POS_NZ" + "REAL_POSSQ" > "HOLLight.hollight.REAL_POSSQ" + "REAL_POS" > "HOLLight.hollight.REAL_POS" + "REAL_POLY_NEG_CLAUSES" > "HOLLight.hollight.REAL_POLY_NEG_CLAUSES" + "REAL_POLY_CLAUSES" > "HOLLight.hollight.REAL_POLY_CLAUSES" + "REAL_OF_NUM_SUM_NUMSEG" > "HOLLight.hollight.REAL_OF_NUM_SUM_NUMSEG" + "REAL_OF_NUM_SUM" > "HOLLight.hollight.REAL_OF_NUM_SUM" + "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_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" + "REAL_NZ_IMP_LT" > "HOLLight.hollight.REAL_NZ_IMP_LT" + "REAL_NOT_LT" > "HOLLight.hollight.REAL_NOT_LT" + "REAL_NOT_LE" > "HOLLight.hollight.real_lt_def" + "REAL_NOT_EQ" > "HOLLight.hollight.REAL_NOT_EQ" + "REAL_NEG_SUB" > "HOLLight.hollight.REAL_NEG_SUB" + "REAL_NEG_RMUL" > "HOLLight.hollight.REAL_NEG_RMUL" + "REAL_NEG_NEG" > "HOLLight.hollight.REAL_NEG_NEG" + "REAL_NEG_MUL2" > "HOLLight.hollight.REAL_NEG_MUL2" + "REAL_NEG_MINUS1" > "HOLLight.hollight.REAL_NEG_MINUS1" + "REAL_NEG_LT0" > "HOLLight.hollight.REAL_NEG_LT0" + "REAL_NEG_LMUL" > "HOLLight.hollight.REAL_NEG_LMUL" + "REAL_NEG_LE0" > "HOLLight.hollight.REAL_NEG_LE0" + "REAL_NEG_INV" > "HOLLight.hollight.REAL_NEG_INV" + "REAL_NEG_GT0" > "HOLLight.hollight.REAL_NEG_GT0" + "REAL_NEG_GE0" > "HOLLight.hollight.REAL_NEG_GE0" + "REAL_NEG_EQ_0" > "HOLLight.hollight.REAL_NEG_EQ_0" + "REAL_NEG_EQ0" > "HOLLight.hollight.REAL_NEG_EQ0" + "REAL_NEG_EQ" > "HOLLight.hollight.REAL_NEG_EQ" + "REAL_NEG_ADD" > "HOLLight.hollight.REAL_NEG_ADD" + "REAL_NEG_0" > "HOLLight.hollight.REAL_NEG_0" + "REAL_NEGNEG" > "HOLLight.hollight.REAL_NEGNEG" + "REAL_MUL_RZERO" > "HOLLight.hollight.REAL_MUL_RZERO" + "REAL_MUL_RNEG" > "HOLLight.hollight.REAL_MUL_RNEG" + "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_LZERO" > "HOLLight.hollight.REAL_MUL_LZERO" + "REAL_MUL_LNEG" > "HOLLight.hollight.REAL_MUL_LNEG" + "REAL_MUL_LINV_UNIQ" > "HOLLight.hollight.REAL_MUL_LINV_UNIQ" + "REAL_MUL_AC" > "HOLLight.hollight.REAL_MUL_AC" + "REAL_MUL_2" > "HOLLight.hollight.REAL_MUL_2" + "REAL_MUL" > "HOLLight.hollight.REAL_MUL" + "REAL_MIN_SYM" > "HOLLight.hollight.REAL_MIN_SYM" + "REAL_MIN_MIN" > "HOLLight.hollight.REAL_MIN_MIN" + "REAL_MIN_MAX" > "HOLLight.hollight.REAL_MIN_MAX" + "REAL_MIN_LT" > "HOLLight.hollight.REAL_MIN_LT" + "REAL_MIN_LE" > "HOLLight.hollight.REAL_MIN_LE" + "REAL_MIN_ASSOC" > "HOLLight.hollight.REAL_MIN_ASSOC" + "REAL_MIN_ACI" > "HOLLight.hollight.REAL_MIN_ACI" + "REAL_MIDDLE2" > "HOLLight.hollight.REAL_MIDDLE2" + "REAL_MIDDLE1" > "HOLLight.hollight.REAL_MIDDLE1" + "REAL_MEAN" > "HOLLight.hollight.REAL_MEAN" + "REAL_MAX_SYM" > "HOLLight.hollight.REAL_MAX_SYM" + "REAL_MAX_MIN" > "HOLLight.hollight.REAL_MAX_MIN" + "REAL_MAX_MAX" > "HOLLight.hollight.REAL_MAX_MAX" + "REAL_MAX_LT" > "HOLLight.hollight.REAL_MAX_LT" + "REAL_MAX_LE" > "HOLLight.hollight.REAL_MAX_LE" + "REAL_MAX_ASSOC" > "HOLLight.hollight.REAL_MAX_ASSOC" + "REAL_MAX_ACI" > "HOLLight.hollight.REAL_MAX_ACI" + "REAL_LT_TRANS" > "HOLLight.hollight.REAL_LT_TRANS" + "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" > "HOLLight.hollight.REAL_LT_SQUARE" + "REAL_LT_RNEG" > "HOLLight.hollight.REAL_LT_RNEG" + "REAL_LT_RMUL_IMP" > "HOLLight.hollight.REAL_LT_RMUL_IMP" + "REAL_LT_RMUL_EQ" > "HOLLight.hollight.REAL_LT_RMUL_EQ" + "REAL_LT_RMUL_0" > "HOLLight.hollight.REAL_LT_RMUL_0" + "REAL_LT_RMUL" > "HOLLight.hollight.REAL_LT_RMUL" + "REAL_LT_REFL" > "HOLLight.hollight.REAL_LT_REFL" + "REAL_LT_RDIV_EQ" > "HOLLight.hollight.REAL_LT_RDIV_EQ" + "REAL_LT_RDIV_0" > "HOLLight.hollight.REAL_LT_RDIV_0" + "REAL_LT_RDIV" > "HOLLight.hollight.REAL_LT_RDIV" + "REAL_LT_RCANCEL_IMP" > "HOLLight.hollight.REAL_LT_RCANCEL_IMP" + "REAL_LT_RADD" > "HOLLight.hollight.REAL_LT_RADD" + "REAL_LT_POW2" > "HOLLight.hollight.REAL_LT_POW2" + "REAL_LT_NZ" > "HOLLight.hollight.REAL_LT_NZ" + "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_MULTIPLE" > "HOLLight.hollight.REAL_LT_MULTIPLE" + "REAL_LT_MUL2_ALT" > "HOLLight.hollight.REAL_LT_MUL2_ALT" + "REAL_LT_MUL2" > "HOLLight.hollight.REAL_LT_MUL2" + "REAL_LT_MUL" > "HOLLight.hollight.REAL_LT_MUL" + "REAL_LT_MIN" > "HOLLight.hollight.REAL_LT_MIN" + "REAL_LT_MAX" > "HOLLight.hollight.REAL_LT_MAX" + "REAL_LT_LNEG" > "HOLLight.hollight.REAL_LT_LNEG" + "REAL_LT_LMUL_IMP" > "HOLLight.hollight.REAL_LT_LMUL_IMP" + "REAL_LT_LMUL_EQ" > "HOLLight.hollight.REAL_LT_LMUL_EQ" + "REAL_LT_LMUL_0" > "HOLLight.hollight.REAL_LT_LMUL_0" + "REAL_LT_LMUL" > "HOLLight.hollight.REAL_LT_LMUL" + "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" + "REAL_LT_LADD_IMP" > "HOLLight.hollight.REAL_LT_LADD_IMP" + "REAL_LT_LADD" > "HOLLight.hollight.REAL_LT_LADD" + "REAL_LT_INV_EQ" > "HOLLight.hollight.REAL_LT_INV_EQ" + "REAL_LT_INV2" > "HOLLight.hollight.REAL_LT_INV2" + "REAL_LT_INV" > "HOLLight.hollight.REAL_LT_INV" + "REAL_LT_IMP_NZ" > "HOLLight.hollight.REAL_LT_IMP_NZ" + "REAL_LT_IMP_NE" > "HOLLight.hollight.REAL_LT_IMP_NE" + "REAL_LT_IMP_LE" > "HOLLight.hollight.REAL_LT_IMP_LE" + "REAL_LT_HALF2" > "HOLLight.hollight.REAL_LT_HALF2" + "REAL_LT_HALF1" > "HOLLight.hollight.REAL_LT_HALF1" + "REAL_LT_GT" > "HOLLight.hollight.REAL_LT_GT" + "REAL_LT_FRACTION_0" > "HOLLight.hollight.REAL_LT_FRACTION_0" + "REAL_LT_FRACTION" > "HOLLight.hollight.REAL_LT_FRACTION" + "REAL_LT_DIV2_EQ" > "HOLLight.hollight.REAL_LT_DIV2_EQ" + "REAL_LT_DIV" > "HOLLight.hollight.REAL_LT_DIV" + "REAL_LT_ANTISYM" > "HOLLight.hollight.REAL_LT_ANTISYM" + "REAL_LT_ADD_SUB" > "HOLLight.hollight.REAL_LT_ADD_SUB" + "REAL_LT_ADDR" > "HOLLight.hollight.REAL_LT_ADDR" + "REAL_LT_ADDNEG2" > "HOLLight.hollight.REAL_LT_ADDNEG2" + "REAL_LT_ADDNEG" > "HOLLight.hollight.REAL_LT_ADDNEG" + "REAL_LT_ADDL" > "HOLLight.hollight.REAL_LT_ADDL" + "REAL_LT_ADD2" > "HOLLight.hollight.REAL_LT_ADD2" + "REAL_LT_ADD1" > "HOLLight.hollight.REAL_LT_ADD1" + "REAL_LT_ADD" > "HOLLight.hollight.REAL_LT_ADD" + "REAL_LT_1" > "HOLLight.hollight.REAL_LT_1" + "REAL_LT_01" > "HOLLight.hollight.REAL_LT_01" + "REAL_LTE_TRANS" > "HOLLight.hollight.REAL_LTE_TRANS" + "REAL_LTE_TOTAL" > "HOLLight.hollight.REAL_LTE_TOTAL" + "REAL_LTE_ANTISYM" > "HOLLight.hollight.REAL_LTE_ANTISYM" + "REAL_LTE_ADD2" > "HOLLight.hollight.REAL_LTE_ADD2" + "REAL_LTE_ADD" > "HOLLight.hollight.REAL_LTE_ADD" + "REAL_LT1_POW2" > "HOLLight.hollight.REAL_LT1_POW2" + "REAL_LT" > "HOLLight.hollight.REAL_LT" + "REAL_LNEG_UNIQ" > "HOLLight.hollight.REAL_LNEG_UNIQ" + "REAL_LINV_UNIQ" > "HOLLight.hollight.REAL_LINV_UNIQ" + "REAL_LE_TRANS" > "HOLLight.hollight.REAL_LE_TRANS" + "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_POW" > "HOLLight.hollight.REAL_LE_SQUARE_POW" + "REAL_LE_SQUARE_ABS" > "HOLLight.hollight.REAL_LE_SQUARE_ABS" + "REAL_LE_SQUARE" > "HOLLight.hollight.REAL_LE_SQUARE" + "REAL_LE_RSQRT" > "HOLLight.hollight.REAL_LE_RSQRT" + "REAL_LE_RNEG" > "HOLLight.hollight.REAL_LE_RNEG" + "REAL_LE_RMUL_IMP" > "HOLLight.hollight.REAL_LE_RMUL_IMP" + "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_RDIV_EQ" > "HOLLight.hollight.REAL_LE_RDIV_EQ" + "REAL_LE_RDIV" > "HOLLight.hollight.REAL_LE_RDIV" + "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_MUL2V" > "HOLLight.hollight.REAL_LE_MUL2V" + "REAL_LE_MUL2" > "HOLLight.hollight.REAL_LE_MUL2" + "REAL_LE_MUL" > "HOLLight.hollight.REAL_LE_MUL" + "REAL_LE_MIN" > "HOLLight.hollight.REAL_LE_MIN" + "REAL_LE_MAX" > "HOLLight.hollight.REAL_LE_MAX" + "REAL_LE_LT" > "HOLLight.hollight.REAL_LE_LT" + "REAL_LE_LSQRT" > "HOLLight.hollight.REAL_LE_LSQRT" + "REAL_LE_LNEG" > "HOLLight.hollight.REAL_LE_LNEG" + "REAL_LE_LMUL_LOCAL" > "HOLLight.hollight.REAL_LE_LMUL_LOCAL" + "REAL_LE_LMUL_IMP" > "HOLLight.hollight.REAL_LE_LMUL_IMP" + "REAL_LE_LMUL_EQ" > "HOLLight.hollight.REAL_LE_LMUL_EQ" + "REAL_LE_LMUL" > "HOLLight.hollight.REAL_LE_LMUL" + "REAL_LE_LDIV_EQ" > "HOLLight.hollight.REAL_LE_LDIV_EQ" + "REAL_LE_LDIV" > "HOLLight.hollight.REAL_LE_LDIV" + "REAL_LE_LCANCEL_IMP" > "HOLLight.hollight.REAL_LE_LCANCEL_IMP" + "REAL_LE_LADD" > "HOLLight.hollight.REAL_LE_LADD" + "REAL_LE_INV_EQ" > "HOLLight.hollight.REAL_LE_INV_EQ" + "REAL_LE_INV2" > "HOLLight.hollight.REAL_LE_INV2" + "REAL_LE_INV" > "HOLLight.hollight.REAL_LE_INV" + "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" + "REAL_LE_ADD" > "HOLLight.hollight.REAL_LE_ADD" + "REAL_LE_01" > "HOLLight.hollight.REAL_LE_01" + "REAL_LET_TRANS" > "HOLLight.hollight.REAL_LET_TRANS" + "REAL_LET_TOTAL" > "HOLLight.hollight.REAL_LET_TOTAL" + "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_LE1_POW2" > "HOLLight.hollight.REAL_LE1_POW2" + "REAL_LE" > "HOLLight.hollight.REAL_LE" + "REAL_INV_POS" > "HOLLight.hollight.REAL_INV_POS" + "REAL_INV_NZ" > "HOLLight.hollight.REAL_INV_NZ" + "REAL_INV_NEG" > "HOLLight.hollight.REAL_INV_NEG" + "REAL_INV_MUL_WEAK" > "HOLLight.hollight.REAL_INV_MUL_WEAK" + "REAL_INV_MUL" > "HOLLight.hollight.REAL_INV_MUL" + "REAL_INV_LT1" > "HOLLight.hollight.REAL_INV_LT1" + "REAL_INV_LE_1" > "HOLLight.hollight.REAL_INV_LE_1" + "REAL_INV_INV" > "HOLLight.hollight.REAL_INV_INV" + "REAL_INV_EQ_0" > "HOLLight.hollight.REAL_INV_EQ_0" + "REAL_INV_DIV" > "HOLLight.hollight.REAL_INV_DIV" + "REAL_INV_1_LE" > "HOLLight.hollight.REAL_INV_1_LE" + "REAL_INV_1OVER" > "HOLLight.hollight.REAL_INV_1OVER" + "REAL_INV_1" > "HOLLight.hollight.REAL_INV_1" + "REAL_INVINV" > "HOLLight.hollight.REAL_INVINV" + "REAL_INV1" > "HOLLight.hollight.REAL_INV1" + "REAL_INJ" > "HOLLight.hollight.REAL_INJ" + "REAL_HREAL_LEMMA2" > "HOLLight.hollight.REAL_HREAL_LEMMA2" + "REAL_HREAL_LEMMA1" > "HOLLight.hollight.REAL_HREAL_LEMMA1" + "REAL_HALF_DOUBLE" > "HOLLight.hollight.REAL_HALF_DOUBLE" + "REAL_FACT_NZ" > "HOLLight.hollight.REAL_FACT_NZ" + "REAL_EXP_TOTAL_LEMMA" > "HOLLight.hollight.REAL_EXP_TOTAL_LEMMA" + "REAL_EXP_TOTAL" > "HOLLight.hollight.REAL_EXP_TOTAL" + "REAL_EXP_SUB" > "HOLLight.hollight.REAL_EXP_SUB" + "REAL_EXP_POS_LT" > "HOLLight.hollight.REAL_EXP_POS_LT" + "REAL_EXP_POS_LE" > "HOLLight.hollight.REAL_EXP_POS_LE" + "REAL_EXP_NZ" > "HOLLight.hollight.REAL_EXP_NZ" + "REAL_EXP_NEG_MUL2" > "HOLLight.hollight.REAL_EXP_NEG_MUL2" + "REAL_EXP_NEG_MUL" > "HOLLight.hollight.REAL_EXP_NEG_MUL" + "REAL_EXP_NEG" > "HOLLight.hollight.REAL_EXP_NEG" + "REAL_EXP_N" > "HOLLight.hollight.REAL_EXP_N" + "REAL_EXP_MONO_LT" > "HOLLight.hollight.REAL_EXP_MONO_LT" + "REAL_EXP_MONO_LE" > "HOLLight.hollight.REAL_EXP_MONO_LE" + "REAL_EXP_MONO_IMP" > "HOLLight.hollight.REAL_EXP_MONO_IMP" + "REAL_EXP_LT_1" > "HOLLight.hollight.REAL_EXP_LT_1" + "REAL_EXP_LN" > "HOLLight.hollight.REAL_EXP_LN" + "REAL_EXP_LE_X" > "HOLLight.hollight.REAL_EXP_LE_X" + "REAL_EXP_INJ" > "HOLLight.hollight.REAL_EXP_INJ" + "REAL_EXP_FDIFF" > "HOLLight.hollight.REAL_EXP_FDIFF" + "REAL_EXP_CONVERGES" > "HOLLight.hollight.REAL_EXP_CONVERGES" + "REAL_EXP_BOUND_LEMMA" > "HOLLight.hollight.REAL_EXP_BOUND_LEMMA" + "REAL_EXP_ADD_MUL" > "HOLLight.hollight.REAL_EXP_ADD_MUL" + "REAL_EXP_ADD" > "HOLLight.hollight.REAL_EXP_ADD" + "REAL_EXP_0" > "HOLLight.hollight.REAL_EXP_0" + "REAL_EQ_SUB_RADD" > "HOLLight.hollight.REAL_EQ_SUB_RADD" + "REAL_EQ_SUB_LADD" > "HOLLight.hollight.REAL_EQ_SUB_LADD" + "REAL_EQ_RMUL_IMP" > "HOLLight.hollight.REAL_EQ_RMUL_IMP" + "REAL_EQ_RMUL" > "HOLLight.hollight.REAL_EQ_RMUL" + "REAL_EQ_RDIV_EQ" > "HOLLight.hollight.REAL_EQ_RDIV_EQ" + "REAL_EQ_RCANCEL_IMP" > "HOLLight.hollight.REAL_EQ_RCANCEL_IMP" + "REAL_EQ_RADD" > "HOLLight.hollight.REAL_EQ_RADD" + "REAL_EQ_NEG2" > "HOLLight.hollight.REAL_EQ_NEG2" + "REAL_EQ_NEG" > "HOLLight.hollight.REAL_EQ_NEG" + "REAL_EQ_MUL_RCANCEL" > "HOLLight.hollight.REAL_EQ_MUL_RCANCEL" + "REAL_EQ_MUL_LCANCEL" > "HOLLight.hollight.REAL_EQ_MUL_LCANCEL" + "REAL_EQ_LMUL_IMP" > "HOLLight.hollight.REAL_EQ_LMUL_IMP" + "REAL_EQ_LMUL2" > "HOLLight.hollight.REAL_EQ_LMUL2" + "REAL_EQ_LMUL" > "HOLLight.hollight.REAL_EQ_LMUL" + "REAL_EQ_LDIV_EQ" > "HOLLight.hollight.REAL_EQ_LDIV_EQ" + "REAL_EQ_LCANCEL_IMP" > "HOLLight.hollight.REAL_EQ_LCANCEL_IMP" + "REAL_EQ_LADD" > "HOLLight.hollight.REAL_EQ_LADD" + "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" + "REAL_EQ_ADD_LCANCEL_0" > "HOLLight.hollight.REAL_EQ_ADD_LCANCEL_0" + "REAL_EQ_ADD_LCANCEL" > "HOLLight.hollight.REAL_EQ_ADD_LCANCEL" + "REAL_ENTIRE" > "HOLLight.hollight.REAL_ENTIRE" + "REAL_DOWN2" > "HOLLight.hollight.REAL_DOWN2" + "REAL_DOWN" > "HOLLight.hollight.REAL_DOWN" + "REAL_DOUBLE" > "HOLLight.hollight.REAL_DOUBLE" + "REAL_DIV_SQRT" > "HOLLight.hollight.REAL_DIV_SQRT" + "REAL_DIV_RMUL" > "HOLLight.hollight.REAL_DIV_RMUL" + "REAL_DIV_REFL" > "HOLLight.hollight.REAL_DIV_REFL" + "REAL_DIV_POW2_ALT" > "HOLLight.hollight.REAL_DIV_POW2_ALT" + "REAL_DIV_POW2" > "HOLLight.hollight.REAL_DIV_POW2" + "REAL_DIV_MUL2" > "HOLLight.hollight.REAL_DIV_MUL2" + "REAL_DIV_LZERO" > "HOLLight.hollight.REAL_DIV_LZERO" + "REAL_DIV_LMUL" > "HOLLight.hollight.REAL_DIV_LMUL" + "REAL_DIV_1" > "HOLLight.hollight.REAL_DIV_1" + "REAL_DIFFSQ" > "HOLLight.hollight.REAL_DIFFSQ" + "REAL_COMPLETE_SOMEPOS" > "HOLLight.hollight.REAL_COMPLETE_SOMEPOS" + "REAL_COMPLETE" > "HOLLight.hollight.REAL_COMPLETE" + "REAL_ATN_POWSER_SUMMABLE" > "HOLLight.hollight.REAL_ATN_POWSER_SUMMABLE" + "REAL_ATN_POWSER_DIFFS_SUMMABLE" > "HOLLight.hollight.REAL_ATN_POWSER_DIFFS_SUMMABLE" + "REAL_ATN_POWSER_DIFFS_SUM" > "HOLLight.hollight.REAL_ATN_POWSER_DIFFS_SUM" + "REAL_ATN_POWSER_DIFFS_DIFFS_SUMMABLE" > "HOLLight.hollight.REAL_ATN_POWSER_DIFFS_DIFFS_SUMMABLE" + "REAL_ATN_POWSER_DIFFL" > "HOLLight.hollight.REAL_ATN_POWSER_DIFFL" + "REAL_ATN_POWSER" > "HOLLight.hollight.REAL_ATN_POWSER" + "REAL_ARCH_SIMPLE" > "HOLLight.hollight.REAL_ARCH_SIMPLE" + "REAL_ARCH_LEAST" > "HOLLight.hollight.REAL_ARCH_LEAST" + "REAL_ARCH" > "HOLLight.hollight.REAL_ARCH" + "REAL_ADD_SUB2" > "HOLLight.hollight.REAL_ADD_SUB2" + "REAL_ADD_SUB" > "HOLLight.hollight.REAL_ADD_SUB" + "REAL_ADD_RINV" > "HOLLight.hollight.REAL_ADD_RINV" + "REAL_ADD_RID_UNIQ" > "HOLLight.hollight.REAL_ADD_RID_UNIQ" + "REAL_ADD_RID" > "HOLLight.hollight.REAL_ADD_RID" + "REAL_ADD_RDISTRIB" > "HOLLight.hollight.REAL_ADD_RDISTRIB" + "REAL_ADD_LID_UNIQ" > "HOLLight.hollight.REAL_ADD_LID_UNIQ" + "REAL_ADD_AC" > "HOLLight.hollight.REAL_ADD_AC" + "REAL_ADD2_SUB2" > "HOLLight.hollight.REAL_ADD2_SUB2" + "REAL_ADD" > "HOLLight.hollight.REAL_ADD" + "REAL_ABS_ZERO" > "HOLLight.hollight.REAL_ABS_ZERO" + "REAL_ABS_TRIANGLE_LT" > "HOLLight.hollight.REAL_ABS_TRIANGLE_LT" + "REAL_ABS_TRIANGLE_LE" > "HOLLight.hollight.REAL_ABS_TRIANGLE_LE" + "REAL_ABS_TRIANGLE" > "HOLLight.hollight.REAL_ABS_TRIANGLE" + "REAL_ABS_SUB_ABS" > "HOLLight.hollight.REAL_ABS_SUB_ABS" + "REAL_ABS_SUB" > "HOLLight.hollight.REAL_ABS_SUB" + "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_REFL" > "HOLLight.hollight.REAL_ABS_REFL" + "REAL_ABS_POW" > "HOLLight.hollight.REAL_ABS_POW" + "REAL_ABS_POS" > "HOLLight.hollight.REAL_ABS_POS" + "REAL_ABS_NZ" > "HOLLight.hollight.REAL_ABS_NZ" + "REAL_ABS_NUM" > "HOLLight.hollight.REAL_ABS_NUM" + "REAL_ABS_NEG" > "HOLLight.hollight.REAL_ABS_NEG" + "REAL_ABS_MUL" > "HOLLight.hollight.REAL_ABS_MUL" + "REAL_ABS_LE" > "HOLLight.hollight.REAL_ABS_LE" + "REAL_ABS_INV" > "HOLLight.hollight.REAL_ABS_INV" + "REAL_ABS_DIV" > "HOLLight.hollight.REAL_ABS_DIV" + "REAL_ABS_CIRCLE" > "HOLLight.hollight.REAL_ABS_CIRCLE" + "REAL_ABS_CASES" > "HOLLight.hollight.REAL_ABS_CASES" + "REAL_ABS_BOUNDS" > "HOLLight.hollight.REAL_ABS_BOUNDS" + "REAL_ABS_BOUND" > "HOLLight.hollight.REAL_ABS_BOUND" + "REAL_ABS_BETWEEN2" > "HOLLight.hollight.REAL_ABS_BETWEEN2" + "REAL_ABS_BETWEEN1" > "HOLLight.hollight.REAL_ABS_BETWEEN1" + "REAL_ABS_BETWEEN" > "HOLLight.hollight.REAL_ABS_BETWEEN" + "REAL_ABS_ABS" > "HOLLight.hollight.REAL_ABS_ABS" + "REAL_ABS_1" > "HOLLight.hollight.REAL_ABS_1" + "REAL_ABS_0" > "HOLLight.hollight.REAL_ABS_0" + "REAL" > "HOLLight.hollight.REAL" + "RAT_LEMMA5" > "HOLLight.hollight.RAT_LEMMA5" + "RAT_LEMMA4" > "HOLLight.hollight.RAT_LEMMA4" + "RAT_LEMMA3" > "HOLLight.hollight.RAT_LEMMA3" + "RAT_LEMMA2" > "HOLLight.hollight.RAT_LEMMA2" + "RAT_LEMMA1" > "HOLLight.hollight.RAT_LEMMA1" + "PSUM_SUM_NUMSEG" > "HOLLight.hollight.PSUM_SUM_NUMSEG" + "PSUM_SUM" > "HOLLight.hollight.PSUM_SUM" + "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_MEMBER" > "HOLLight.hollight.PSUBSET_MEMBER" + "PSUBSET_IRREFL" > "HOLLight.hollight.PSUBSET_IRREFL" + "PSUBSET_INSERT_SUBSET" > "HOLLight.hollight.PSUBSET_INSERT_SUBSET" + "PRE_ELIM_THM" > "HOLLight.hollight.PRE_ELIM_THM" + "POW_ZERO_EQ" > "HOLLight.hollight.POW_ZERO_EQ" + "POW_ZERO" > "HOLLight.hollight.POW_ZERO" + "POW_ROOT_POS" > "HOLLight.hollight.POW_ROOT_POS" + "POW_POS_LT" > "HOLLight.hollight.POW_POS_LT" + "POW_POS" > "HOLLight.hollight.POW_POS" + "POW_PLUS1" > "HOLLight.hollight.POW_PLUS1" + "POW_NZ" > "HOLLight.hollight.POW_NZ" + "POW_MUL" > "HOLLight.hollight.POW_MUL" + "POW_MINUS1" > "HOLLight.hollight.POW_MINUS1" + "POW_M1" > "HOLLight.hollight.POW_M1" + "POW_LT" > "HOLLight.hollight.POW_LT" + "POW_LE" > "HOLLight.hollight.POW_LE" + "POW_INV" > "HOLLight.hollight.POW_INV" + "POW_EQ" > "HOLLight.hollight.POW_EQ" + "POW_ADD" > "HOLLight.hollight.POW_ADD" + "POW_ABS" > "HOLLight.hollight.POW_ABS" + "POW_2_SQRT_ABS" > "HOLLight.hollight.POW_2_SQRT_ABS" + "POW_2_SQRT" > "HOLLight.hollight.POW_2_SQRT" + "POW_2_LT" > "HOLLight.hollight.POW_2_LT" + "POW_2_LE1" > "HOLLight.hollight.POW_2_LE1" + "POW_2" > "HOLLight.hollight.POW_2" + "POW_1" > "HOLLight.hollight.POW_1" + "POW_0" > "HOLLight.hollight.POW_0" + "POWSER_LIMIT_0_STRONG" > "HOLLight.hollight.POWSER_LIMIT_0_STRONG" + "POWSER_LIMIT_0" > "HOLLight.hollight.POWSER_LIMIT_0" + "POWSER_INSIDEA" > "HOLLight.hollight.POWSER_INSIDEA" + "POWSER_INSIDE" > "HOLLight.hollight.POWSER_INSIDE" + "POWSER_EQUAL_0" > "HOLLight.hollight.POWSER_EQUAL_0" + "POWSER_EQUAL" > "HOLLight.hollight.POWSER_EQUAL" + "POWSER_0" > "HOLLight.hollight.POWSER_0" + "POWREV" > "HOLLight.hollight.POWREV" + "POWDIFF_LEMMA" > "HOLLight.hollight.POWDIFF_LEMMA" + "POWDIFF" > "HOLLight.hollight.POWDIFF" + "PI_POS" > "HOLLight.hollight.PI_POS" + "PI2_PI4" > "HOLLight.hollight.PI2_PI4" + "PI2_BOUNDS" > "HOLLight.hollight.PI2_BOUNDS" + "PI2" > "HOLLight.hollight.PI2" + "PASTECART_FST_SND" > "HOLLight.hollight.PASTECART_FST_SND" + "PASTECART_EQ" > "HOLLight.hollight.PASTECART_EQ" + "PASSOC_def" > "HOLLight.hollight.PASSOC_def" + "PAIR_SURJECTIVE" > "Datatype.prod.nchotomy" + "PAIR_EXISTS_THM" > "HOLLight.hollight.PAIR_EXISTS_THM" + "PAIR_EQ" > "Datatype.prod.simps_1" + "PAIRWISE_def" > "HOLLight.hollight.PAIRWISE_def" + "PAIR" > "HOL4Compat.PAIR" + "OUTR_def" > "HOLLight.hollight.OUTR_def" + "OUTL_def" > "HOLLight.hollight.OUTL_def" + "OR_EXISTS_THM" > "HOL.ex_disj_distrib" + "OR_CLAUSES" > "HOLLight.hollight.OR_CLAUSES" + "OPEN_UNOPEN" > "HOLLight.hollight.OPEN_UNOPEN" + "OPEN_SUBOPEN" > "HOLLight.hollight.OPEN_SUBOPEN" + "OPEN_OWN_NEIGH" > "HOLLight.hollight.OPEN_OWN_NEIGH" + "OPEN_NEIGH" > "HOLLight.hollight.OPEN_NEIGH" + "ONE" > "HOLLight.hollight.ONE" + "ODD_def" > "HOLLight.hollight.ODD_def" + "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_DOUBLE" > "HOLLight.hollight.ODD_DOUBLE" + "ODD_ADD" > "HOLLight.hollight.ODD_ADD" + "NUM_INTEGRAL_LEMMA" > "HOLLight.hollight.NUM_INTEGRAL_LEMMA" + "NUM_INTEGRAL" > "HOLLight.hollight.NUM_INTEGRAL" + "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_RREC" > "HOLLight.hollight.NUMSEG_RREC" + "NUMSEG_REC" > "HOLLight.hollight.NUMSEG_REC" + "NUMSEG_OFFSET_IMAGE" > "HOLLight.hollight.NUMSEG_OFFSET_IMAGE" + "NUMSEG_LREC" > "HOLLight.hollight.NUMSEG_LREC" + "NUMSEG_EMPTY" > "HOLLight.hollight.NUMSEG_EMPTY" + "NUMSEG_COMBINE_R" > "HOLLight.hollight.NUMSEG_COMBINE_R" + "NUMSEG_COMBINE_L" > "HOLLight.hollight.NUMSEG_COMBINE_L" + "NUMSEG_CLAUSES" > "HOLLight.hollight.NUMSEG_CLAUSES" + "NUMSEG_ADD_SPLIT" > "HOLLight.hollight.NUMSEG_ADD_SPLIT" + "NUMRIGHT_def" > "HOLLight.hollight.NUMRIGHT_def" + "NUMPAIR_def" > "HOLLight.hollight.NUMPAIR_def" + "NUMPAIR_INJ_LEMMA" > "HOLLight.hollight.NUMPAIR_INJ_LEMMA" + "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_LZERO" > "HOLLight.hollight.NSUM_UNION_LZERO" + "NSUM_UNION_EQ" > "HOLLight.hollight.NSUM_UNION_EQ" + "NSUM_UNION" > "HOLLight.hollight.NSUM_UNION" + "NSUM_TRIV_NUMSEG" > "HOLLight.hollight.NSUM_TRIV_NUMSEG" + "NSUM_SWAP_NUMSEG" > "HOLLight.hollight.NSUM_SWAP_NUMSEG" + "NSUM_SWAP" > "HOLLight.hollight.NSUM_SWAP" + "NSUM_SUPPORT" > "HOLLight.hollight.NSUM_SUPPORT" + "NSUM_SUPERSET" > "HOLLight.hollight.NSUM_SUPERSET" + "NSUM_SUBSET_SIMPLE" > "HOLLight.hollight.NSUM_SUBSET_SIMPLE" + "NSUM_SUBSET" > "HOLLight.hollight.NSUM_SUBSET" + "NSUM_SING_NUMSEG" > "HOLLight.hollight.NSUM_SING_NUMSEG" + "NSUM_SING" > "HOLLight.hollight.NSUM_SING" + "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_OFFSET_0" > "HOLLight.hollight.NSUM_OFFSET_0" + "NSUM_OFFSET" > "HOLLight.hollight.NSUM_OFFSET" + "NSUM_NSUM_RESTRICT" > "HOLLight.hollight.NSUM_NSUM_RESTRICT" + "NSUM_MULTICOUNT_GEN" > "HOLLight.hollight.NSUM_MULTICOUNT_GEN" + "NSUM_MULTICOUNT" > "HOLLight.hollight.NSUM_MULTICOUNT" + "NSUM_LT_ALL" > "HOLLight.hollight.NSUM_LT_ALL" + "NSUM_LT" > "HOLLight.hollight.NSUM_LT" + "NSUM_LE_NUMSEG" > "HOLLight.hollight.NSUM_LE_NUMSEG" + "NSUM_LE" > "HOLLight.hollight.NSUM_LE" + "NSUM_IMAGE_GEN" > "HOLLight.hollight.NSUM_IMAGE_GEN" + "NSUM_IMAGE" > "HOLLight.hollight.NSUM_IMAGE" + "NSUM_EQ_SUPERSET" > "HOLLight.hollight.NSUM_EQ_SUPERSET" + "NSUM_EQ_NUMSEG" > "HOLLight.hollight.NSUM_EQ_NUMSEG" + "NSUM_EQ_0_NUMSEG" > "HOLLight.hollight.NSUM_EQ_0_NUMSEG" + "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_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_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_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" + "NSUM_BOUND_GEN" > "HOLLight.hollight.NSUM_BOUND_GEN" + "NSUM_BOUND" > "HOLLight.hollight.NSUM_BOUND" + "NSUM_ADD_SPLIT" > "HOLLight.hollight.NSUM_ADD_SPLIT" + "NSUM_ADD_NUMSEG" > "HOLLight.hollight.NSUM_ADD_NUMSEG" + "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_ODD" > "HOLLight.hollight.NOT_ODD" + "NOT_LT" > "HOLLight.hollight.NOT_LT" + "NOT_LE" > "HOLLight.hollight.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_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_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" + "NEUTRAL_REAL_MUL" > "HOLLight.hollight.NEUTRAL_REAL_MUL" + "NEUTRAL_REAL_ADD" > "HOLLight.hollight.NEUTRAL_REAL_ADD" + "NEUTRAL_MUL" > "HOLLight.hollight.NEUTRAL_MUL" + "NEUTRAL_ADD" > "HOLLight.hollight.NEUTRAL_ADD" + "NET_SUM" > "HOLLight.hollight.NET_SUM" + "NET_SUB" > "HOLLight.hollight.NET_SUB" + "NET_NULL_MUL" > "HOLLight.hollight.NET_NULL_MUL" + "NET_NULL_CMUL" > "HOLLight.hollight.NET_NULL_CMUL" + "NET_NULL_ADD" > "HOLLight.hollight.NET_NULL_ADD" + "NET_NULL" > "HOLLight.hollight.NET_NULL" + "NET_NEG" > "HOLLight.hollight.NET_NEG" + "NET_MUL" > "HOLLight.hollight.NET_MUL" + "NET_LE" > "HOLLight.hollight.NET_LE" + "NET_INV" > "HOLLight.hollight.NET_INV" + "NET_DIV" > "HOLLight.hollight.NET_DIV" + "NET_CONV_NZ" > "HOLLight.hollight.NET_CONV_NZ" + "NET_CONV_IBOUNDED" > "HOLLight.hollight.NET_CONV_IBOUNDED" + "NET_CONV_BOUNDED" > "HOLLight.hollight.NET_CONV_BOUNDED" + "NET_ADD" > "HOLLight.hollight.NET_ADD" + "NET_ABS" > "HOLLight.hollight.NET_ABS" + "NEST_LEMMA_UNIQ" > "HOLLight.hollight.NEST_LEMMA_UNIQ" + "NEST_LEMMA" > "HOLLight.hollight.NEST_LEMMA" + "NADD_UBOUND" > "HOLLight.hollight.NADD_UBOUND" + "NADD_SUC" > "HOLLight.hollight.NADD_SUC" + "NADD_RDISTRIB" > "HOLLight.hollight.NADD_RDISTRIB" + "NADD_OF_NUM_WELLDEF" > "HOLLight.hollight.NADD_OF_NUM_WELLDEF" + "NADD_OF_NUM_MUL" > "HOLLight.hollight.NADD_OF_NUM_MUL" + "NADD_OF_NUM_LE" > "HOLLight.hollight.NADD_OF_NUM_LE" + "NADD_OF_NUM_EQ" > "HOLLight.hollight.NADD_OF_NUM_EQ" + "NADD_OF_NUM_ADD" > "HOLLight.hollight.NADD_OF_NUM_ADD" + "NADD_OF_NUM" > "HOLLight.hollight.NADD_OF_NUM" + "NADD_NONZERO" > "HOLLight.hollight.NADD_NONZERO" + "NADD_MUL_WELLDEF_LEMMA" > "HOLLight.hollight.NADD_MUL_WELLDEF_LEMMA" + "NADD_MUL_WELLDEF" > "HOLLight.hollight.NADD_MUL_WELLDEF" + "NADD_MUL_SYM" > "HOLLight.hollight.NADD_MUL_SYM" + "NADD_MUL_LINV_LEMMA8" > "HOLLight.hollight.NADD_MUL_LINV_LEMMA8" + "NADD_MUL_LINV_LEMMA7a" > "HOLLight.hollight.NADD_MUL_LINV_LEMMA7a" + "NADD_MUL_LINV_LEMMA7" > "HOLLight.hollight.NADD_MUL_LINV_LEMMA7" + "NADD_MUL_LINV_LEMMA6" > "HOLLight.hollight.NADD_MUL_LINV_LEMMA6" + "NADD_MUL_LINV_LEMMA5" > "HOLLight.hollight.NADD_MUL_LINV_LEMMA5" + "NADD_MUL_LINV_LEMMA4" > "HOLLight.hollight.NADD_MUL_LINV_LEMMA4" + "NADD_MUL_LINV_LEMMA3" > "HOLLight.hollight.NADD_MUL_LINV_LEMMA3" + "NADD_MUL_LINV_LEMMA2" > "HOLLight.hollight.NADD_MUL_LINV_LEMMA2" + "NADD_MUL_LINV_LEMMA1" > "HOLLight.hollight.NADD_MUL_LINV_LEMMA1" + "NADD_MUL_LINV_LEMMA0" > "HOLLight.hollight.NADD_MUL_LINV_LEMMA0" + "NADD_MUL_LINV" > "HOLLight.hollight.NADD_MUL_LINV" + "NADD_MUL_LID" > "HOLLight.hollight.NADD_MUL_LID" + "NADD_MUL_ASSOC" > "HOLLight.hollight.NADD_MUL_ASSOC" + "NADD_MULTIPLICATIVE" > "HOLLight.hollight.NADD_MULTIPLICATIVE" + "NADD_MUL" > "HOLLight.hollight.NADD_MUL" + "NADD_LE_WELLDEF_LEMMA" > "HOLLight.hollight.NADD_LE_WELLDEF_LEMMA" + "NADD_LE_WELLDEF" > "HOLLight.hollight.NADD_LE_WELLDEF" + "NADD_LE_TRANS" > "HOLLight.hollight.NADD_LE_TRANS" + "NADD_LE_TOTAL_LEMMA" > "HOLLight.hollight.NADD_LE_TOTAL_LEMMA" + "NADD_LE_TOTAL" > "HOLLight.hollight.NADD_LE_TOTAL" + "NADD_LE_RMUL" > "HOLLight.hollight.NADD_LE_RMUL" + "NADD_LE_REFL" > "HOLLight.hollight.NADD_LE_REFL" + "NADD_LE_RADD" > "HOLLight.hollight.NADD_LE_RADD" + "NADD_LE_LMUL" > "HOLLight.hollight.NADD_LE_LMUL" + "NADD_LE_LADD" > "HOLLight.hollight.NADD_LE_LADD" + "NADD_LE_EXISTS" > "HOLLight.hollight.NADD_LE_EXISTS" + "NADD_LE_ANTISYM" > "HOLLight.hollight.NADD_LE_ANTISYM" + "NADD_LE_ADD" > "HOLLight.hollight.NADD_LE_ADD" + "NADD_LE_0" > "HOLLight.hollight.NADD_LE_0" + "NADD_LDISTRIB" > "HOLLight.hollight.NADD_LDISTRIB" + "NADD_LBOUND" > "HOLLight.hollight.NADD_LBOUND" + "NADD_INV_WELLDEF" > "HOLLight.hollight.NADD_INV_WELLDEF" + "NADD_INV_0" > "HOLLight.hollight.NADD_INV_0" + "NADD_INV" > "HOLLight.hollight.NADD_INV" + "NADD_EQ_TRANS" > "HOLLight.hollight.NADD_EQ_TRANS" + "NADD_EQ_SYM" > "HOLLight.hollight.NADD_EQ_SYM" + "NADD_EQ_REFL" > "HOLLight.hollight.NADD_EQ_REFL" + "NADD_EQ_IMP_LE" > "HOLLight.hollight.NADD_EQ_IMP_LE" + "NADD_DIST_LEMMA" > "HOLLight.hollight.NADD_DIST_LEMMA" + "NADD_DIST" > "HOLLight.hollight.NADD_DIST" + "NADD_COMPLETE" > "HOLLight.hollight.NADD_COMPLETE" + "NADD_CAUCHY" > "HOLLight.hollight.NADD_CAUCHY" + "NADD_BOUND" > "HOLLight.hollight.NADD_BOUND" + "NADD_ARCH_ZERO" > "HOLLight.hollight.NADD_ARCH_ZERO" + "NADD_ARCH_MULT" > "HOLLight.hollight.NADD_ARCH_MULT" + "NADD_ARCH_LEMMA" > "HOLLight.hollight.NADD_ARCH_LEMMA" + "NADD_ARCH" > "HOLLight.hollight.NADD_ARCH" + "NADD_ALTMUL" > "HOLLight.hollight.NADD_ALTMUL" + "NADD_ADD_WELLDEF" > "HOLLight.hollight.NADD_ADD_WELLDEF" + "NADD_ADD_SYM" > "HOLLight.hollight.NADD_ADD_SYM" + "NADD_ADD_LID" > "HOLLight.hollight.NADD_ADD_LID" + "NADD_ADD_LCANCEL" > "HOLLight.hollight.NADD_ADD_LCANCEL" + "NADD_ADD_ASSOC" > "HOLLight.hollight.NADD_ADD_ASSOC" + "NADD_ADDITIVE" > "HOLLight.hollight.NADD_ADDITIVE" + "NADD_ADD" > "HOLLight.hollight.NADD_ADD" + "MVT_LEMMA" > "HOLLight.hollight.MVT_LEMMA" + "MVT_ALT" > "HOLLight.hollight.MVT_ALT" + "MVT" > "HOLLight.hollight.MVT" + "MULT_SYM" > "IntDef.zmult_ac_2" + "MULT_SUC" > "Nat.mult_Suc_right" + "MULT_EXP" > "HOLLight.hollight.MULT_EXP" + "MULT_EQ_1" > "HOLLight.hollight.MULT_EQ_1" + "MULT_EQ_0" > "Nat.mult_is_0" + "MULT_DIV_2" > "HOLLight.hollight.MULT_DIV_2" + "MULT_CLAUSES" > "HOLLight.hollight.MULT_CLAUSES" + "MULT_ASSOC" > "IntDef.zmult_ac_1" + "MULT_AC" > "HOLLight.hollight.MULT_AC" + "MULT_2" > "HOLLight.hollight.MULT_2" + "MULT_0" > "Nat.mult_0_right" + "MTOP_TENDS_UNIQ" > "HOLLight.hollight.MTOP_TENDS_UNIQ" + "MTOP_TENDS" > "HOLLight.hollight.MTOP_TENDS" + "MTOP_OPEN" > "HOLLight.hollight.MTOP_OPEN" + "MTOP_LIMPT" > "HOLLight.hollight.MTOP_LIMPT" + "MR1_SUB_LT" > "HOLLight.hollight.MR1_SUB_LT" + "MR1_SUB_LE" > "HOLLight.hollight.MR1_SUB_LE" + "MR1_SUB" > "HOLLight.hollight.MR1_SUB" + "MR1_LIMPT" > "HOLLight.hollight.MR1_LIMPT" + "MR1_DEF" > "HOLLight.hollight.MR1_DEF" + "MR1_BOUNDED" > "HOLLight.hollight.MR1_BOUNDED" + "MR1_BETWEEN1" > "HOLLight.hollight.MR1_BETWEEN1" + "MR1_ADD_LT" > "HOLLight.hollight.MR1_ADD_LT" + "MR1_ADD_LE" > "HOLLight.hollight.MR1_ADD_LE" + "MR1_ADD" > "HOLLight.hollight.MR1_ADD" + "MONO_SUC" > "HOLLight.hollight.MONO_SUC" + "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" + "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" + "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_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_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" + "METRIC_ZERO" > "HOLLight.hollight.METRIC_ZERO" + "METRIC_TRIANGLE" > "HOLLight.hollight.METRIC_TRIANGLE" + "METRIC_SYM" > "HOLLight.hollight.METRIC_SYM" + "METRIC_SAME" > "HOLLight.hollight.METRIC_SAME" + "METRIC_POS" > "HOLLight.hollight.METRIC_POS" + "METRIC_NZ" > "HOLLight.hollight.METRIC_NZ" + "METRIC_ISMET" > "HOLLight.hollight.METRIC_ISMET" + "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" + "MEASURE_LE" > "HOLLight.hollight.MEASURE_LE" + "MCLAURIN_ZERO" > "HOLLight.hollight.MCLAURIN_ZERO" + "MCLAURIN_SIN" > "HOLLight.hollight.MCLAURIN_SIN" + "MCLAURIN_NEG" > "HOLLight.hollight.MCLAURIN_NEG" + "MCLAURIN_LN_POS" > "HOLLight.hollight.MCLAURIN_LN_POS" + "MCLAURIN_LN_NEG" > "HOLLight.hollight.MCLAURIN_LN_NEG" + "MCLAURIN_EXP_LT" > "HOLLight.hollight.MCLAURIN_EXP_LT" + "MCLAURIN_EXP_LEMMA" > "HOLLight.hollight.MCLAURIN_EXP_LEMMA" + "MCLAURIN_EXP_LE" > "HOLLight.hollight.MCLAURIN_EXP_LE" + "MCLAURIN_COS" > "HOLLight.hollight.MCLAURIN_COS" + "MCLAURIN_BI_LE" > "HOLLight.hollight.MCLAURIN_BI_LE" + "MCLAURIN_ATN" > "HOLLight.hollight.MCLAURIN_ATN" + "MCLAURIN_ALL_LT" > "HOLLight.hollight.MCLAURIN_ALL_LT" + "MCLAURIN_ALL_LE" > "HOLLight.hollight.MCLAURIN_ALL_LE" + "MCLAURIN" > "HOLLight.hollight.MCLAURIN" + "MAX_LEMMA" > "HOLLight.hollight.MAX_LEMMA" + "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" + "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_LMULT" > "HOLLight.hollight.LT_LMULT" + "LT_LE" > "HOLLight.hollight.LT_LE" + "LT_IMP_LE" > "HOLLight.hollight.LT_IMP_LE" + "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_ADDR" > "HOLLight.hollight.LT_ADDR" + "LT_ADD2" > "HOLLight.hollight.LT_ADD2" + "LT_ADD" > "HOLLight.hollight.LT_ADD" + "LT_0" > "HOLLight.hollight.LT_0" + "LTE_TRANS" > "HOLLight.hollight.LTE_TRANS" + "LTE_CASES" > "HOLLight.hollight.LTE_CASES" + "LTE_ANTISYM" > "HOLLight.hollight.LTE_ANTISYM" + "LTE_ADD2" > "HOLLight.hollight.LTE_ADD2" + "LN_POW" > "HOLLight.hollight.LN_POW" + "LN_POS_LT" > "HOLLight.hollight.LN_POS_LT" + "LN_POS" > "HOLLight.hollight.LN_POS" + "LN_MUL" > "HOLLight.hollight.LN_MUL" + "LN_MONO_LT" > "HOLLight.hollight.LN_MONO_LT" + "LN_MONO_LE" > "HOLLight.hollight.LN_MONO_LE" + "LN_LT_X" > "HOLLight.hollight.LN_LT_X" + "LN_LE" > "HOLLight.hollight.LN_LE" + "LN_INV" > "HOLLight.hollight.LN_INV" + "LN_INJ" > "HOLLight.hollight.LN_INJ" + "LN_EXP" > "HOLLight.hollight.LN_EXP" + "LN_DIV" > "HOLLight.hollight.LN_DIV" + "LN_1" > "HOLLight.hollight.LN_1" + "LIST_OF_SET_PROPERTIES" > "HOLLight.hollight.LIST_OF_SET_PROPERTIES" + "LIM_X" > "HOLLight.hollight.LIM_X" + "LIM_UNIQ" > "HOLLight.hollight.LIM_UNIQ" + "LIM_TRANSFORM" > "HOLLight.hollight.LIM_TRANSFORM" + "LIM_TENDS2" > "HOLLight.hollight.LIM_TENDS2" + "LIM_TENDS" > "HOLLight.hollight.LIM_TENDS" + "LIM_SUM" > "HOLLight.hollight.LIM_SUM" + "LIM_SUB" > "HOLLight.hollight.LIM_SUB" + "LIM_NULL" > "HOLLight.hollight.LIM_NULL" + "LIM_NEG" > "HOLLight.hollight.LIM_NEG" + "LIM_MUL" > "HOLLight.hollight.LIM_MUL" + "LIM_INV" > "HOLLight.hollight.LIM_INV" + "LIM_EQUAL" > "HOLLight.hollight.LIM_EQUAL" + "LIM_DIV" > "HOLLight.hollight.LIM_DIV" + "LIM_CONST" > "HOLLight.hollight.LIM_CONST" + "LIM_ADD" > "HOLLight.hollight.LIM_ADD" + "LIM" > "HOLLight.hollight.LIM" + "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" + "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_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" + "LET_END_def" > "HOLLight.hollight.LET_END_def" + "LET_CASES" > "HOLLight.hollight.LET_CASES" + "LET_ANTISYM" > "HOLLight.hollight.LET_ANTISYM" + "LET_ADD2" > "HOLLight.hollight.LET_ADD2" + "LESS_SUC_EQ" > "HOLLight.hollight.LESS_SUC_EQ" + "LESS_1" > "HOLLight.hollight.LESS_1" + "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" + "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_FORALL_OR_THM" > "HOL.all_simps_3" + "LEFT_FORALL_IMP_THM" > "HOL.imp_ex" + "LEFT_EXISTS_IMP_THM" > "HOL.imp_all" + "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" + "LAMBDA_UNIQUE" > "HOLLight.hollight.LAMBDA_UNIQUE" + "LAMBDA_ETA" > "HOLLight.hollight.LAMBDA_ETA" + "LAMBDA_BETA" > "HOLLight.hollight.LAMBDA_BETA" + "I_THM" > "Fun.id_apply" + "I_O_ID" > "HOLLight.hollight.I_O_ID" + "IVT_DERIVATIVE_POS" > "HOLLight.hollight.IVT_DERIVATIVE_POS" + "IVT_DERIVATIVE_NEG" > "HOLLight.hollight.IVT_DERIVATIVE_NEG" + "IVT_DERIVATIVE_0" > "HOLLight.hollight.IVT_DERIVATIVE_0" + "IVT2" > "HOLLight.hollight.IVT2" + "IVT" > "HOLLight.hollight.IVT" + "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" + "ITERATE_UNION_GEN" > "HOLLight.hollight.ITERATE_UNION_GEN" + "ITERATE_UNION" > "HOLLight.hollight.ITERATE_UNION" + "ITERATE_SUPPORT" > "HOLLight.hollight.ITERATE_SUPPORT" + "ITERATE_SING" > "HOLLight.hollight.ITERATE_SING" + "ITERATE_RELATED" > "HOLLight.hollight.ITERATE_RELATED" + "ITERATE_IMAGE" > "HOLLight.hollight.ITERATE_IMAGE" + "ITERATE_EQ_NEUTRAL" > "HOLLight.hollight.ITERATE_EQ_NEUTRAL" + "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_CLOSED" > "HOLLight.hollight.ITERATE_CLOSED" + "ITERATE_CLAUSES_GEN" > "HOLLight.hollight.ITERATE_CLAUSES_GEN" + "ITERATE_CLAUSES" > "HOLLight.hollight.ITERATE_CLAUSES" + "ISO_def" > "HOLLight.hollight.ISO_def" + "ISO_USAGE" > "HOLLight.hollight.ISO_USAGE" + "ISO_REFL" > "HOLLight.hollight.ISO_REFL" + "ISO_FUN" > "HOLLight.hollight.ISO_FUN" + "ISMET_R1" > "HOLLight.hollight.ISMET_R1" + "IN_def" > "HOLLight.hollight.IN_def" + "IN_UNIV" > "HOLLight.hollight.IN_UNIV" + "IN_UNIONS" > "HOLLight.hollight.IN_UNIONS" + "IN_UNION" > "HOLLight.hollight.IN_UNION" + "IN_SUPPORT" > "HOLLight.hollight.IN_SUPPORT" + "IN_SING" > "HOLLight.hollight.IN_SING" + "IN_SET_OF_LIST" > "HOLLight.hollight.IN_SET_OF_LIST" + "IN_REST" > "HOLLight.hollight.IN_REST" + "IN_NUMSEG" > "HOLLight.hollight.IN_NUMSEG" + "IN_INTERS" > "HOLLight.hollight.IN_INTERS" + "IN_INTER" > "HOLLight.hollight.IN_INTER" + "IN_INSERT" > "HOLLight.hollight.IN_INSERT" + "IN_IMAGE" > "HOLLight.hollight.IN_IMAGE" + "IN_ELIM_THM" > "HOLLight.hollight.IN_ELIM_THM" + "IN_DISJOINT" > "HOLLight.hollight.IN_DISJOINT" + "IN_DIFF" > "HOLLight.hollight.IN_DIFF" + "IN_DELETE_EQ" > "HOLLight.hollight.IN_DELETE_EQ" + "IN_DELETE" > "HOLLight.hollight.IN_DELETE" + "INT_POW" > "HOLLight.hollight.INT_POW" + "INT_LT_DISCRETE" > "HOLLight.hollight.INT_LT_DISCRETE" + "INT_LT" > "HOLLight.hollight.INT_LT" + "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_FORALL_POS" > "HOLLight.hollight.INT_FORALL_POS" + "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_SUBSET" > "HOLLight.hollight.INTER_SUBSET" + "INTER_OVER_UNION" > "HOLLight.hollight.INTER_OVER_UNION" + "INTER_IDEMPOT" > "HOLLight.hollight.INTER_IDEMPOT" + "INTER_EMPTY" > "HOLLight.hollight.INTER_EMPTY" + "INTER_COMM" > "HOLLight.hollight.INTER_COMM" + "INTER_ASSOC" > "HOLLight.hollight.INTER_ASSOC" + "INTER_ACI" > "HOLLight.hollight.INTER_ACI" + "INTERVAL_LEMMA_LT" > "HOLLight.hollight.INTERVAL_LEMMA_LT" + "INTERVAL_LEMMA" > "HOLLight.hollight.INTERVAL_LEMMA" + "INTERVAL_ABS" > "HOLLight.hollight.INTERVAL_ABS" + "INTERS_def" > "HOLLight.hollight.INTERS_def" + "INTEGRAL_NULL" > "HOLLight.hollight.INTEGRAL_NULL" + "INSERT_def" > "HOLLight.hollight.INSERT_def" + "INSERT_UNIV" > "HOLLight.hollight.INSERT_UNIV" + "INSERT_UNION_EQ" > "HOLLight.hollight.INSERT_UNION_EQ" + "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_AC" > "HOLLight.hollight.INSERT_AC" + "INSERT" > "HOLLight.hollight.INSERT" + "INR_def" > "HOLLight.hollight.INR_def" + "INL_def" > "HOLLight.hollight.INL_def" + "INJ_def" > "HOLLight.hollight.INJ_def" + "INJ_INVERSE2" > "HOLLight.hollight.INJ_INVERSE2" + "INJP_def" > "HOLLight.hollight.INJP_def" + "INJP_INJ" > "HOLLight.hollight.INJP_INJ" + "INJN_def" > "HOLLight.hollight.INJN_def" + "INJN_INJ" > "HOLLight.hollight.INJN_INJ" + "INJF_def" > "HOLLight.hollight.INJF_def" + "INJF_INJ" > "HOLLight.hollight.INJF_INJ" + "INJECTIVE_ON_LEFT_INVERSE" > "HOLLight.hollight.INJECTIVE_ON_LEFT_INVERSE" + "INJECTIVE_LEFT_INVERSE" > "HOLLight.hollight.INJECTIVE_LEFT_INVERSE" + "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_IMAGE_INJ" > "HOLLight.hollight.INFINITE_IMAGE_INJ" + "INFINITE_DIFF_FINITE" > "HOLLight.hollight.INFINITE_DIFF_FINITE" + "IMP_EQ_CLAUSE" > "HOLLight.hollight.IMP_EQ_CLAUSE" + "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_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_DIFF_INJ" > "HOLLight.hollight.IMAGE_DIFF_INJ" + "IMAGE_DELETE_INJ" > "HOLLight.hollight.IMAGE_DELETE_INJ" + "IMAGE_CONST" > "HOLLight.hollight.IMAGE_CONST" + "IMAGE_CLAUSES" > "HOLLight.hollight.IMAGE_CLAUSES" + "HREAL_MUL_RZERO" > "HOLLight.hollight.HREAL_MUL_RZERO" + "HREAL_MUL_LZERO" > "HOLLight.hollight.HREAL_MUL_LZERO" + "HREAL_LE_MUL_RCANCEL_IMP" > "HOLLight.hollight.HREAL_LE_MUL_RCANCEL_IMP" + "HREAL_LE_EXISTS_DEF" > "HOLLight.hollight.HREAL_LE_EXISTS_DEF" + "HREAL_LE_ADD_RCANCEL" > "HOLLight.hollight.HREAL_LE_ADD_RCANCEL" + "HREAL_LE_ADD_LCANCEL" > "HOLLight.hollight.HREAL_LE_ADD_LCANCEL" + "HREAL_LE_ADD2" > "HOLLight.hollight.HREAL_LE_ADD2" + "HREAL_EQ_ADD_RCANCEL" > "HOLLight.hollight.HREAL_EQ_ADD_RCANCEL" + "HREAL_EQ_ADD_LCANCEL" > "HOLLight.hollight.HREAL_EQ_ADD_LCANCEL" + "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" + "HAS_SIZE_def" > "HOLLight.hollight.HAS_SIZE_def" + "HAS_SIZE_UNIONS" > "HOLLight.hollight.HAS_SIZE_UNIONS" + "HAS_SIZE_UNION" > "HOLLight.hollight.HAS_SIZE_UNION" + "HAS_SIZE_SUC" > "HOLLight.hollight.HAS_SIZE_SUC" + "HAS_SIZE_PRODUCT_DEPENDENT" > "HOLLight.hollight.HAS_SIZE_PRODUCT_DEPENDENT" + "HAS_SIZE_PRODUCT" > "HOLLight.hollight.HAS_SIZE_PRODUCT" + "HAS_SIZE_POWERSET" > "HOLLight.hollight.HAS_SIZE_POWERSET" + "HAS_SIZE_NUMSEG_LT" > "HOLLight.hollight.HAS_SIZE_NUMSEG_LT" + "HAS_SIZE_NUMSEG_LE" > "HOLLight.hollight.HAS_SIZE_NUMSEG_LE" + "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" > "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_CLAUSES" > "HOLLight.hollight.HAS_SIZE_CLAUSES" + "HAS_SIZE_CARD" > "HOLLight.hollight.HAS_SIZE_CARD" + "HAS_SIZE_0" > "HOLLight.hollight.HAS_SIZE_0" + "GSPEC_def" > "HOLLight.hollight.GSPEC_def" + "GP_FINITE" > "HOLLight.hollight.GP_FINITE" + "GP" > "HOLLight.hollight.GP" + "GEQ_def" > "HOLLight.hollight.GEQ_def" + "GAUGE_MIN" > "HOLLight.hollight.GAUGE_MIN" + "GABS_def" > "HOLLight.hollight.GABS_def" + "FUN_EQ_THM" > "Fun.expand_fun_eq" + "FUNCTION_FACTORS_RIGHT" > "HOLLight.hollight.FUNCTION_FACTORS_RIGHT" + "FUNCTION_FACTORS_LEFT" > "HOLLight.hollight.FUNCTION_FACTORS_LEFT" + "FTC1" > "HOLLight.hollight.FTC1" + "FSTCART_PASTECART" > "HOLLight.hollight.FSTCART_PASTECART" + "FST" > "Product_Type.fst_conv" + "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_IN_UNIONS" > "HOLLight.hollight.FORALL_IN_UNIONS" + "FORALL_IN_IMAGE" > "HOLLight.hollight.FORALL_IN_IMAGE" + "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" + "FNIL_def" > "HOLLight.hollight.FNIL_def" + "FINREC_def" > "HOLLight.hollight.FINREC_def" + "FINREC_UNIQUE_LEMMA" > "HOLLight.hollight.FINREC_UNIQUE_LEMMA" + "FINREC_SUC_LEMMA" > "HOLLight.hollight.FINREC_SUC_LEMMA" + "FINREC_FUN_LEMMA" > "HOLLight.hollight.FINREC_FUN_LEMMA" + "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_UNIONS" > "HOLLight.hollight.FINITE_UNIONS" + "FINITE_UNION" > "HOLLight.hollight.FINITE_UNION" + "FINITE_SUPPORT_DELTA" > "HOLLight.hollight.FINITE_SUPPORT_DELTA" + "FINITE_SUPPORT" > "HOLLight.hollight.FINITE_SUPPORT" + "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_RESTRICT" > "HOLLight.hollight.FINITE_RESTRICT" + "FINITE_RECURSION_DELETE" > "HOLLight.hollight.FINITE_RECURSION_DELETE" + "FINITE_RECURSION" > "HOLLight.hollight.FINITE_RECURSION" + "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_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_INJ" > "HOLLight.hollight.FINITE_INDEX_INJ" + "FINITE_IMAGE_INJ_GENERAL" > "HOLLight.hollight.FINITE_IMAGE_INJ_GENERAL" + "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_FUNSPACE" > "HOLLight.hollight.FINITE_FUNSPACE" + "FINITE_FINITE_IMAGE" > "HOLLight.hollight.FINITE_FINITE_IMAGE" + "FINITE_DIFF" > "HOLLight.hollight.FINITE_DIFF" + "FINITE_DELETE_IMP" > "HOLLight.hollight.FINITE_DELETE_IMP" + "FINITE_DELETE" > "HOLLight.hollight.FINITE_DELETE" + "FINE_MIN" > "HOLLight.hollight.FINE_MIN" + "FILTER_def" > "HOLLight.hollight.FILTER_def" + "FILTER_MAP" > "HOLLight.hollight.FILTER_MAP" + "FILTER_APPEND" > "HOLLight.hollight.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" + "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" + "EXISTS_UNIQUE_REFL" > "HOL.ex1_eq_1" + "EXISTS_UNIQUE_ALT" > "HOLLight.hollight.EXISTS_UNIQUE_ALT" + "EXISTS_UNIQUE" > "HOL.Ex1_def" + "EXISTS_THM" > "HOL4Setup.EXISTS_DEF" + "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_OR_THM" > "HOL.ex_disj_distrib" + "EXISTS_ONE_REP" > "HOLLight.hollight.EXISTS_ONE_REP" + "EXISTS_NOT_THM" > "Inductive.basic_monos_15" + "EXISTS_IN_IMAGE" > "HOLLight.hollight.EXISTS_IN_IMAGE" + "EXISTS_IN_CLAUSES" > "HOLLight.hollight.EXISTS_IN_CLAUSES" + "EXISTS_EX" > "HOLLight.hollight.EXISTS_EX" + "EXISTS_BOOL_THM" > "Set.ex_bool_eq" + "EXCLUDED_MIDDLE" > "HOLLight.hollight.EXCLUDED_MIDDLE" + "EVEN_def" > "HOLLight.hollight.EVEN_def" + "EVEN_OR_ODD" > "HOLLight.hollight.EVEN_OR_ODD" + "EVEN_ODD_DECOMPOSITION" > "HOLLight.hollight.EVEN_ODD_DECOMPOSITION" + "EVEN_MULT" > "HOLLight.hollight.EVEN_MULT" + "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_DOUBLE" > "HOLLight.hollight.EVEN_DOUBLE" + "EVEN_DIV2" > "HOLLight.hollight.EVEN_DIV2" + "EVEN_AND_ODD" > "HOLLight.hollight.EVEN_AND_ODD" + "EVEN_ADD" > "HOLLight.hollight.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_MULT_RCANCEL" > "Nat.mult_cancel2" + "EQ_MULT_LCANCEL" > "NatBin.nat_mult_eq_cancel_disj" + "EQ_IMP_LE" > "HOLLight.hollight.EQ_IMP_LE" + "EQ_EXT" > "HOL.meta_eq_to_obj_eq" + "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_LCANCEL_0" > "HOLLight.hollight.EQ_ADD_LCANCEL_0" + "EQ_ADD_LCANCEL" > "Nat.nat_add_left_cancel" + "EMPTY_def" > "HOLLight.hollight.EMPTY_def" + "EMPTY_UNIONS" > "HOLLight.hollight.EMPTY_UNIONS" + "EMPTY_UNION" > "HOLLight.hollight.EMPTY_UNION" + "EMPTY_SUBSET" > "HOLLight.hollight.EMPTY_SUBSET" + "EMPTY_NOT_UNIV" > "HOLLight.hollight.EMPTY_NOT_UNIV" + "EMPTY_GSPEC" > "HOLLight.hollight.EMPTY_GSPEC" + "EMPTY_DIFF" > "HOLLight.hollight.EMPTY_DIFF" + "EMPTY_DELETE" > "HOLLight.hollight.EMPTY_DELETE" + "EL_def" > "HOLLight.hollight.EL_def" + "DORDER_TENDSTO" > "HOLLight.hollight.DORDER_TENDSTO" + "DORDER_NGE" > "HOLLight.hollight.DORDER_NGE" + "DORDER_LEMMA" > "HOLLight.hollight.DORDER_LEMMA" + "DIV_def" > "HOLLight.hollight.DIV_def" + "DIV_UNIQ" > "HOLLight.hollight.DIV_UNIQ" + "DIV_REFL" > "HOLLight.hollight.DIV_REFL" + "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_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" + "DIVMOD_EXIST" > "HOLLight.hollight.DIVMOD_EXIST" + "DIVMOD_ELIM_THM" > "HOLLight.hollight.DIVMOD_ELIM_THM" + "DIVISION_UBOUND_LT" > "HOLLight.hollight.DIVISION_UBOUND_LT" + "DIVISION_UBOUND" > "HOLLight.hollight.DIVISION_UBOUND" + "DIVISION_THM" > "HOLLight.hollight.DIVISION_THM" + "DIVISION_SINGLE" > "HOLLight.hollight.DIVISION_SINGLE" + "DIVISION_RHS" > "HOLLight.hollight.DIVISION_RHS" + "DIVISION_LT_GEN" > "HOLLight.hollight.DIVISION_LT_GEN" + "DIVISION_LT" > "HOLLight.hollight.DIVISION_LT" + "DIVISION_LHS" > "HOLLight.hollight.DIVISION_LHS" + "DIVISION_LE" > "HOLLight.hollight.DIVISION_LE" + "DIVISION_LBOUND_LT" > "HOLLight.hollight.DIVISION_LBOUND_LT" + "DIVISION_LBOUND" > "HOLLight.hollight.DIVISION_LBOUND" + "DIVISION_GT" > "HOLLight.hollight.DIVISION_GT" + "DIVISION_EXISTS" > "HOLLight.hollight.DIVISION_EXISTS" + "DIVISION_EQ" > "HOLLight.hollight.DIVISION_EQ" + "DIVISION_APPEND_LEMMA2" > "HOLLight.hollight.DIVISION_APPEND_LEMMA2" + "DIVISION_APPEND_LEMMA1" > "HOLLight.hollight.DIVISION_APPEND_LEMMA1" + "DIVISION_APPEND" > "HOLLight.hollight.DIVISION_APPEND" + "DIVISION_1" > "HOLLight.hollight.DIVISION_1" + "DIVISION_0" > "HOLLight.hollight.DIVISION_0" + "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" + "DIST_REFL" > "HOLLight.hollight.DIST_REFL" + "DIST_RADD_0" > "HOLLight.hollight.DIST_RADD_0" + "DIST_RADD" > "HOLLight.hollight.DIST_RADD" + "DIST_LZERO" > "HOLLight.hollight.DIST_LZERO" + "DIST_LMUL" > "HOLLight.hollight.DIST_LMUL" + "DIST_LE_CASES" > "HOLLight.hollight.DIST_LE_CASES" + "DIST_LADD_0" > "HOLLight.hollight.DIST_LADD_0" + "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_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" + "DISJOINT_INSERT" > "HOLLight.hollight.DISJOINT_INSERT" + "DISJOINT_EMPTY_REFL" > "HOLLight.hollight.DISJOINT_EMPTY_REFL" + "DISJOINT_EMPTY" > "HOLLight.hollight.DISJOINT_EMPTY" + "DISJOINT_DELETE_SYM" > "HOLLight.hollight.DISJOINT_DELETE_SYM" + "DINT_UNIQ" > "HOLLight.hollight.DINT_UNIQ" + "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_XM1" > "HOLLight.hollight.DIFF_XM1" + "DIFF_X" > "HOLLight.hollight.DIFF_X" + "DIFF_UNIV" > "HOLLight.hollight.DIFF_UNIV" + "DIFF_UNIQ" > "HOLLight.hollight.DIFF_UNIQ" + "DIFF_TAN_COMPOSITE" > "HOLLight.hollight.DIFF_TAN_COMPOSITE" + "DIFF_TAN" > "HOLLight.hollight.DIFF_TAN" + "DIFF_SUM" > "HOLLight.hollight.DIFF_SUM" + "DIFF_SUB" > "HOLLight.hollight.DIFF_SUB" + "DIFF_SIN" > "HOLLight.hollight.DIFF_SIN" + "DIFF_POW" > "HOLLight.hollight.DIFF_POW" + "DIFF_NEG" > "HOLLight.hollight.DIFF_NEG" + "DIFF_MUL" > "HOLLight.hollight.DIFF_MUL" + "DIFF_LN_COMPOSITE" > "HOLLight.hollight.DIFF_LN_COMPOSITE" + "DIFF_LN" > "HOLLight.hollight.DIFF_LN" + "DIFF_LMIN" > "HOLLight.hollight.DIFF_LMIN" + "DIFF_LMAX" > "HOLLight.hollight.DIFF_LMAX" + "DIFF_LINC" > "HOLLight.hollight.DIFF_LINC" + "DIFF_LDEC" > "HOLLight.hollight.DIFF_LDEC" + "DIFF_LCONST" > "HOLLight.hollight.DIFF_LCONST" + "DIFF_ISCONST_END_SIMPLE" > "HOLLight.hollight.DIFF_ISCONST_END_SIMPLE" + "DIFF_ISCONST_END" > "HOLLight.hollight.DIFF_ISCONST_END" + "DIFF_ISCONST_ALL" > "HOLLight.hollight.DIFF_ISCONST_ALL" + "DIFF_ISCONST" > "HOLLight.hollight.DIFF_ISCONST" + "DIFF_INVERSE_LT" > "HOLLight.hollight.DIFF_INVERSE_LT" + "DIFF_INVERSE" > "HOLLight.hollight.DIFF_INVERSE" + "DIFF_INV" > "HOLLight.hollight.DIFF_INV" + "DIFF_INSERT" > "HOLLight.hollight.DIFF_INSERT" + "DIFF_EXP" > "HOLLight.hollight.DIFF_EXP" + "DIFF_EQ_EMPTY" > "HOLLight.hollight.DIFF_EQ_EMPTY" + "DIFF_EMPTY" > "HOLLight.hollight.DIFF_EMPTY" + "DIFF_DIV" > "HOLLight.hollight.DIFF_DIV" + "DIFF_DIFF" > "HOLLight.hollight.DIFF_DIFF" + "DIFF_COS" > "HOLLight.hollight.DIFF_COS" + "DIFF_CONT" > "HOLLight.hollight.DIFF_CONT" + "DIFF_CONST" > "HOLLight.hollight.DIFF_CONST" + "DIFF_COMPOSITE" > "HOLLight.hollight.DIFF_COMPOSITE" + "DIFF_CMUL" > "HOLLight.hollight.DIFF_CMUL" + "DIFF_CHAIN" > "HOLLight.hollight.DIFF_CHAIN" + "DIFF_CARAT" > "HOLLight.hollight.DIFF_CARAT" + "DIFF_ATN_COMPOSITE" > "HOLLight.hollight.DIFF_ATN_COMPOSITE" + "DIFF_ATN" > "HOLLight.hollight.DIFF_ATN" + "DIFF_ASN_COS" > "HOLLight.hollight.DIFF_ASN_COS" + "DIFF_ASN" > "HOLLight.hollight.DIFF_ASN" + "DIFF_ADD" > "HOLLight.hollight.DIFF_ADD" + "DIFF_ACS_SIN" > "HOLLight.hollight.DIFF_ACS_SIN" + "DIFF_ACS" > "HOLLight.hollight.DIFF_ACS" + "DIFFS_NEG" > "HOLLight.hollight.DIFFS_NEG" + "DIFFS_LEMMA2" > "HOLLight.hollight.DIFFS_LEMMA2" + "DIFFS_LEMMA" > "HOLLight.hollight.DIFFS_LEMMA" + "DIFFS_EQUIV" > "HOLLight.hollight.DIFFS_EQUIV" + "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_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" + "DEF_treal_le" > "HOLLight.hollight.DEF_treal_le" + "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_tendsto" > "HOLLight.hollight.DEF_tendsto" + "DEF_tends_real_real" > "HOLLight.hollight.DEF_tends_real_real" + "DEF_tends_num_real" > "HOLLight.hollight.DEF_tends_num_real" + "DEF_tends" > "HOLLight.hollight.DEF_tends" + "DEF_tdiv" > "HOLLight.hollight.DEF_tdiv" + "DEF_tan" > "HOLLight.hollight.DEF_tan" + "DEF_tailadmissible" > "HOLLight.hollight.DEF_tailadmissible" + "DEF_support" > "HOLLight.hollight.DEF_support" + "DEF_superadmissible" > "HOLLight.hollight.DEF_superadmissible" + "DEF_sup" > "HOLLight.hollight.DEF_sup" + "DEF_sums" > "HOLLight.hollight.DEF_sums" + "DEF_summable" > "HOLLight.hollight.DEF_summable" + "DEF_suminf" > "HOLLight.hollight.DEF_suminf" + "DEF_sum" > "HOLLight.hollight.DEF_sum" + "DEF_subseq" > "HOLLight.hollight.DEF_subseq" + "DEF_sqrt" > "HOLLight.hollight.DEF_sqrt" + "DEF_sndcart" > "HOLLight.hollight.DEF_sndcart" + "DEF_sin" > "HOLLight.hollight.DEF_sin" + "DEF_set_of_list" > "HOLLight.hollight.DEF_set_of_list" + "DEF_rsum" > "HOLLight.hollight.DEF_rsum" + "DEF_root" > "HOLLight.hollight.DEF_root" + "DEF_real_sub" > "HOLLight.hollight.DEF_real_sub" + "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_min" > "HOLLight.hollight.DEF_real_min" + "DEF_real_max" > "HOLLight.hollight.DEF_real_max" + "DEF_real_lt" > "HOLLight.hollight.DEF_real_lt" + "DEF_real_le" > "HOLLight.hollight.DEF_real_le" + "DEF_real_inv" > "HOLLight.hollight.DEF_real_inv" + "DEF_real_gt" > "HOLLight.hollight.DEF_real_gt" + "DEF_real_ge" > "HOLLight.hollight.DEF_real_ge" + "DEF_real_div" > "HOLLight.hollight.DEF_real_div" + "DEF_real_add" > "HOLLight.hollight.DEF_real_add" + "DEF_real_abs" > "HOLLight.hollight.DEF_real_abs" + "DEF_re_universe" > "HOLLight.hollight.DEF_re_universe" + "DEF_re_union" > "HOLLight.hollight.DEF_re_union" + "DEF_re_subset" > "HOLLight.hollight.DEF_re_subset" + "DEF_re_null" > "HOLLight.hollight.DEF_re_null" + "DEF_re_intersect" > "HOLLight.hollight.DEF_re_intersect" + "DEF_re_compl" > "HOLLight.hollight.DEF_re_compl" + "DEF_re_Union" > "HOLLight.hollight.DEF_re_Union" + "DEF_psum" > "HOLLight.hollight.DEF_psum" + "DEF_pi" > "HOLLight.hollight.DEF_pi" + "DEF_pastecart" > "HOLLight.hollight.DEF_pastecart" + "DEF_pairwise" > "HOLLight.hollight.DEF_pairwise" + "DEF_o" > "Fun.o_apply" + "DEF_nsum" > "HOLLight.hollight.DEF_nsum" + "DEF_neutral" > "HOLLight.hollight.DEF_neutral" + "DEF_neigh" > "HOLLight.hollight.DEF_neigh" + "DEF_nadd_rinv" > "HOLLight.hollight.DEF_nadd_rinv" + "DEF_nadd_of_num" > "HOLLight.hollight.DEF_nadd_of_num" + "DEF_nadd_mul" > "HOLLight.hollight.DEF_nadd_mul" + "DEF_nadd_le" > "HOLLight.hollight.DEF_nadd_le" + "DEF_nadd_inv" > "HOLLight.hollight.DEF_nadd_inv" + "DEF_nadd_eq" > "HOLLight.hollight.DEF_nadd_eq" + "DEF_nadd_add" > "HOLLight.hollight.DEF_nadd_add" + "DEF_mtop" > "HOLLight.hollight.DEF_mtop" + "DEF_mr1" > "HOLLight.hollight.DEF_mr1" + "DEF_monoidal" > "HOLLight.hollight.DEF_monoidal" + "DEF_mono" > "HOLLight.hollight.DEF_mono" + "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_ln" > "HOLLight.hollight.DEF_ln" + "DEF_list_of_set" > "HOLLight.hollight.DEF_list_of_set" + "DEF_limpt" > "HOLLight.hollight.DEF_limpt" + "DEF_lim" > "HOLLight.hollight.DEF_lim" + "DEF_lambda" > "HOLLight.hollight.DEF_lambda" + "DEF_iterate" > "HOLLight.hollight.DEF_iterate" + "DEF_istopology" > "HOLLight.hollight.DEF_istopology" + "DEF_ismet" > "HOLLight.hollight.DEF_ismet" + "DEF_is_nadd" > "HOLLight.hollight.DEF_is_nadd" + "DEF_is_int" > "HOLLight.hollight.DEF_is_int" + "DEF_int_sub" > "HOLLight.hollight.DEF_int_sub" + "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_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_add" > "HOLLight.hollight.DEF_int_add" + "DEF_int_abs" > "HOLLight.hollight.DEF_int_abs" + "DEF_hreal_of_num" > "HOLLight.hollight.DEF_hreal_of_num" + "DEF_hreal_mul" > "HOLLight.hollight.DEF_hreal_mul" + "DEF_hreal_le" > "HOLLight.hollight.DEF_hreal_le" + "DEF_hreal_inv" > "HOLLight.hollight.DEF_hreal_inv" + "DEF_hreal_add" > "HOLLight.hollight.DEF_hreal_add" + "DEF_gauge" > "HOLLight.hollight.DEF_gauge" + "DEF_fstcart" > "HOLLight.hollight.DEF_fstcart" + "DEF_finite_index" > "HOLLight.hollight.DEF_finite_index" + "DEF_fine" > "HOLLight.hollight.DEF_fine" + "DEF_exp" > "HOLLight.hollight.DEF_exp" + "DEF_dsize" > "HOLLight.hollight.DEF_dsize" + "DEF_dorder" > "HOLLight.hollight.DEF_dorder" + "DEF_division" > "HOLLight.hollight.DEF_division" + "DEF_dist" > "HOLLight.hollight.DEF_dist" + "DEF_dimindex" > "HOLLight.hollight.DEF_dimindex" + "DEF_diffs" > "HOLLight.hollight.DEF_diffs" + "DEF_diffl" > "HOLLight.hollight.DEF_diffl" + "DEF_differentiable" > "HOLLight.hollight.DEF_differentiable" + "DEF_defint" > "HOLLight.hollight.DEF_defint" + "DEF_cos" > "HOLLight.hollight.DEF_cos" + "DEF_convergent" > "HOLLight.hollight.DEF_convergent" + "DEF_contl" > "HOLLight.hollight.DEF_contl" + "DEF_closed" > "HOLLight.hollight.DEF_closed" + "DEF_cauchy" > "HOLLight.hollight.DEF_cauchy" + "DEF_bounded" > "HOLLight.hollight.DEF_bounded" + "DEF_ball" > "HOLLight.hollight.DEF_ball" + "DEF_atn" > "HOLLight.hollight.DEF_atn" + "DEF_asn" > "HOLLight.hollight.DEF_asn" + "DEF_admissible" > "HOLLight.hollight.DEF_admissible" + "DEF_acs" > "HOLLight.hollight.DEF_acs" + "DEF__star_" > "HOLLightCompat.mult_altdef" + "DEF__slash__backslash_" > "HOLLightCompat.light_and_def" + "DEF__questionmark__exclamationmark_" > "HOL4Compat.EXISTS_UNIQUE_DEF" + "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__exclamationmark_" > "HOL.All_def" + "DEF__equal__equal__greaterthan_" > "HOLLightCompat.light_imp_def" + "DEF__equal__equal_" > "HOLLight.hollight.DEF__equal__equal_" + "DEF__dot__dot_" > "HOLLight.hollight.DEF__dot__dot_" + "DEF__backslash__slash_" > "HOL.or_def" + "DEF__FALSITY_" > "HOLLight.hollight.DEF__FALSITY_" + "DEF__10314" > "HOLLight.hollight.DEF__10314" + "DEF__10313" > "HOLLight.hollight.DEF__10313" + "DEF__10312" > "HOLLight.hollight.DEF__10312" + "DEF__10289" > "HOLLight.hollight.DEF__10289" + "DEF__10288" > "HOLLight.hollight.DEF__10288" + "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_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_SING" > "HOLLight.hollight.DEF_SING" + "DEF_SETSPEC" > "HOLLight.hollight.DEF_SETSPEC" + "DEF_REVERSE" > "HOLLight.hollight.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_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_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_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_ITSET" > "HOLLight.hollight.DEF_ITSET" + "DEF_ITLIST2" > "HOLLight.hollight.DEF_ITLIST2" + "DEF_ITLIST" > "HOLLight.hollight.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_INR" > "HOLLight.hollight.DEF_INR" + "DEF_INL" > "HOLLight.hollight.DEF_INL" + "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_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_FNIL" > "HOLLight.hollight.DEF_FNIL" + "DEF_FINREC" > "HOLLight.hollight.DEF_FINREC" + "DEF_FINITE" > "HOLLight.hollight.DEF_FINITE" + "DEF_FILTER" > "HOLLight.hollight.DEF_FILTER" + "DEF_FCONS" > "HOLLight.hollight.DEF_FCONS" + "DEF_FACT" > "HOLLight.hollight.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_DECIMAL" > "HOLLight.hollight.DEF_DECIMAL" + "DEF_CURRY" > "HOLLight.hollight.DEF_CURRY" + "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_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_BOTTOM" > "HOLLight.hollight.DEF_BOTTOM" + "DEF_BIT1" > "HOLLightCompat.NUMERAL_BIT1_altdef" + "DEF_BIT0" > "HOLLightCompat.NUMERAL_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_$" > "HOLLight.hollight.DEF_$" + "DECOMPOSITION" > "HOLLight.hollight.DECOMPOSITION" + "DECIMAL_def" > "HOLLight.hollight.DECIMAL_def" + "CURRY_def" > "HOLLight.hollight.CURRY_def" + "COUNTABLE_def" > "HOLLight.hollight.COUNTABLE_def" + "COS_ZERO_LEMMA" > "HOLLight.hollight.COS_ZERO_LEMMA" + "COS_ZERO" > "HOLLight.hollight.COS_ZERO" + "COS_TOTAL" > "HOLLight.hollight.COS_TOTAL" + "COS_SIN_SQRT" > "HOLLight.hollight.COS_SIN_SQRT" + "COS_SIN" > "HOLLight.hollight.COS_SIN" + "COS_POS_PI2" > "HOLLight.hollight.COS_POS_PI2" + "COS_POS_PI" > "HOLLight.hollight.COS_POS_PI" + "COS_PI2" > "HOLLight.hollight.COS_PI2" + "COS_PI" > "HOLLight.hollight.COS_PI" + "COS_PERIODIC_PI" > "HOLLight.hollight.COS_PERIODIC_PI" + "COS_PERIODIC" > "HOLLight.hollight.COS_PERIODIC" + "COS_PAIRED" > "HOLLight.hollight.COS_PAIRED" + "COS_ONE_2PI" > "HOLLight.hollight.COS_ONE_2PI" + "COS_NPI" > "HOLLight.hollight.COS_NPI" + "COS_NEG" > "HOLLight.hollight.COS_NEG" + "COS_ISZERO" > "HOLLight.hollight.COS_ISZERO" + "COS_FDIFF" > "HOLLight.hollight.COS_FDIFF" + "COS_DOUBLE" > "HOLLight.hollight.COS_DOUBLE" + "COS_CONVERGES" > "HOLLight.hollight.COS_CONVERGES" + "COS_BOUNDS" > "HOLLight.hollight.COS_BOUNDS" + "COS_BOUND" > "HOLLight.hollight.COS_BOUND" + "COS_ATN_NZ" > "HOLLight.hollight.COS_ATN_NZ" + "COS_ASN_NZ" > "HOLLight.hollight.COS_ASN_NZ" + "COS_ADD" > "HOLLight.hollight.COS_ADD" + "COS_ACS" > "HOLLight.hollight.COS_ACS" + "COS_ABS" > "HOLLight.hollight.COS_ABS" + "COS_2" > "HOLLight.hollight.COS_2" + "COS_0" > "HOLLight.hollight.COS_0" + "CONT_X" > "HOLLight.hollight.CONT_X" + "CONT_SUB" > "HOLLight.hollight.CONT_SUB" + "CONT_NEG" > "HOLLight.hollight.CONT_NEG" + "CONT_MUL" > "HOLLight.hollight.CONT_MUL" + "CONT_INVERSE" > "HOLLight.hollight.CONT_INVERSE" + "CONT_INV" > "HOLLight.hollight.CONT_INV" + "CONT_INJ_RANGE" > "HOLLight.hollight.CONT_INJ_RANGE" + "CONT_INJ_LEMMA2" > "HOLLight.hollight.CONT_INJ_LEMMA2" + "CONT_INJ_LEMMA" > "HOLLight.hollight.CONT_INJ_LEMMA" + "CONT_HASSUP" > "HOLLight.hollight.CONT_HASSUP" + "CONT_DIV" > "HOLLight.hollight.CONT_DIV" + "CONT_CONST" > "HOLLight.hollight.CONT_CONST" + "CONT_COMPOSE" > "HOLLight.hollight.CONT_COMPOSE" + "CONT_BOUNDED_ABS" > "HOLLight.hollight.CONT_BOUNDED_ABS" + "CONT_BOUNDED" > "HOLLight.hollight.CONT_BOUNDED" + "CONT_ATTAINS_ALL" > "HOLLight.hollight.CONT_ATTAINS_ALL" + "CONT_ATTAINS2" > "HOLLight.hollight.CONT_ATTAINS2" + "CONT_ATTAINS" > "HOLLight.hollight.CONT_ATTAINS" + "CONT_ADD" > "HOLLight.hollight.CONT_ADD" + "CONTL_LIM" > "HOLLight.hollight.CONTL_LIM" + "CONS_def" > "HOLLight.hollight.CONS_def" + "CONS_11" > "HOLLight.hollight.CONS_11" + "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_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_EXPAND" > "HOLLight.hollight.COND_EXPAND" + "COND_EQ_CLAUSE" > "HOLLight.hollight.COND_EQ_CLAUSE" + "COND_ELIM_THM" > "HOLLight.hollight.COND_ELIM_THM" + "COND_CLAUSES" > "HOLLight.hollight.COND_CLAUSES" + "COND_ABS" > "HOLLight.hollight.COND_ABS" + "COMPONENT" > "HOLLight.hollight.COMPONENT" + "COMPL_MEM" > "HOLLight.hollight.COMPL_MEM" + "CLOSED_LIMPT" > "HOLLight.hollight.CLOSED_LIMPT" + "CIRCLE_SINCOS" > "HOLLight.hollight.CIRCLE_SINCOS" + "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" > "HOLLight.hollight.CART_EQ" + "CARD_def" > "HOLLight.hollight.CARD_def" + "CARD_UNION_LE" > "HOLLight.hollight.CARD_UNION_LE" + "CARD_UNION_EQ" > "HOLLight.hollight.CARD_UNION_EQ" + "CARD_UNIONS_LE" > "HOLLight.hollight.CARD_UNIONS_LE" + "CARD_UNION" > "HOLLight.hollight.CARD_UNION" + "CARD_SUBSET_LE" > "HOLLight.hollight.CARD_SUBSET_LE" + "CARD_SUBSET_EQ" > "HOLLight.hollight.CARD_SUBSET_EQ" + "CARD_SUBSET" > "HOLLight.hollight.CARD_SUBSET" + "CARD_PSUBSET" > "HOLLight.hollight.CARD_PSUBSET" + "CARD_PRODUCT" > "HOLLight.hollight.CARD_PRODUCT" + "CARD_POWERSET" > "HOLLight.hollight.CARD_POWERSET" + "CARD_NUMSEG_LT" > "HOLLight.hollight.CARD_NUMSEG_LT" + "CARD_NUMSEG_LEMMA" > "HOLLight.hollight.CARD_NUMSEG_LEMMA" + "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" > "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_0" > "HOLLight.hollight.CARD_EQ_0" + "CARD_DELETE" > "HOLLight.hollight.CARD_DELETE" + "CARD_CLAUSES" > "HOLLight.hollight.CARD_CLAUSES" + "BOUNDS_NOTZERO" > "HOLLight.hollight.BOUNDS_NOTZERO" + "BOUNDS_LINEAR_0" > "HOLLight.hollight.BOUNDS_LINEAR_0" + "BOUNDS_LINEAR" > "HOLLight.hollight.BOUNDS_LINEAR" + "BOUNDS_IGNORE" > "HOLLight.hollight.BOUNDS_IGNORE" + "BOUNDS_DIVIDED" > "HOLLight.hollight.BOUNDS_DIVIDED" + "BOTTOM_def" > "HOLLight.hollight.BOTTOM_def" + "BOOL_CASES_AX" > "Datatype.bool.nchotomy" + "BOLZANO_LEMMA" > "HOLLight.hollight.BOLZANO_LEMMA" + "BIT1_THM" > "HOLLightCompat.NUMERAL_BIT1_altdef" + "BIT0_THM" > "HOLLightCompat.NUMERAL_BIT0_def" + "BIJ_def" > "HOLLight.hollight.BIJ_def" + "BETA_THM" > "Presburger.fm_modd_pinf" + "BALL_OPEN" > "HOLLight.hollight.BALL_OPEN" + "BALL_NEIGH" > "HOLLight.hollight.BALL_NEIGH" + "ATN_TAN" > "HOLLight.hollight.ATN_TAN" + "ATN_POS_LT" > "HOLLight.hollight.ATN_POS_LT" + "ATN_POS_LE" > "HOLLight.hollight.ATN_POS_LE" + "ATN_NEG" > "HOLLight.hollight.ATN_NEG" + "ATN_MONO_LT_EQ" > "HOLLight.hollight.ATN_MONO_LT_EQ" + "ATN_MONO_LT" > "HOLLight.hollight.ATN_MONO_LT" + "ATN_MONO_LE_EQ" > "HOLLight.hollight.ATN_MONO_LE_EQ" + "ATN_LT_PI4_POS" > "HOLLight.hollight.ATN_LT_PI4_POS" + "ATN_LT_PI4_NEG" > "HOLLight.hollight.ATN_LT_PI4_NEG" + "ATN_LT_PI4" > "HOLLight.hollight.ATN_LT_PI4" + "ATN_LE_PI4" > "HOLLight.hollight.ATN_LE_PI4" + "ATN_INJ" > "HOLLight.hollight.ATN_INJ" + "ATN_BOUNDS" > "HOLLight.hollight.ATN_BOUNDS" + "ATN_1" > "HOLLight.hollight.ATN_1" + "ATN_0" > "HOLLight.hollight.ATN_0" + "ATN" > "HOLLight.hollight.ATN" + "ASSOC_def" > "HOLLight.hollight.ASSOC_def" + "ASN_SIN" > "HOLLight.hollight.ASN_SIN" + "ASN_BOUNDS_LT" > "HOLLight.hollight.ASN_BOUNDS_LT" + "ASN_BOUNDS" > "HOLLight.hollight.ASN_BOUNDS" + "ASN" > "HOLLight.hollight.ASN" + "ARITH_ZERO" > "HOLLight.hollight.ARITH_ZERO" + "ARITH_SUC" > "HOLLight.hollight.ARITH_SUC" + "ARITH_SUB" > "HOLLight.hollight.ARITH_SUB" + "ARITH_PRE" > "HOLLight.hollight.ARITH_PRE" + "ARITH_ODD" > "HOLLight.hollight.ARITH_ODD" + "ARITH_MULT" > "HOLLight.hollight.ARITH_MULT" + "ARITH_LT" > "HOLLight.hollight.ARITH_LT" + "ARITH_LE" > "HOLLight.hollight.ARITH_LE" + "ARITH_EXP" > "HOLLight.hollight.ARITH_EXP" + "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" + "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" + "ADMISSIBLE_NEST" > "HOLLight.hollight.ADMISSIBLE_NEST" + "ADMISSIBLE_LAMBDA" > "HOLLight.hollight.ADMISSIBLE_LAMBDA" + "ADMISSIBLE_IMP_SUPERADMISSIBLE" > "HOLLight.hollight.ADMISSIBLE_IMP_SUPERADMISSIBLE" + "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_SUC" > "Nat.add_Suc_right" + "ADD_SUBR2" > "Nat.diff_add_0" + "ADD_SUBR" > "HOLLight.hollight.ADD_SUBR" + "ADD_SUB2" > "Nat.diff_add_inverse" + "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_AC" > "HOLLight.hollight.ADD_AC" + "ADD_0" > "Finite_Set.AC_add.f_e.ident" + "ADD1" > "HOLLight.hollight.ADD1" + "ACS_MONO_LT" > "HOLLight.hollight.ACS_MONO_LT" + "ACS_COS" > "HOLLight.hollight.ACS_COS" + "ACS_BOUNDS_LT" > "HOLLight.hollight.ACS_BOUNDS_LT" + "ACS_BOUNDS" > "HOLLight.hollight.ACS_BOUNDS" + "ACS" > "HOLLight.hollight.ACS" + "ABS_ZERO" > "HOLLight.hollight.ABS_ZERO" + "ABS_TRIANGLE" > "HOLLight.hollight.ABS_TRIANGLE" + "ABS_SUM" > "HOLLight.hollight.ABS_SUM" + "ABS_SUB_ABS" > "HOLLight.hollight.ABS_SUB_ABS" + "ABS_SUB" > "HOLLight.hollight.ABS_SUB" + "ABS_STILLNZ" > "HOLLight.hollight.ABS_STILLNZ" + "ABS_SIMP" > "Presburger.fm_modd_pinf" + "ABS_SIGN2" > "HOLLight.hollight.ABS_SIGN2" + "ABS_SIGN" > "HOLLight.hollight.ABS_SIGN" + "ABS_REFL" > "HOLLight.hollight.ABS_REFL" + "ABS_POW2" > "HOLLight.hollight.ABS_POW2" + "ABS_POS" > "HOLLight.hollight.ABS_POS" + "ABS_NZ" > "HOLLight.hollight.ABS_NZ" + "ABS_NEG_LEMMA" > "HOLLight.hollight.ABS_NEG_LEMMA" + "ABS_NEG" > "HOLLight.hollight.ABS_NEG" + "ABS_N" > "HOLLight.hollight.ABS_N" + "ABS_MUL" > "HOLLight.hollight.ABS_MUL" + "ABS_LT_MUL2" > "HOLLight.hollight.ABS_LT_MUL2" + "ABS_LE" > "HOLLight.hollight.ABS_LE" + "ABS_INV" > "HOLLight.hollight.ABS_INV" + "ABS_DIV" > "HOLLight.hollight.ABS_DIV" + "ABS_CIRCLE" > "HOLLight.hollight.ABS_CIRCLE" + "ABS_CASES" > "HOLLight.hollight.ABS_CASES" + "ABS_BOUNDS" > "HOLLight.hollight.ABS_BOUNDS" + "ABS_BOUND" > "HOLLight.hollight.ABS_BOUND" + "ABS_BETWEEN2" > "HOLLight.hollight.ABS_BETWEEN2" + "ABS_BETWEEN1" > "HOLLight.hollight.ABS_BETWEEN1" + "ABS_BETWEEN" > "HOLLight.hollight.ABS_BETWEEN" + "ABS_ABS" > "HOLLight.hollight.ABS_ABS" + "ABS_1" > "HOLLight.hollight.ABS_1" + "ABS_0" > "HOLLight.hollight.ABS_0" + "ABSORPTION" > "HOLLight.hollight.ABSORPTION" + ">_def" > "HOLLight.hollight.>_def" + ">=_def" > "HOLLight.hollight.>=_def" + "<_def" > "HOLLight.hollight.<_def" + "<=_def" > "HOLLight.hollight.<=_def" + "$_def" > "HOLLight.hollight.$_def" + +ignore_thms + "TYDEF_prod" + "TYDEF_num" + "TYDEF_1" + "IND_SUC_0_EXISTS" + "DEF_one" + "DEF__0" + "DEF_SUC" + "DEF_NUM_REP" + "DEF_IND_SUC" + "DEF_IND_0" + +end diff -r 7e417a7cbf9f -r bd59bfd4bf37 src/HOL/Import/hol4rews.ML --- a/src/HOL/Import/hol4rews.ML Mon Sep 26 02:06:44 2005 +0200 +++ b/src/HOL/Import/hol4rews.ML Mon Sep 26 02:27:14 2005 +0200 @@ -333,9 +333,10 @@ end fun add_hol4_mapping bthy bthm isathm thy = - let - val isathm = follow_name isathm thy - val _ = message ("Adding theorem map: " ^ bthy ^ "." ^ bthm ^ " --> " ^ isathm) + let + (* val _ = writeln ("Before follow_name: "^isathm) *) + val isathm = follow_name isathm thy + (* val _ = writeln ("Adding theorem map: " ^ bthy ^ "." ^ bthm ^ " --> " ^ isathm)*) val curmaps = HOL4Maps.get thy val newmaps = StringPair.update_new ((bthy,bthm),SOME isathm) curmaps val upd_thy = HOL4Maps.put newmaps thy diff -r 7e417a7cbf9f -r bd59bfd4bf37 src/HOL/Import/import_package.ML --- a/src/HOL/Import/import_package.ML Mon Sep 26 02:06:44 2005 +0200 +++ b/src/HOL/Import/import_package.ML Mon Sep 26 02:27:14 2005 +0200 @@ -41,6 +41,7 @@ let val sg = sign_of_thm th val prem = hd (prems_of th) + val _ = message ("Import_tac: thyname="^thyname^", thmname="^thmname) val _ = message ("Import trying to prove " ^ (string_of_cterm (cterm_of sg prem))) val int_thms = case ImportData.get thy of NONE => fst (Replay.setup_int_thms thyname thy) @@ -52,6 +53,7 @@ val thm = equal_elim rew thm val prew = ProofKernel.rewrite_hol4_term prem thy val prem' = #2 (Logic.dest_equals (prop_of prew)) + val thm = Drule.disambiguate_frees thm val _ = message ("Import proved " ^ (string_of_thm thm)) in case Shuffler.set_prop thy prem' [("",thm)] of diff -r 7e417a7cbf9f -r bd59bfd4bf37 src/HOL/Import/import_syntax.ML --- a/src/HOL/Import/import_syntax.ML Mon Sep 26 02:06:44 2005 +0200 +++ b/src/HOL/Import/import_syntax.ML Mon Sep 26 02:27:14 2005 +0200 @@ -50,6 +50,8 @@ val segname = get_import_segment thy val (int_thms,facts) = Replay.setup_int_thms thyname thy val thmnames = List.filter (not o should_ignore thyname thy) facts + (* val _ = set show_types + val _ = set show_sorts*) in thy |> clear_import_thy |> set_segment thyname segname diff -r 7e417a7cbf9f -r bd59bfd4bf37 src/HOL/Import/proof_kernel.ML --- a/src/HOL/Import/proof_kernel.ML Mon Sep 26 02:06:44 2005 +0200 +++ b/src/HOL/Import/proof_kernel.ML Mon Sep 26 02:27:14 2005 +0200 @@ -1290,13 +1290,14 @@ let val sg = sign_of thy val (hth' as HOLThm (args as (_,th))) = norm_hthm sg hth - val _ = if has_ren hth' then warning ("Theorem " ^ thmname ^ " needs variable-disambiguating") - else () val rew = rewrite_hol4_term (concl_of th) thy val th = equal_elim rew th val thy' = add_hol4_pending thyname thmname args thy + val th = Drule.disambiguate_frees th val thy2 = if gen_output - then add_dump ("lemma " ^ (quotename thmname) ^ ": " ^ (smart_string_of_thm th) ^ "\n by (import " ^ thyname ^ " " ^ (quotename thmname) ^ ")") thy' + then add_dump ("lemma " ^ (quotename thmname) ^ ": " ^ + (smart_string_of_thm th) ^ "\n by (import " ^ + thyname ^ " " ^ (quotename thmname) ^ ")") thy' else thy' in (thy2,hth')