diff -r 01b8c8d17f13 -r 2a77bc3b4eac src/CCL/Term.thy --- a/src/CCL/Term.thy Fri Sep 20 23:36:33 2024 +0200 +++ b/src/CCL/Term.thy Fri Sep 20 23:37:00 2024 +0200 @@ -12,7 +12,8 @@ definition one :: "i" where "one == true" -definition "if" :: "[i,i,i]\i" (\(3if _/ then _/ else _)\ [0,0,60] 60) +definition "if" :: "[i,i,i]\i" + (\(\indent=3 notation=\mixfix if then else\\if _/ then _/ else _)\ [0,0,60] 60) where "if b then t else u == case(b, t, u, \ x y. bot, \v. bot)" definition inl :: "i\i" @@ -48,7 +49,8 @@ definition "let1" :: "[i,i\i]\i" where let_def: "let1(t, f) == case(t,f(true),f(false), \x y. f(), \u. f(lam x. u(x)))" -syntax "_let1" :: "[idt,i,i]\i" (\(3let _ be _/ in _)\ [0,0,60] 60) +syntax "_let1" :: "[idt,i,i]\i" + (\(\indent=3 notation=\mixfix let be in\\let _ be _/ in _)\ [0,0,60] 60) syntax_consts "_let1" == let1 translations "let x be a in e" == "CONST let1(a, \x. e)" @@ -65,9 +67,12 @@ \g'. f(\x y z. g'(>)))" syntax - "_letrec" :: "[idt,idt,i,i]\i" (\(3letrec _ _ be _/ in _)\ [0,0,0,60] 60) - "_letrec2" :: "[idt,idt,idt,i,i]\i" (\(3letrec _ _ _ be _/ in _)\ [0,0,0,0,60] 60) - "_letrec3" :: "[idt,idt,idt,idt,i,i]\i" (\(3letrec _ _ _ _ be _/ in _)\ [0,0,0,0,0,60] 60) + "_letrec" :: "[idt,idt,i,i]\i" + (\(\indent=3 notation=\mixfix letrec be in\\letrec _ _ be _/ in _)\ [0,0,0,60] 60) + "_letrec2" :: "[idt,idt,idt,i,i]\i" + (\(\indent=3 notation=\mixfix letrec be in\\letrec _ _ _ be _/ in _)\ [0,0,0,0,60] 60) + "_letrec3" :: "[idt,idt,idt,idt,i,i]\i" + (\(\indent=3 notation=\mixfix letrec be in\\letrec _ _ _ _ be _/ in _)\ [0,0,0,0,0,60] 60) syntax_consts "_letrec" == letrec and "_letrec2" == letrec2 and @@ -143,7 +148,7 @@ 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)" -definition napply :: "[i\i,i,i]\i" (\(_ ^ _ ` _)\ [56,56,56] 56) +definition napply :: "[i\i,i,i]\i" (\(\notation=\mixfix napply\\_ ^ _ ` _)\ [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