From 07c51b1801bd701bef61cedd571a944d9bd37e8b Mon Sep 17 00:00:00 2001 From: Josh Chen Date: Wed, 8 Jul 2020 15:55:48 +0200 Subject: 1. Initial `Definition` keyword. 2. ifelse. --- spartan/lib/More_Types.thy | 50 ++++++++++++++++++++++++++++++++++++++++++---- 1 file changed, 46 insertions(+), 4 deletions(-) (limited to 'spartan/lib/More_Types.thy') diff --git a/spartan/lib/More_Types.thy b/spartan/lib/More_Types.thy index 1d7abb9..38ba2aa 100644 --- a/spartan/lib/More_Types.thy +++ b/spartan/lib/More_Types.thy @@ -1,5 +1,3 @@ -chapter \Some standard types\ - theory More_Types imports Spartan @@ -96,9 +94,53 @@ Lemma Bool_false: "false: Bool" unfolding Bool_def true_def false_def by typechk+ -lemmas [intros] = BoolF Bool_true Bool_false +\ \Definitions like these should be handled by a future function package\ +Definition ifelse [rotated 1]: + assumes *[unfolded Bool_def true_def false_def]: + "\x. x: Bool \ C x: U i" + "x: Bool" + "a: C true" + "b: C false" + shows "C x" + by (elim x) (elim, rule *)+ + +Lemma if_true: + assumes + "\x. x: Bool \ C x: U i" + "a: C true" + "b: C false" + shows "ifelse C true a b \ a" + unfolding ifelse_def true_def + supply assms[unfolded Bool_def true_def false_def] + by reduce + +Lemma if_false: + assumes + "\x. x: Bool \ C x: U i" + "a: C true" + "b: C false" + shows "ifelse C false a b \ b" + unfolding ifelse_def false_def + supply assms[unfolded Bool_def true_def false_def] + by reduce + +lemmas + [intros] = BoolF Bool_true Bool_false and + [comps] = if_true if_false and + [elims ?x] = ifelse +lemmas + BoolE = ifelse + +subsection \Notation\ + +definition ifelse_i ("if _ then _ else _") + where [implicit]: "if x then a else b \ ifelse ? x a b" + +no_translations "if x then a else b" \ "CONST ifelse C x a b" + +subsection \Logical connectives\ -\ \Can define if-then-else etc.\ +definition not ("!_") where "not x \ ifelse (K Bool) x false true" end -- cgit v1.2.3