# HG changeset patch # User wenzelm # Date 1452543398 -3600 # Node ID 3c9a0985e6be9f08e2c997ef3694bb6b9a26ecca # Parent eed7ca453573b773e83997e59e62d987952623bc eliminated old defs; diff -r eed7ca453573 -r 3c9a0985e6be src/CCL/CCL.thy --- a/src/CCL/CCL.thy Mon Jan 11 07:44:20 2016 +0100 +++ b/src/CCL/CCL.thy Mon Jan 11 21:16:38 2016 +0100 @@ -320,17 +320,17 @@ apply blast done -lemma POgenXH: - " : POgen(R) \ t= bot | (t=true \ t'=true) | (t=false \ t'=false) | - (EX a a' b b'. t= \ t'= \ : R \ : R) | +lemma POgenXH: + " : POgen(R) \ t= bot | (t=true \ t'=true) | (t=false \ t'=false) | + (EX a a' b b'. t= \ t'= \ : R \ : R) | (EX f f'. t=lam x. f(x) \ t'=lam x. f'(x) \ (ALL x. : R))" apply (unfold POgen_def SIM_def) apply (rule iff_refl [THEN XHlemma2]) done -lemma poXH: - "t [= t' \ t=bot | (t=true \ t'=true) | (t=false \ t'=false) | - (EX a a' b b'. t= \ t'= \ a [= a' \ b [= b') | +lemma poXH: + "t [= t' \ t=bot | (t=true \ t'=true) | (t=false \ t'=false) | + (EX a a' b b'. t= \ t'= \ a [= a' \ b [= b') | (EX f f'. t=lam x. f(x) \ t'=lam x. f'(x) \ (ALL x. f(x) [= f'(x)))" apply (simp add: PO_iff del: ex_simps) apply (rule POgen_mono @@ -438,18 +438,18 @@ apply blast done -lemma EQgenXH: - " : EQgen(R) \ (t=bot \ t'=bot) | (t=true \ t'=true) | - (t=false \ t'=false) | - (EX a a' b b'. t= \ t'= \ : R \ : R) | +lemma EQgenXH: + " : EQgen(R) \ (t=bot \ t'=bot) | (t=true \ t'=true) | + (t=false \ t'=false) | + (EX a a' b b'. t= \ t'= \ : R \ : R) | (EX f f'. t=lam x. f(x) \ t'=lam x. f'(x) \ (ALL x. : R))" apply (unfold EQgen_def SIM_def) apply (rule iff_refl [THEN XHlemma2]) done -lemma eqXH: - "t=t' \ (t=bot \ t'=bot) | (t=true \ t'=true) | (t=false \ t'=false) | - (EX a a' b b'. t= \ t'= \ a=a' \ b=b') | +lemma eqXH: + "t=t' \ (t=bot \ t'=bot) | (t=true \ t'=true) | (t=false \ t'=false) | + (EX a a' b b'. t= \ t'= \ a=a' \ b=b') | (EX f f'. t=lam x. f(x) \ t'=lam x. f'(x) \ (ALL x. f(x)=f'(x)))" apply (subgoal_tac " : EQ \ (t=bot \ t'=bot) | (t=true \ t'=true) | (t=false \ t'=false) | diff -r eed7ca453573 -r 3c9a0985e6be src/CCL/Set.thy --- a/src/CCL/Set.thy Mon Jan 11 07:44:20 2016 +0100 +++ b/src/CCL/Set.thy Mon Jan 11 21:16:38 2016 +0100 @@ -9,62 +9,18 @@ typedecl 'a set instance set :: ("term") "term" .. -consts - Collect :: "['a \ o] \ 'a set" (*comprehension*) - Compl :: "('a set) \ 'a set" (*complement*) - Int :: "['a set, 'a set] \ 'a set" (infixl "Int" 70) - Un :: "['a set, 'a set] \ 'a set" (infixl "Un" 65) - Union :: "(('a set)set) \ 'a set" (*...of a set*) - Inter :: "(('a set)set) \ 'a set" (*...of a set*) - UNION :: "['a set, 'a \ 'b set] \ 'b set" (*general*) - INTER :: "['a set, 'a \ 'b set] \ 'b set" (*general*) - Ball :: "['a set, 'a \ o] \ o" (*bounded quants*) - Bex :: "['a set, 'a \ o] \ o" (*bounded quants*) - mono :: "['a set \ 'b set] \ o" (*monotonicity*) - mem :: "['a, 'a set] \ o" (infixl ":" 50) (*membership*) - subset :: "['a set, 'a set] \ o" (infixl "<=" 50) - singleton :: "'a \ 'a set" ("{_}") - empty :: "'a set" ("{}") + +subsection \Set comprehension and membership\ + +axiomatization Collect :: "['a \ o] \ 'a set" + and mem :: "['a, 'a set] \ o" (infixl ":" 50) +where mem_Collect_iff: "(a : Collect(P)) \ P(a)" + and set_extension: "A = B \ (ALL x. x:A \ x:B)" syntax - "_Coll" :: "[idt, o] \ 'a set" ("(1{_./ _})") (*collection*) - - (* Big Intersection / Union *) - - "_INTER" :: "[idt, 'a set, 'b set] \ 'b set" ("(INT _:_./ _)" [0, 0, 0] 10) - "_UNION" :: "[idt, 'a set, 'b set] \ 'b set" ("(UN _:_./ _)" [0, 0, 0] 10) - - (* Bounded Quantifiers *) - - "_Ball" :: "[idt, 'a set, o] \ o" ("(ALL _:_./ _)" [0, 0, 0] 10) - "_Bex" :: "[idt, 'a set, o] \ o" ("(EX _:_./ _)" [0, 0, 0] 10) - + "_Coll" :: "[idt, o] \ 'a set" ("(1{_./ _})") translations - "{x. P}" == "CONST Collect(\x. P)" - "INT x:A. B" == "CONST INTER(A, \x. B)" - "UN x:A. B" == "CONST UNION(A, \x. B)" - "ALL x:A. P" == "CONST Ball(A, \x. P)" - "EX x:A. P" == "CONST Bex(A, \x. P)" - -axiomatization where - mem_Collect_iff: "(a : {x. P(x)}) \ P(a)" and - set_extension: "A = B \ (ALL x. x:A \ x:B)" - -defs - Ball_def: "Ball(A, P) == ALL x. x:A \ P(x)" - Bex_def: "Bex(A, P) == EX x. x:A \ P(x)" - mono_def: "mono(f) == (ALL A B. A <= B \ f(A) <= f(B))" - subset_def: "A <= B == ALL x:A. x:B" - singleton_def: "{a} == {x. x=a}" - empty_def: "{} == {x. False}" - Un_def: "A Un B == {x. x:A | x:B}" - Int_def: "A Int B == {x. x:A \ x:B}" - Compl_def: "Compl(A) == {x. \x:A}" - INTER_def: "INTER(A, B) == {y. ALL x:A. y: B(x)}" - UNION_def: "UNION(A, B) == {y. EX x:A. y: B(x)}" - Inter_def: "Inter(S) == (INT x:S. x)" - Union_def: "Union(S) == (UN x:S. x)" - + "{x. P}" == "CONST Collect(\x. P)" lemma CollectI: "P(a) \ a : {x. P(x)}" apply (rule mem_Collect_iff [THEN iffD2]) @@ -85,6 +41,19 @@ subsection \Bounded quantifiers\ +definition Ball :: "['a set, 'a \ o] \ o" + where "Ball(A, P) == ALL x. x:A \ P(x)" + +definition Bex :: "['a set, 'a \ o] \ o" + where "Bex(A, P) == EX x. x:A \ P(x)" + +syntax + "_Ball" :: "[idt, 'a set, o] \ o" ("(ALL _:_./ _)" [0, 0, 0] 10) + "_Bex" :: "[idt, 'a set, o] \ o" ("(EX _:_./ _)" [0, 0, 0] 10) +translations + "ALL x:A. P" == "CONST Ball(A, \x. P)" + "EX x:A. P" == "CONST Bex(A, \x. P)" + lemma ballI: "(\x. x:A \ P(x)) \ ALL x:A. P(x)" by (simp add: Ball_def) @@ -107,8 +76,7 @@ lemma ball_rew: "(ALL x:A. True) \ True" by (blast intro: ballI) - -subsection \Congruence rules\ +subsubsection \Congruence rules\ lemma ball_cong: "\A = A'; \x. x:A' \ P(x) \ P'(x)\ \ @@ -121,6 +89,52 @@ by (blast intro: bexI elim: bexE) +subsection \Further operations\ + +definition subset :: "['a set, 'a set] \ o" (infixl "<=" 50) + where "A <= B == ALL x:A. x:B" + +definition mono :: "['a set \ 'b set] \ o" + where "mono(f) == (ALL A B. A <= B \ f(A) <= f(B))" + +definition singleton :: "'a \ 'a set" ("{_}") + where "{a} == {x. x=a}" + +definition empty :: "'a set" ("{}") + where "{} == {x. False}" + +definition Un :: "['a set, 'a set] \ 'a set" (infixl "Un" 65) + where "A Un B == {x. x:A | x:B}" + +definition Int :: "['a set, 'a set] \ 'a set" (infixl "Int" 70) + where "A Int B == {x. x:A \ x:B}" + +definition Compl :: "('a set) \ 'a set" + where "Compl(A) == {x. \x:A}" + + +subsection \Big Intersection / Union\ + +definition INTER :: "['a set, 'a \ 'b set] \ 'b set" + where "INTER(A, B) == {y. ALL x:A. y: B(x)}" + +definition UNION :: "['a set, 'a \ 'b set] \ 'b set" + where "UNION(A, B) == {y. EX x:A. y: B(x)}" + +syntax + "_INTER" :: "[idt, 'a set, 'b set] \ 'b set" ("(INT _:_./ _)" [0, 0, 0] 10) + "_UNION" :: "[idt, 'a set, 'b set] \ 'b set" ("(UN _:_./ _)" [0, 0, 0] 10) +translations + "INT x:A. B" == "CONST INTER(A, \x. B)" + "UN x:A. B" == "CONST UNION(A, \x. B)" + +definition Inter :: "(('a set)set) \ 'a set" + where "Inter(S) == (INT x:S. x)" + +definition Union :: "(('a set)set) \ 'a set" + where "Union(S) == (UN x:S. x)" + + subsection \Rules for subsets\ lemma subsetI: "(\x. x:A \ x:B) \ A <= B" diff -r eed7ca453573 -r 3c9a0985e6be src/CCL/Term.thy --- a/src/CCL/Term.thy Mon Jan 11 07:44:20 2016 +0100 +++ b/src/CCL/Term.thy Mon Jan 11 21:16:38 2016 +0100 @@ -9,48 +9,69 @@ imports CCL begin -consts +definition one :: "i" + where "one == true" - one :: "i" +definition "if" :: "[i,i,i]\i" ("(3if _/ then _/ else _)" [0,0,60] 60) + where "if b then t else u == case(b, t, u, \ x y. bot, \v. bot)" - "if" :: "[i,i,i]\i" ("(3if _/ then _/ else _)" [0,0,60] 60) +definition inl :: "i\i" + where "inl(a) == " - inl :: "i\i" - inr :: "i\i" - "when" :: "[i,i\i,i\i]\i" +definition inr :: "i\i" + where "inr(b) == " + +definition split :: "[i,[i,i]\i]\i" + where "split(t,f) == case(t, bot, bot, f, \u. bot)" - split :: "[i,[i,i]\i]\i" - fst :: "i\i" - snd :: "i\i" - thd :: "i\i" +definition "when" :: "[i,i\i,i\i]\i" + where "when(t,f,g) == split(t, \b x. if b then f(x) else g(x))" + +definition fst :: "i\i" + where "fst(t) == split(t, \x y. x)" - zero :: "i" - succ :: "i\i" - ncase :: "[i,i,i\i]\i" - nrec :: "[i,i,[i,i]\i]\i" +definition snd :: "i\i" + where "snd(t) == split(t, \x y. y)" + +definition thd :: "i\i" + where "thd(t) == split(t, \x p. split(p, \y z. z))" + +definition zero :: "i" + where "zero == inl(one)" - nil :: "i" ("([])") - cons :: "[i,i]\i" (infixr "$" 80) - lcase :: "[i,i,[i,i]\i]\i" - lrec :: "[i,i,[i,i,i]\i]\i" +definition succ :: "i\i" + where "succ(n) == inr(n)" + +definition ncase :: "[i,i,i\i]\i" + where "ncase(n,b,c) == when(n, \x. b, \y. c(y))" - "let" :: "[i,i\i]\i" - letrec :: "[[i,i\i]\i,(i\i)\i]\i" - letrec2 :: "[[i,i,i\i\i]\i,(i\i\i)\i]\i" - letrec3 :: "[[i,i,i,i\i\i\i]\i,(i\i\i\i)\i]\i" + +no_syntax + "_Let" :: "[letbinds, 'a] => 'a" ("(let (_)/ in (_))" 10) + +definition "let" :: "[i,i\i]\i" + where "let(t, f) == case(t,f(true),f(false), \x y. f(), \u. f(lam x. u(x)))" syntax - "_let" :: "[id,i,i]\i" ("(3let _ be _/ in _)" - [0,0,60] 60) + "_Let" :: "[letbinds, 'a] => 'a" ("(let (_)/ in (_))" 10) - "_letrec" :: "[id,id,i,i]\i" ("(3letrec _ _ be _/ in _)" - [0,0,0,60] 60) +definition letrec :: "[[i,i\i]\i,(i\i)\i]\i" + where "letrec(h, b) == b(\x. fix(\f. lam x. h(x,\y. f`y))`x)" + +definition letrec2 :: "[[i,i,i\i\i]\i,(i\i\i)\i]\i" + where "letrec2 (h, f) == + letrec (\p g'. split(p,\x y. h(x,y,\u v. g'())), \g'. f(\x y. g'()))" - "_letrec2" :: "[id,id,id,i,i]\i" ("(3letrec _ _ _ be _/ in _)" - [0,0,0,0,60] 60) +definition letrec3 :: "[[i,i,i,i\i\i\i]\i,(i\i\i\i)\i]\i" + where "letrec3 (h, f) == + letrec (\p g'. split(p,\x xs. split(xs,\y z. h(x,y,z,\u v w. g'(>)))), + \g'. f(\x y z. g'(>)))" - "_letrec3" :: "[id,id,id,id,i,i]\i" ("(3letrec _ _ _ _ be _/ in _)" - [0,0,0,0,0,60] 60) +syntax + "_let" :: "[id,i,i]\i" ("(3let _ be _/ in _)" [0,0,60] 60) + "_letrec" :: "[id,id,i,i]\i" ("(3letrec _ _ be _/ in _)" [0,0,0,60] 60) + "_letrec2" :: "[id,id,id,i,i]\i" ("(3letrec _ _ _ be _/ in _)" [0,0,0,0,60] 60) + "_letrec3" :: "[id,id,id,id,i,i]\i" ("(3letrec _ _ _ _ be _/ in _)" [0,0,0,0,0,60] 60) ML \ (** Quantifier translations: variable binding **) @@ -107,42 +128,24 @@ (@{const_syntax letrec3}, K letrec3_tr')] \ -consts - napply :: "[i\i,i,i]\i" ("(_ ^ _ ` _)" [56,56,56] 56) + +definition nrec :: "[i,i,[i,i]\i]\i" + where "nrec(n,b,c) == letrec g x be ncase(x, b, \y. c(y,g(y))) in g(n)" + +definition nil :: "i" ("([])") + where "[] == inl(one)" -defs - one_def: "one == true" - if_def: "if b then t else u == case(b, t, u, \ x y. bot, \v. bot)" - inl_def: "inl(a) == " - inr_def: "inr(b) == " - when_def: "when(t,f,g) == split(t, \b x. if b then f(x) else g(x))" - split_def: "split(t,f) == case(t, bot, bot, f, \u. bot)" - fst_def: "fst(t) == split(t, \x y. x)" - snd_def: "snd(t) == split(t, \x y. y)" - thd_def: "thd(t) == split(t, \x p. split(p, \y z. z))" - zero_def: "zero == inl(one)" - succ_def: "succ(n) == inr(n)" - ncase_def: "ncase(n,b,c) == when(n, \x. b, \y. c(y))" - nrec_def: " nrec(n,b,c) == letrec g x be ncase(x, b, \y. c(y,g(y))) in g(n)" - nil_def: "[] == inl(one)" - cons_def: "h$t == inr()" - lcase_def: "lcase(l,b,c) == when(l, \x. b, \y. split(y,c))" - lrec_def: "lrec(l,b,c) == letrec g x be lcase(x, b, \h t. c(h,t,g(t))) in g(l)" +definition cons :: "[i,i]\i" (infixr "$" 80) + where "h$t == inr()" + +definition lcase :: "[i,i,[i,i]\i]\i" + where "lcase(l,b,c) == when(l, \x. b, \y. split(y,c))" - let_def: "let x be t in f(x) == case(t,f(true),f(false), \x y. f(), \u. f(lam x. u(x)))" - letrec_def: - "letrec g x be h(x,g) in b(g) == b(\x. fix(\f. lam x. h(x,\y. f`y))`x)" +definition lrec :: "[i,i,[i,i,i]\i]\i" + where "lrec(l,b,c) == letrec g x be lcase(x, b, \h t. c(h,t,g(t))) in g(l)" - letrec2_def: "letrec g x y be h(x,y,g) in f(g)== - letrec g' p be split(p,\x y. h(x,y,\u v. g'())) - in f(\x y. g'())" - - letrec3_def: "letrec g x y z be h(x,y,z,g) in f(g) == - letrec g' p be split(p,\x xs. split(xs,\y z. h(x,y,z,\u v w. g'(>)))) - in f(\x y z. g'(>))" - - napply_def: "f ^n` a == nrec(n,a,\x g. f(g))" - +definition napply :: "[i\i,i,i]\i" ("(_ ^ _ ` _)" [56,56,56] 56) + where "f ^n` a == nrec(n,a,\x g. f(g))" lemmas simp_can_defs = one_def inl_def inr_def and simp_ncan_defs = if_def when_def split_def fst_def snd_def thd_def diff -r eed7ca453573 -r 3c9a0985e6be src/CCL/Type.thy --- a/src/CCL/Type.thy Mon Jan 11 07:44:20 2016 +0100 +++ b/src/CCL/Type.thy Mon Jan 11 21:16:38 2016 +0100 @@ -9,42 +9,39 @@ imports Term begin -consts - - Subtype :: "['a set, 'a \ o] \ 'a set" - Bool :: "i set" - Unit :: "i set" - Plus :: "[i set, i set] \ i set" (infixr "+" 55) - Pi :: "[i set, i \ i set] \ i set" - Sigma :: "[i set, i \ i set] \ i set" - Nat :: "i set" - List :: "i set \ i set" - Lists :: "i set \ i set" - ILists :: "i set \ i set" - TAll :: "(i set \ i set) \ i set" (binder "TALL " 55) - TEx :: "(i set \ i set) \ i set" (binder "TEX " 55) - Lift :: "i set \ i set" ("(3[_])") - - SPLIT :: "[i, [i, i] \ i set] \ i set" +definition Subtype :: "['a set, 'a \ o] \ 'a set" + where "Subtype(A, P) == {x. x:A \ P(x)}" syntax - "_Pi" :: "[idt, i set, i set] \ i set" ("(3PROD _:_./ _)" - [0,0,60] 60) + "_Subtype" :: "[idt, 'a set, o] \ 'a set" ("(1{_: _ ./ _})") +translations + "{x: A. B}" == "CONST Subtype(A, \x. B)" - "_Sigma" :: "[idt, i set, i set] \ i set" ("(3SUM _:_./ _)" - [0,0,60] 60) +definition Unit :: "i set" + where "Unit == {x. x=one}" + +definition Bool :: "i set" + where "Bool == {x. x=true | x=false}" + +definition Plus :: "[i set, i set] \ i set" (infixr "+" 55) + where "A+B == {x. (EX a:A. x=inl(a)) | (EX b:B. x=inr(b))}" - "_arrow" :: "[i set, i set] \ i set" ("(_ ->/ _)" [54, 53] 53) - "_star" :: "[i set, i set] \ i set" ("(_ */ _)" [56, 55] 55) - "_Subtype" :: "[idt, 'a set, o] \ 'a set" ("(1{_: _ ./ _})") +definition Pi :: "[i set, i \ i set] \ i set" + where "Pi(A,B) == {x. EX b. x=lam x. b(x) \ (ALL x:A. b(x):B(x))}" + +definition Sigma :: "[i set, i \ i set] \ i set" + where "Sigma(A,B) == {x. EX a:A. EX b:B(a).x=}" +syntax + "_Pi" :: "[idt, i set, i set] \ i set" ("(3PROD _:_./ _)" [0,0,60] 60) + "_Sigma" :: "[idt, i set, i set] \ i set" ("(3SUM _:_./ _)" [0,0,60] 60) + "_arrow" :: "[i set, i set] \ i set" ("(_ ->/ _)" [54, 53] 53) + "_star" :: "[i set, i set] \ i set" ("(_ */ _)" [56, 55] 55) translations - "PROD x:A. B" => "CONST Pi(A, \x. B)" - "A -> B" => "CONST Pi(A, \_. B)" - "SUM x:A. B" => "CONST Sigma(A, \x. B)" - "A * B" => "CONST Sigma(A, \_. B)" - "{x: A. B}" == "CONST Subtype(A, \x. B)" - + "PROD x:A. B" \ "CONST Pi(A, \x. B)" + "A -> B" \ "CONST Pi(A, \_. B)" + "SUM x:A. B" \ "CONST Sigma(A, \x. B)" + "A * B" \ "CONST Sigma(A, \_. B)" print_translation \ [(@{const_syntax Pi}, fn _ => Syntax_Trans.dependent_tr' (@{syntax_const "_Pi"}, @{syntax_const "_arrow"})), @@ -52,28 +49,34 @@ fn _ => Syntax_Trans.dependent_tr' (@{syntax_const "_Sigma"}, @{syntax_const "_star"}))] \ -defs - Subtype_def: "{x:A. P(x)} == {x. x:A \ P(x)}" - Unit_def: "Unit == {x. x=one}" - Bool_def: "Bool == {x. x=true | x=false}" - Plus_def: "A+B == {x. (EX a:A. x=inl(a)) | (EX b:B. x=inr(b))}" - Pi_def: "Pi(A,B) == {x. EX b. x=lam x. b(x) \ (ALL x:A. b(x):B(x))}" - Sigma_def: "Sigma(A,B) == {x. EX a:A. EX b:B(a).x=}" - Nat_def: "Nat == lfp(\X. Unit + X)" - List_def: "List(A) == lfp(\X. Unit + A*X)" +definition Nat :: "i set" + where "Nat == lfp(\X. Unit + X)" + +definition List :: "i set \ i set" + where "List(A) == lfp(\X. Unit + A*X)" + +definition Lists :: "i set \ i set" + where "Lists(A) == gfp(\X. Unit + A*X)" + +definition ILists :: "i set \ i set" + where "ILists(A) == gfp(\X.{} + A*X)" - Lists_def: "Lists(A) == gfp(\X. Unit + A*X)" - ILists_def: "ILists(A) == gfp(\X.{} + A*X)" + +definition TAll :: "(i set \ i set) \ i set" (binder "TALL " 55) + where "TALL X. B(X) == Inter({X. EX Y. X=B(Y)})" - Tall_def: "TALL X. B(X) == Inter({X. EX Y. X=B(Y)})" - Tex_def: "TEX X. B(X) == Union({X. EX Y. X=B(Y)})" - Lift_def: "[A] == A Un {bot}" +definition TEx :: "(i set \ i set) \ i set" (binder "TEX " 55) + where "TEX X. B(X) == Union({X. EX Y. X=B(Y)})" - SPLIT_def: "SPLIT(p,B) == Union({A. EX x y. p= \ A=B(x,y)})" +definition Lift :: "i set \ i set" ("(3[_])") + where "[A] == A Un {bot}" + +definition SPLIT :: "[i, [i, i] \ i set] \ i set" + where "SPLIT(p,B) == Union({A. EX x y. p= \ A=B(x,y)})" lemmas simp_type_defs = - Subtype_def Unit_def Bool_def Plus_def Sigma_def Pi_def Lift_def Tall_def Tex_def + Subtype_def Unit_def Bool_def Plus_def Sigma_def Pi_def Lift_def TAll_def TEx_def and ind_type_defs = Nat_def List_def and simp_data_defs = one_def inl_def inr_def and ind_data_defs = zero_def succ_def nil_def cons_def