playground/coq/unfinished/metacoq-get-constr.v

346 lines
9.2 KiB
Coq
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

From MetaCoq.Template Require Import utils All.
Require Import List String.
Import ListNotations MonadNotation.
Locate List.hd.
Definition default_global_decl :=
ConstantDecl (Build_constant_body
(tVar ""%string)
None
(Monomorphic_ctx
(LevelSet.Mkt []%list, ConstraintSet.Mkt []%list))).
Locate kername.
Check (1,2,3).
(* (1, 2, 3) : (nat × nat) × nat *)
Check ((1,2),3).
(* (1, 2, 3) : (nat × nat) × nat *)
Check ind_bodies.
Definition default_kername : kername :=
(MPfile []%list, ""%string).
Definition default_kg : (kername * global_decl) :=
(default_kername, default_global_decl).
Inductive C :=
| r : C
| g : nat -> C
| b : bool -> nat -> C.
MetaCoq Quote Recursively Definition qC := C.
(* TODO *)
Definition default_oib :=
{|
ind_name := "Empty";
ind_type := tSort
{|
Universe.t_set := {|
UnivExprSet.this := [UnivExpr.npe
(NoPropLevel.lSet, false)];
UnivExprSet.is_ok := UnivExprSet.Raw.singleton_ok
(UnivExpr.npe
(NoPropLevel.lSet, false)) |};
Universe.t_ne := eq_refl |};
ind_kelim := InType;
ind_ctors := [];
ind_projs := [] |}.
Definition aux1 (p : program) : global_env := fst p.
Compute aux1 qC.
Locate ident.
Definition aux2 (g : global_env) : (kername * global_decl) := List.hd default_kg g.
Compute aux2 (aux1 qC).
Definition aux3 (kg : kername * global_decl) : global_decl := snd kg.
Compute aux3 (aux2 (aux1 qC)).
Definition t1 := aux3 (aux2 (aux1 qC)).
Check t1.
Print global_decl.
Print universes_decl.
Check (Monomorphic_ctx (LevelSet.Mkt []%list, ConstraintSet.Mkt []%list)).
Compute Build_mutual_inductive_body Finite 0 []%list []%list
(Monomorphic_ctx (LevelSet.Mkt []%list, ConstraintSet.Mkt []%list)) None.
Definition default_mib := Build_mutual_inductive_body Finite 0 []%list []%list
(Monomorphic_ctx (LevelSet.Mkt []%list, ConstraintSet.Mkt []%list)) None.
Check InductiveDecl.
Definition aux4 (g : global_decl) : list one_inductive_body :=
match g with
| ConstantDecl _ => []%list
| InductiveDecl mib => mib.(ind_bodies) (* TODO Map *)
end.
Compute aux4 (aux3 (aux2 (aux1 qC))).
Check map.
Definition aux5 (oibs : list one_inductive_body)
: list (list (ident * term * nat)) :=
map (fun o:one_inductive_body => o.(ind_ctors)) oibs.
Compute aux5 (aux4 (aux3 (aux2 (aux1 qC)))). (* CONSTRUCTORS! *)
Definition aux (p : program) : list (list (ident * term * nat)) :=
let genv := fst p in
let kgd := List.hd default_kg genv in
let gd := snd kgd in
let oibs :=
match gd with
| ConstantDecl _ => []%list
| InductiveDecl mib => mib.(ind_bodies)
end in
map (fun o:one_inductive_body => o.(ind_ctors)) oibs.
Compute aux qC.
Notation "f $ x" := (f x)
(at level 60, right associativity, only parsing).
Definition aux'' (qname : qualid) : TemplateMonad unit :=
x <- tmLocate1 qname ;;
(*x2 <- tmQuoteInductive x;;*)
match x with
| IndRef ind => (tmQuoteInductive ind.(inductive_mind)) >>= tmPrint
| _ => tmFail "not an inductive"
end.
MetaCoq Run (aux'' "C").
Compute let a:=3 in a.
Definition aux'2_1 (qname : qualid) : TemplateMonad unit :=
x <- tmLocate1 qname ;;
(*x2 <- tmQuoteInductive x;;*)
match x with
| IndRef ind =>
oibs <- tmQuoteInductive ind.(inductive_mind) ;;
(* list one_inductive_body *)
tmFail "test"
(*
rs <- map (fun o:one_inductive_body => o.(ind_ctors)) oibs ;;
tmPrint rs *)
| _ => tmFail "not an inductive"
end.
Compute aux'2_1 "C".
Definition aux'2 (qname : qualid) : TemplateMonad unit :=
x <- tmLocate1 qname ;;
(*x2 <- tmQuoteInductive x;;*)
match x with
| IndRef ind =>
mib <- tmQuoteInductive ind.(inductive_mind) ;;
(* mib : mutual_inductive_body *)
(* o.(ind_ctors) : list one_inductive_body *)
let rs := map (fun o:one_inductive_body => o.(ind_ctors))
mib.(ind_bodies) in
tmPrint rs
(* rs : list (list ((ident * term) * nat) ) *)
| _ => tmFail "not an inductive"
end.
MetaCoq Run (aux'2 "C").
Definition aux'3 (qname : qualid) : TemplateMonad unit :=
x <- tmLocate1 qname ;;
(*x2 <- tmQuoteInductive x;;*)
match x with
| IndRef ind =>
mib <- tmQuoteInductive ind.(inductive_mind) ;;
(* mib : mutual_inductive_body *)
(* o.(ind_ctors) : list one_inductive_body *)
let rs := map (fun o:one_inductive_body => o.(ind_ctors))
mib.(ind_bodies) in
res <- tmEval cbn rs ;;
tmPrint res
(* rs : list (list ((ident * term) * nat) ) *)
| _ => tmFail "not an inductive"
end.
MetaCoq Run (aux'3 "C").
Definition aux'4 (qname : qualid) : TemplateMonad unit :=
x <- tmLocate1 qname ;;
(*x2 <- tmQuoteInductive x;;*)
match x with
| IndRef ind =>
mib <- tmQuoteInductive ind.(inductive_mind) ;;
(* mib : mutual_inductive_body *)
(* m.(ind_bodies) : list one_inductive_body *)
(* o.(ind_ctors) : list ((ident * term) * nat) *)
let rs := map (fun o:one_inductive_body =>
map (fun k:((ident*term)*nat) =>
let '(i, t, n) := k in
(i, t)) o.(ind_ctors))
mib.(ind_bodies) in
res <- tmEval cbn rs ;;
tmPrint res
(* rs : list (list ((ident * term) * nat) ) *)
| _ => tmFail "not an inductive"
end.
MetaCoq Run (aux'4 "C").
Print typed_term.
Print dirpath.
Definition ident_from_inductive (ind : inductive) : ident :=
snd ind.(inductive_mind).
Compute ident_from_inductive (
{|
inductive_mind := (MPfile ["Datatypes"; "Init"; "Coq"], "nat");
inductive_ind := 0 |}).
Unset Printing Notations.
Check True \/ False.
Print or.
Set Printing Notations.
Check True \/ False.
(*
Lemma pq' : forall (p q : Prop), (p -> q) -> (~p \/ q).
*)
Lemma a : forall (n : nat), n + 0 = n.
Proof.
intros.
induction n.
- reflexivity.
- simpl.
rewrite IHn.
reflexivity.
Qed.
Lemma pq : forall (p q : Prop), (p \/ ~p) -> (p -> q) -> (~p \/ q).
Proof.
intros p q H H0.
destruct H.
- right.
apply H0.
exact H.
- left.
exact H.
Qed.
Print pq.
(*
Definition helper (ind: inductive) : TemplateMonad typed_term :=
x <- tmLocate1 ind.(inductive_mind)
*)
Definition aux'5 (qname : qualid) : TemplateMonad unit :=
x <- tmLocate1 qname ;;
(*x2 <- tmQuoteInductive x;;*)
match x with
| IndRef ind =>
mib <- tmQuoteInductive ind.(inductive_mind) ;;
(* mib : mutual_inductive_body *)
(* m.(ind_bodies) : list one_inductive_body *)
(* o.(ind_ctors) : list ((ident * term) * nat) *)
let rs := map (fun o:one_inductive_body =>
map (fun k:((ident*term)*nat) =>
let '(i, t, n) := k in
(i, t)) o.(ind_ctors))
mib.(ind_bodies) in
res <- tmEval cbn rs ;;
tmPrint res
(* rs : list (list ((ident * term) * nat) ) *)
| _ => tmFail "not an inductive"
end.
MetaCoq Run (aux'5 "C").
(* list (string * Type) *)
Definition aux'2 (qname : qualid) : TemplateMonad unit :=
x <- tmLocate1 qname ;;
(*x2 <- tmQuoteInductive x;;*)
match x with
| IndRef ind =>
oibs <- tmQuoteInductive ind.(inductive_mind) ;;
(* list one_inductive_body *)
let rs := (map (fun o:one_inductive_body => o.(ind_ctors)) oibs) in
tmPrint rs
| _ => tmFail "not an inductive"
end.
MetaCoq Run (aux'2 "C").
Definition aux'2 (qname : qualid) : TemplateMonad unit :=
x <- tmLocate1 qname ;;
(*x2 <- tmQuoteInductive x;;*)
match x with
| IndRef ind =>
x3 <- tmQuoteInductive ind.(inductive_mind) ;;
(* list one_inductive_body *)
tmPrint x3
| _ => tmFail "not an inductive"
end.
MetaCoq Run (aux'2 "C").
Definition aux' (p : program) : list (list (ident * term * nat)) :=
let gd := snd $ List.hd default_kg $ (fst p) in
map (fun o:one_inductive_body => o.(ind_ctors))
match gd with
| ConstantDecl _ => []%list
| InductiveDecl mib => mib.(ind_bodies)
end.
Compute aux' qC.
Definition aux4 (g : global_decl) : one_inductive_body :=
match g with
| ConstantDecl _ => default_oib
| InductiveDecl mib => List.hd default_oib mib.(ind_bodies) (* TODO Map *)
end.
Compute aux4 (aux3 (aux2 (aux1 qC))).
Print one_inductive_body.
Definition aux5 (o : one_inductive_body) : list (ident * term * nat) := o.(ind_ctors).
Definition default_itxn := ((""%string, tVar ""%string), 3).
Check default_itxn : ident * term * nat.
Compute aux5 (aux4 (aux3 (aux2 (aux1 qC)))). (* CONSTRUCTORS! *)
Definition t2 := aux5 (aux4 (aux3 (aux2 (aux1 qC)))). (* CONSTRUCTORS! *)
Compute t2.
Definition sample_ctor := List.last t2 default_itxn.
Compute sample_ctor.
Definition sample_ctor' := fst sample_ctor.
Compute sample_ctor'.
Definition sample_ctor'' := snd sample_ctor'.
Compute sample_ctor''.
MetaCoq Test Unquote (tRel 0).
= [("r", tRel 0, 0);
("g",
tProd nAnon
(tInd
{|
inductive_mind := (MPfile ["Datatypes"; "Init"; "Coq"], "nat");
inductive_ind := 0 |} []) (tRel 1), 1);
("b",
tProd nAnon
(tInd
{|
inductive_mind := (MPfile ["Datatypes"; "Init"; "Coq"], "bool");
inductive_ind := 0 |} [])
(tProd nAnon
(tInd
{|
inductive_mind := (MPfile ["Datatypes"; "Init"; "Coq"], "nat");
inductive_ind := 0 |} []) (tRel 2)), 2)]
: list ((ident × term) × nat)