aboutsummaryrefslogtreecommitdiff
path: root/spartan/theories/Spartan.thy
diff options
context:
space:
mode:
authorJosh Chen2020-04-02 23:27:06 +0200
committerJosh Chen2020-04-02 23:27:06 +0200
commit2781c68f0fdb435827097efc497c2172d6050e50 (patch)
tree4f0e05000a490b9b9e33e8d082d96bfd5d83152f /spartan/theories/Spartan.thy
parent0ddab0fe11c33fc559fc8fb58528618efdbc93a4 (diff)
1. make id function an abbrev. 2. fix reduce method
Diffstat (limited to '')
-rw-r--r--spartan/theories/Spartan.thy12
1 files changed, 5 insertions, 7 deletions
diff --git a/spartan/theories/Spartan.thy b/spartan/theories/Spartan.thy
index ac12937..a79f209 100644
--- a/spartan/theories/Spartan.thy
+++ b/spartan/theories/Spartan.thy
@@ -242,7 +242,7 @@ setup \<open>
ctxt addSolver (mk_solver "" typechk_tac))
\<close>
-method reduce uses add = (simp add: comps add | subst comps, reduce add: add)+
+method reduce uses add = (simp add: comps add | subst comps)+
section \<open>Implicit notations\<close>
@@ -355,23 +355,21 @@ translations "g \<circ> f" \<leftharpoondown> "g \<circ>\<^bsub>A\<^esub> f"
subsection \<open>Identity function\<close>
-definition id where "id A \<equiv> \<lambda>x: A. x"
+abbreviation id where "id A \<equiv> \<lambda>x: A. x"
lemma
- idI [typechk]: "A: U i \<Longrightarrow> id A: A \<rightarrow> A" and
- id_comp [comps]: "x: A \<Longrightarrow> (id A) x \<equiv> x"
- unfolding id_def by reduce
+ id_type[typechk]: "A: U i \<Longrightarrow> id A: A \<rightarrow> A" and
+ id_comp [comps]: "x: A \<Longrightarrow> (id A) x \<equiv> x" \<comment> \<open>for the occasional manual rewrite\<close>
+ by reduce
lemma id_left [comps]:
assumes "f: A \<rightarrow> B" "A: U i" "B: U i"
shows "(id B) \<circ>\<^bsub>A\<^esub> f \<equiv> f"
- unfolding id_def
by (subst eta_exp[of f]) (reduce, rule eta)
lemma id_right [comps]:
assumes "f: A \<rightarrow> B" "A: U i" "B: U i"
shows "f \<circ>\<^bsub>A\<^esub> (id A) \<equiv> f"
- unfolding id_def
by (subst eta_exp[of f]) (reduce, rule eta)
lemma id_U [typechk]: