diff -r 4a584b094aba -r 0cb35fa9c6fa src/HOL/Lambda/Type.thy --- a/src/HOL/Lambda/Type.thy Fri Jan 25 23:05:24 2008 +0100 +++ b/src/HOL/Lambda/Type.thy Fri Jan 25 23:05:25 2008 +0100 @@ -56,12 +56,14 @@ "e \ t \ u : T" "e \ Abs t : T" -consts +primrec typings :: "(nat \ type) \ dB list \ type list \ bool" - -abbreviation - funs :: "type list \ type \ type" (infixr "=>>" 200) where - "Ts =>> T == foldr Fun Ts T" +where + "typings e [] Ts = (Ts = [])" + | "typings e (t # ts) Ts = + (case Ts of + [] \ False + | T # Ts \ e \ t : T \ typings e ts Ts)" abbreviation typings_rel :: "(nat \ type) \ dB list \ type list \ bool" @@ -69,15 +71,14 @@ "env ||- ts : Ts == typings env ts Ts" notation (latex) - funs (infixr "\" 200) and typings_rel ("_ \ _ : _" [50, 50, 50] 50) -primrec - "(e \ [] : Ts) = (Ts = [])" - "(e \ (t # ts) : Ts) = - (case Ts of - [] \ False - | T # Ts \ e \ t : T \ e \ ts : Ts)" +abbreviation + funs :: "type list \ type \ type" (infixr "=>>" 200) where + "Ts =>> T == foldr Fun Ts T" + +notation (latex) + funs (infixr "\" 200) subsection {* Some examples *} @@ -123,11 +124,11 @@ apply simp+ done -lemma rev_exhaust2 [case_names Nil snoc, extraction_expand]: - "(xs = [] \ P) \ (\ys y. xs = ys @ [y] \ P) \ P" +lemma rev_exhaust2 [extraction_expand]: + obtains (Nil) "xs = []" | (snoc) ys y where "xs = ys @ [y]" -- {* Cannot use @{text rev_exhaust} from the @{text List} theory, since it is not constructive *} - apply (subgoal_tac "\ys. xs = rev ys \ P") + apply (subgoal_tac "\ys. xs = rev ys \ thesis") apply (erule_tac x="rev xs" in allE) apply simp apply (rule allI)